| author | blanchet | 
| Thu, 17 Jan 2013 13:11:44 +0100 | |
| changeset 50926 | c7f910a596ad | 
| parent 46822 | 95f1e700b712 | 
| child 58871 | c399ae4b836f | 
| permissions | -rw-r--r-- | 
| 35762 | 1 | (* Title: ZF/ex/misc.thy | 
| 11399 | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 3 | Copyright 1993 University of Cambridge | |
| 4 | ||
| 5 | Composition of homomorphisms, Pastre's examples, ... | |
| 6 | *) | |
| 7 | ||
| 14120 | 8 | header{*Miscellaneous ZF Examples*}
 | 
| 9 | ||
| 16417 | 10 | theory misc imports Main begin | 
| 11399 | 11 | |
| 17093 | 12 | |
| 14120 | 13 | subsection{*Various Small Problems*}
 | 
| 11399 | 14 | |
| 17093 | 15 | text{*The singleton problems are much harder in HOL.*}
 | 
| 16 | lemma singleton_example_1: | |
| 17 |      "\<forall>x \<in> S. \<forall>y \<in> S. x \<subseteq> y \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
 | |
| 18 | by blast | |
| 19 | ||
| 20 | lemma singleton_example_2: | |
| 21 |      "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
 | |
| 22 |   -- {*Variant of the problem above. *}
 | |
| 23 | by blast | |
| 24 | ||
| 25 | lemma "\<exists>!x. f (g(x)) = x \<Longrightarrow> \<exists>!y. g (f(y)) = y" | |
| 26 |   -- {* A unique fixpoint theorem --- @{text fast}/@{text best}/@{text auto} all fail. *} 
 | |
| 27 | apply (erule ex1E, rule ex1I, erule subst_context) | |
| 28 | apply (rule subst, assumption, erule allE, rule subst_context, erule mp) | |
| 29 | apply (erule subst_context) | |
| 30 | done | |
| 31 | ||
| 32 | ||
| 14120 | 33 | text{*A weird property of ordered pairs.*}
 | 
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
39991diff
changeset | 34 | lemma "b\<noteq>c ==> <a,b> \<inter> <a,c> = <a,a>" | 
| 14120 | 35 | by (simp add: Pair_def Int_cons_left Int_cons_right doubleton_eq_iff, blast) | 
| 36 | ||
| 37 | text{*These two are cited in Benzmueller and Kohlhase's system description of
 | |
| 38 | LEO, CADE-15, 1998 (page 139-143) as theorems LEO could not prove.*} | |
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
39991diff
changeset | 39 | lemma "(X = Y \<union> Z) \<longleftrightarrow> (Y \<subseteq> X & Z \<subseteq> X & (\<forall>V. Y \<subseteq> V & Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" | 
| 11399 | 40 | by (blast intro!: equalityI) | 
| 41 | ||
| 39991 | 42 | text{*the dual of the previous one*}
 | 
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
39991diff
changeset | 43 | lemma "(X = Y \<inter> Z) \<longleftrightarrow> (X \<subseteq> Y & X \<subseteq> Z & (\<forall>V. V \<subseteq> Y & V \<subseteq> Z \<longrightarrow> V \<subseteq> X))" | 
| 11399 | 44 | by (blast intro!: equalityI) | 
| 45 | ||
| 39991 | 46 | text{*trivial example of term synthesis: apparently hard for some provers!*}
 | 
