(* Title: ZF/ex/misc.ML ID: $Id: misc.thy,v 1.5 2005/08/17 13:10:00 paulson Exp $ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge Composition of homomorphisms, Pastre's examples, ... *) header{*Miscellaneous ZF Examples*} theory misc imports Main begin subsection{*Various Small Problems*} text{*The singleton problems are much harder in HOL.*} lemma singleton_example_1: "∀x ∈ S. ∀y ∈ S. x ⊆ y ==> ∃z. S ⊆ {z}" by blast lemma singleton_example_2: "∀x ∈ S. \<Union>S ⊆ x ==> ∃z. S ⊆ {z}" -- {*Variant of the problem above. *} by blast lemma "∃!x. f (g(x)) = x ==> ∃!y. g (f(y)) = y" -- {* A unique fixpoint theorem --- @{text fast}/@{text best}/@{text auto} all fail. *} apply (erule ex1E, rule ex1I, erule subst_context) apply (rule subst, assumption, erule allE, rule subst_context, erule mp) apply (erule subst_context) done text{*A weird property of ordered pairs.*} lemma "b≠c ==> <a,b> Int <a,c> = <a,a>" by (simp add: Pair_def Int_cons_left Int_cons_right doubleton_eq_iff, blast) text{*These two are cited in Benzmueller and Kohlhase's system description of LEO, CADE-15, 1998 (page 139-143) as theorems LEO could not prove.*} lemma "(X = Y Un Z) <-> (Y ⊆ X & Z ⊆ X & (∀V. Y ⊆ V & Z ⊆ V --> X ⊆ V))" by (blast intro!: equalityI) text{*the dual of the previous one} lemma "(X = Y Int Z) <-> (X ⊆ Y & X ⊆ Z & (∀V. V ⊆ Y & V ⊆ Z --> V ⊆ X))" by (blast intro!: equalityI) text{*trivial example of term synthesis: apparently hard for some provers!} lemma "a ≠ b ==> a:?X & b ∉ ?X" by blast text{*Nice Blast_tac benchmark. Proved in 0.3s; old tactics can't manage it!} lemma "∀x ∈ S. ∀y ∈ S. x ⊆ y ==> ∃z. S ⊆ {z}" by blast text{*variant of the benchmark above} lemma "∀x ∈ S. Union(S) ⊆ x ==> ∃z. S ⊆ {z}" by blast (*Example 12 (credited to Peter Andrews) from W. Bledsoe. A Maximal Method for Set Variables in Automatic Theorem-proving. In: J. Hayes and D. Michie and L. Mikulich, eds. Machine Intelligence 9. Ellis Horwood, 53-100 (1979). *) lemma "(∀F. {x} ∈ F --> {y} ∈ F) --> (∀A. x ∈ A --> y ∈ A)" by best text{*A characterization of functions suggested by Tobias Nipkow*} lemma "r ∈ domain(r)->B <-> r ⊆ domain(r)*B & (∀X. r `` (r -`` X) ⊆ X)" by (unfold Pi_def function_def, best) subsection{*Composition of homomorphisms is a Homomorphism*} text{*Given as a challenge problem in R. Boyer et al., Set Theory in First-Order Logic: Clauses for G\"odel's Axioms, JAR 2 (1986), 287-327 *} text{*collecting the relevant lemmas} declare comp_fun [simp] SigmaI [simp] apply_funtype [simp] (*Force helps prove conditions of rewrites such as comp_fun_apply, since rewriting does not instantiate Vars.*) lemma "(∀A f B g. hom(A,f,B,g) = {H ∈ A->B. f ∈ A*A->A & g ∈ B*B->B & (∀x ∈ A. ∀y ∈ A. H`(f`<x,y>) = g`<H`x,H`y>)}) --> J ∈ hom(A,f,B,g) & K ∈ hom(B,g,C,h) --> (K O J) ∈ hom(A,f,C,h)" by force text{*Another version, with meta-level rewriting} lemma "(!! A f B g. hom(A,f,B,g) == {H ∈ A->B. f ∈ A*A->A & g ∈ B*B->B & (∀x ∈ A. ∀y ∈ A. H`(f`<x,y>) = g`<H`x,H`y>)}) ==> J ∈ hom(A,f,B,g) & K ∈ hom(B,g,C,h) --> (K O J) ∈ hom(A,f,C,h)" by force subsection{*Pastre's Examples*} text{*D Pastre. Automatic theorem proving in set theory. Artificial Intelligence, 10:1--27, 1978. Previously, these were done using ML code, but blast manages fine.*} lemmas compIs [intro] = comp_surj comp_inj comp_fun [intro] lemmas compDs [dest] = comp_mem_injD1 comp_mem_surjD1 comp_mem_injD2 comp_mem_surjD2 lemma pastre1: "[| (h O g O f) ∈ inj(A,A); (f O h O g) ∈ surj(B,B); (g O f O h) ∈ surj(C,C); f ∈ A->B; g ∈ B->C; h ∈ C->A |] ==> h ∈ bij(C,A)"; by (unfold bij_def, blast) lemma pastre3: "[| (h O g O f) ∈ surj(A,A); (f O h O g) ∈ surj(B,B); (g O f O h) ∈ inj(C,C); f ∈ A->B; g ∈ B->C; h ∈ C->A |] ==> h ∈ bij(C,A)" by (unfold bij_def, blast) lemma pastre4: "[| (h O g O f) ∈ surj(A,A); (f O h O g) ∈ inj(B,B); (g O f O h) ∈ inj(C,C); f ∈ A->B; g ∈ B->C; h ∈ C->A |] ==> h ∈ bij(C,A)" by (unfold bij_def, blast) lemma pastre5: "[| (h O g O f) ∈ inj(A,A); (f O h O g) ∈ surj(B,B); (g O f O h) ∈ inj(C,C); f ∈ A->B; g ∈ B->C; h ∈ C->A |] ==> h ∈ bij(C,A)" by (unfold bij_def, blast) lemma pastre6: "[| (h O g O f) ∈ inj(A,A); (f O h O g) ∈ inj(B,B); (g O f O h) ∈ surj(C,C); f ∈ A->B; g ∈ B->C; h ∈ C->A |] ==> h ∈ bij(C,A)" by (unfold bij_def, blast) end
lemma singleton_example_1:
∀x∈S. ∀y∈S. x ⊆ y ==> ∃z. S ⊆ {z}
lemma singleton_example_2:
∀x∈S. \<Union>S ⊆ x ==> ∃z. S ⊆ {z}
lemma
∃!x. f(g(x)) = x ==> ∃!y. g(f(y)) = y
lemma
b ≠ c ==> 〈a, b〉 ∩ 〈a, c〉 = 〈a, a〉
lemma
X = Y ∪ Z <-> Y ⊆ X ∧ Z ⊆ X ∧ (∀V. Y ⊆ V ∧ Z ⊆ V --> X ⊆ V)
lemma
r ∈ domain(r) -> B <-> r ⊆ domain(r) × B ∧ (∀X. r `` (r -`` X) ⊆ X)
lemma compIs:
[| g ∈ surj(A, B); f ∈ surj(B, C) |] ==> f O g ∈ surj(A, C)
[| g ∈ inj(A, B); f ∈ inj(B, C) |] ==> f O g ∈ inj(A, C)
[| g ∈ A -> B; f ∈ B -> C |] ==> f O g ∈ A -> C
lemma compDs:
[| f O g ∈ inj(A, C); g ∈ A -> B; f ∈ B -> C |] ==> g ∈ inj(A, B)
[| f O g ∈ surj(A, C); g ∈ A -> B; f ∈ B -> C |] ==> f ∈ surj(B, C)
[| f O g ∈ inj(A, C); g ∈ surj(A, B); f ∈ B -> C |] ==> f ∈ inj(B, C)
[| f O g ∈ surj(A, C); g ∈ A -> B; f ∈ inj(B, C) |] ==> g ∈ surj(A, B)
lemma pastre1:
[| h O g O f ∈ inj(A, A); f O h O g ∈ surj(B, B); g O f O h ∈ surj(C, C);
f ∈ A -> B; g ∈ B -> C; h ∈ C -> A |]
==> h ∈ bij(C, A)
lemma pastre3:
[| h O g O f ∈ surj(A, A); f O h O g ∈ surj(B, B); g O f O h ∈ inj(C, C);
f ∈ A -> B; g ∈ B -> C; h ∈ C -> A |]
==> h ∈ bij(C, A)
lemma pastre4:
[| h O g O f ∈ surj(A, A); f O h O g ∈ inj(B, B); g O f O h ∈ inj(C, C);
f ∈ A -> B; g ∈ B -> C; h ∈ C -> A |]
==> h ∈ bij(C, A)
lemma pastre5:
[| h O g O f ∈ inj(A, A); f O h O g ∈ surj(B, B); g O f O h ∈ inj(C, C);
f ∈ A -> B; g ∈ B -> C; h ∈ C -> A |]
==> h ∈ bij(C, A)
lemma pastre6:
[| h O g O f ∈ inj(A, A); f O h O g ∈ inj(B, B); g O f O h ∈ surj(C, C);
f ∈ A -> B; g ∈ B -> C; h ∈ C -> A |]
==> h ∈ bij(C, A)