Theory Superposition_Calculus.HOL_Extra
theory HOL_Extra
imports Main
begin
lemmas = Uniq_I
lemma :
assumes "⋀x1 y1 x2 y2. P x1 y1 ⟹ P x2 y2 ⟹ (x1, y1) = (x2, y2)"
shows "∃⇩≤⇩1(x, y). P x y"
using assms
by (metis UniqI case_prodE)
lemma : "∃⇩≤⇩1x. P x ⟹ P y ⟹ ∃!x. P x"
by (iprover intro: ex1I dest: Uniq_D)
lemma : "Q ≤ P ⟹ Uniq Q ≥ Uniq P"
unfolding le_fun_def le_bool_def
by (rule impI) (simp only: Uniq_I Uniq_D)
lemma : "(⋀x. Q x ⟹ P x) ⟹ Uniq P ⟹ Uniq Q"
by (fact Uniq_antimono[unfolded le_fun_def le_bool_def, rule_format])
lemma : "(∃⇩≤⇩1x. P x) ⟹ {x. P x} = {} ∨ (∃x. {x. P x} = {x})"
using Uniq_D by fastforce
lemma :
"(∃⇩≤⇩1(x, y). P x y) ⟹ {(x, y). P x y} = {} ∨ (∃x y. {(x, y). P x y} = {(x, y)})"
using Collect_eq_if_Uniq by fastforce
lemma :
"(∀x ∈ X. ∃f. P (f x) x) ⟹ (∃f. ∀x ∈ X. P (f x) x)"
"(∃f. ∀x ∈ X. P (f x) x) ⟹ (∀x ∈ X. ∃f. P (f x) x)"
by meson+
lemma :
assumes "x ∈ set X" "f x ∉ set X" "map f X = X"
shows False
using assms
by(induction X) auto
end