| 47 | schematic_lemma "a \<noteq> b ==> a:?X & b \<notin> ?X" | |
| 11399 | 48 | by blast | 
| 49 | ||
| 39991 | 50 | text{*Nice blast benchmark.  Proved in 0.3s; old tactics can't manage it!*}
 | 
| 11399 | 51 | lemma "\<forall>x \<in> S. \<forall>y \<in> S. x \<subseteq> y ==> \<exists>z. S \<subseteq> {z}"
 | 
| 52 | by blast | |
| 53 | ||
| 39991 | 54 | text{*variant of the benchmark above*}
 | 
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
39991diff
changeset | 55 | lemma "\<forall>x \<in> S. \<Union>(S) \<subseteq> x ==> \<exists>z. S \<subseteq> {z}"
 | 
| 11399 | 56 | by blast | 
| 57 | ||
| 58 | (*Example 12 (credited to Peter Andrews) from | |
| 59 | W. Bledsoe. A Maximal Method for Set Variables in Automatic Theorem-proving. | |
| 60 | In: J. Hayes and D. Michie and L. Mikulich, eds. Machine Intelligence 9. | |
| 61 | Ellis Horwood, 53-100 (1979). *) | |
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
39991diff
changeset | 62 | lemma "(\<forall>F. {x} \<in> F \<longrightarrow> {y} \<in> F) \<longrightarrow> (\<forall>A. x \<in> A \<longrightarrow> y \<in> A)"
 | 
| 11399 | 63 | by best | 
| 64 | ||
| 14120 | 65 | text{*A characterization of functions suggested by Tobias Nipkow*}
 | 
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
39991diff
changeset | 66 | lemma "r \<in> domain(r)->B \<longleftrightarrow> r \<subseteq> domain(r)*B & (\<forall>X. r `` (r -`` X) \<subseteq> X)" | 
| 14120 | 67 | by (unfold Pi_def function_def, best) | 
| 11399 | 68 | |
| 69 | ||
| 14120 | 70 | subsection{*Composition of homomorphisms is a Homomorphism*}
 | 
| 71 | ||
| 72 | text{*Given as a challenge problem in
 | |
| 11399 | 73 | R. Boyer et al., | 
| 74 | Set Theory in First-Order Logic: Clauses for G\"odel's Axioms, | |
| 14120 | 75 | JAR 2 (1986), 287-327 *} | 
| 11399 | 76 | |
| 39991 | 77 | text{*collecting the relevant lemmas*}
 | 
| 11399 | 78 | declare comp_fun [simp] SigmaI [simp] apply_funtype [simp] | 
| 79 | ||
| 80 | (*Force helps prove conditions of rewrites such as comp_fun_apply, since | |
| 81 | rewriting does not instantiate Vars.*) | |
| 82 | lemma "(\<forall>A f B g. hom(A,f,B,g) = | |
| 83 |            {H \<in> A->B. f \<in> A*A->A & g \<in> B*B->B &  
 | |
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
39991diff
changeset | 84 | (\<forall>x \<in> A. \<forall>y \<in> A. H`(f`<x,y>) = g`<H`x,H`y>)}) \<longrightarrow> | 
| 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
39991diff
changeset | 85 | J \<in> hom(A,f,B,g) & K \<in> hom(B,g,C,h) \<longrightarrow> | 
| 11399 | 86 | (K O J) \<in> hom(A,f,C,h)" | 
| 87 | by force | |
| 88 | ||
| 39991 | 89 | text{*Another version, with meta-level rewriting*}
 | 
| 11399 | 90 | lemma "(!! A f B g. hom(A,f,B,g) == | 
| 91 |            {H \<in> A->B. f \<in> A*A->A & g \<in> B*B->B &  
 | |
| 92 | (\<forall>x \<in> A. \<forall>y \<in> A. H`(f`<x,y>) = g`<H`x,H`y>)}) | |
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
39991diff
changeset | 93 | ==> J \<in> hom(A,f,B,g) & K \<in> hom(B,g,C,h) \<longrightarrow> (K O J) \<in> hom(A,f,C,h)" | 
| 11399 | 94 | by force | 
| 95 | ||
| 96 | ||
| 14120 | 97 | subsection{*Pastre's Examples*}
 | 
| 11399 | 98 | |
| 14120 | 99 | text{*D Pastre.  Automatic theorem proving in set theory. 
 | 
| 100 | Artificial Intelligence, 10:1--27, 1978. | |
| 101 | Previously, these were done using ML code, but blast manages fine.*} | |
| 11399 | 102 | |
| 103 | lemmas compIs [intro] = comp_surj comp_inj comp_fun [intro] | |
| 104 | lemmas compDs [dest] = comp_mem_injD1 comp_mem_surjD1 | |
| 105 | comp_mem_injD2 comp_mem_surjD2 | |
| 106 | ||
| 107 | lemma pastre1: | |
| 108 | "[| (h O g O f) \<in> inj(A,A); | |
| 109 | (f O h O g) \<in> surj(B,B); | |
| 110 | (g O f O h) \<in> surj(C,C); | |
| 39991 | 111 | f \<in> A->B; g \<in> B->C; h \<in> C->A |] ==> h \<in> bij(C,A)" | 
| 11399 | 112 | by (unfold bij_def, blast) | 
| 113 | ||
| 114 | lemma pastre3: | |
| 115 | "[| (h O g O f) \<in> surj(A,A); | |
| 116 | (f O h O g) \<in> surj(B,B); | |
| 117 | (g O f O h) \<in> inj(C,C); | |
| 118 | f \<in> A->B; g \<in> B->C; h \<in> C->A |] ==> h \<in> bij(C,A)" | |
| 119 | by (unfold bij_def, blast) | |
| 120 | ||
| 121 | lemma pastre4: | |
| 122 | "[| (h O g O f) \<in> surj(A,A); | |
| 123 | (f O h O g) \<in> inj(B,B); | |
| 124 | (g O f O h) \<in> inj(C,C); | |
| 125 | f \<in> A->B; g \<in> B->C; h \<in> C->A |] ==> h \<in> bij(C,A)" | |
| 126 | by (unfold bij_def, blast) | |
| 127 | ||
| 128 | lemma pastre5: | |
| 129 | "[| (h O g O f) \<in> inj(A,A); | |
| 130 | (f O h O g) \<in> surj(B,B); | |
| 131 | (g O f O h) \<in> inj(C,C); | |
| 132 | f \<in> A->B; g \<in> B->C; h \<in> C->A |] ==> h \<in> bij(C,A)" | |
| 133 | by (unfold bij_def, blast) | |
| 134 | ||
| 135 | lemma pastre6: | |
| 136 | "[| (h O g O f) \<in> inj(A,A); | |
| 137 | (f O h O g) \<in> inj(B,B); | |
| 138 | (g O f O h) \<in> surj(C,C); | |
| 139 | f \<in> A->B; g \<in> B->C; h \<in> C->A |] ==> h \<in> bij(C,A)" | |
| 140 | by (unfold bij_def, blast) | |
| 141 | ||
| 142 | ||
| 143 | end | |
| 144 |