| author | nipkow | 
| Tue, 29 Oct 2024 07:41:52 +0100 | |
| changeset 81284 | f77c6448d4d7 | 
| parent 76217 | 8655344f1cf6 | 
| permissions | -rw-r--r-- | 
| 6093 | 1 | (* Title: ZF/QUniv.thy | 
| 1478 | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 0 | 3 | Copyright 1993 University of Cambridge | 
| 4 | *) | |
| 5 | ||
| 60770 | 6 | section\<open>A Small Universe for Lazy Recursive Types\<close> | 
| 13285 | 7 | |
| 16417 | 8 | theory QUniv imports Univ QPair begin | 
| 3923 | 9 | |
| 6112 | 10 | (*Disjoint sums as a datatype*) | 
| 46820 | 11 | rep_datatype | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 12 | elimination sumE | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 13 | induction TrueI | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 14 | case_eqns case_Inl case_Inr | 
| 6112 | 15 | |
| 16 | (*Variant disjoint sums as a datatype*) | |
| 46820 | 17 | rep_datatype | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 18 | elimination qsumE | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 19 | induction TrueI | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 20 | case_eqns qcase_QInl qcase_QInr | 
| 6112 | 21 | |
| 24893 | 22 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 23 | quniv :: "i \<Rightarrow> i" where | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
63040diff
changeset | 24 | "quniv(A) \<equiv> Pow(univ(eclose(A)))" | 
| 0 | 25 | |
| 13285 | 26 | |
| 60770 | 27 | subsection\<open>Properties involving Transset and Sum\<close> | 
| 13356 | 28 | |
| 13285 | 29 | lemma Transset_includes_summands: | 
| 76214 | 30 | "\<lbrakk>Transset(C); A+B \<subseteq> C\<rbrakk> \<Longrightarrow> A \<subseteq> C \<and> B \<subseteq> C" | 
| 46820 | 31 | apply (simp add: sum_def Un_subset_iff) | 
| 13285 | 32 | apply (blast dest: Transset_includes_range) | 
| 33 | done | |
| 34 | ||
| 35 | lemma Transset_sum_Int_subset: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
63040diff
changeset | 36 | "Transset(C) \<Longrightarrow> (A+B) \<inter> C \<subseteq> (A \<inter> C) + (B \<inter> C)" | 
| 46820 | 37 | apply (simp add: sum_def Int_Un_distrib2) | 
| 13285 | 38 | apply (blast dest: Transset_Pair_D) | 
| 39 | done | |
| 40 | ||
| 60770 | 41 | subsection\<open>Introduction and Elimination Rules\<close> | 
| 13285 | 42 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
63040diff
changeset | 43 | lemma qunivI: "X \<subseteq> univ(eclose(A)) \<Longrightarrow> X \<in> quniv(A)" | 
| 13285 | 44 | by (simp add: quniv_def) | 
| 45 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
63040diff
changeset | 46 | lemma qunivD: "X \<in> quniv(A) \<Longrightarrow> X \<subseteq> univ(eclose(A))" | 
| 13285 | 47 | by (simp add: quniv_def) | 
| 48 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
63040diff
changeset | 49 | lemma quniv_mono: "A<=B \<Longrightarrow> quniv(A) \<subseteq> quniv(B)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 50 | unfolding quniv_def | 
| 13285 | 51 | apply (erule eclose_mono [THEN univ_mono, THEN Pow_mono]) | 
| 52 | done | |
| 53 | ||
| 60770 | 54 | subsection\<open>Closure Properties\<close> | 
| 13285 | 55 | |
| 46820 | 56 | lemma univ_eclose_subset_quniv: "univ(eclose(A)) \<subseteq> quniv(A)" | 
| 57 | apply (simp add: quniv_def Transset_iff_Pow [symmetric]) | |
| 13285 | 58 | apply (rule Transset_eclose [THEN Transset_univ]) | 
| 59 | done | |
| 60 | ||
| 63040 | 61 | (*Key property for proving A_subset_quniv; requires eclose in definition of quniv*) | 
| 46820 | 62 | lemma univ_subset_quniv: "univ(A) \<subseteq> quniv(A)" | 
| 13285 | 63 | apply (rule arg_subset_eclose [THEN univ_mono, THEN subset_trans]) | 
| 64 | apply (rule univ_eclose_subset_quniv) | |
| 65 | done | |
| 66 | ||
| 45602 | 67 | lemmas univ_into_quniv = univ_subset_quniv [THEN subsetD] | 
| 13285 | 68 | |
| 46820 | 69 | lemma Pow_univ_subset_quniv: "Pow(univ(A)) \<subseteq> quniv(A)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 70 | unfolding quniv_def | 
| 13285 | 71 | apply (rule arg_subset_eclose [THEN univ_mono, THEN Pow_mono]) | 
| 72 | done | |
| 73 | ||
| 74 | lemmas univ_subset_into_quniv = | |
| 45602 | 75 | PowI [THEN Pow_univ_subset_quniv [THEN subsetD]] | 
| 13285 | 76 | |
| 45602 | 77 | lemmas zero_in_quniv = zero_in_univ [THEN univ_into_quniv] | 
| 78 | lemmas one_in_quniv = one_in_univ [THEN univ_into_quniv] | |
| 79 | lemmas two_in_quniv = two_in_univ [THEN univ_into_quniv] | |
| 13285 | 80 | |
| 81 | lemmas A_subset_quniv = subset_trans [OF A_subset_univ univ_subset_quniv] | |
| 82 | ||
| 45602 | 83 | lemmas A_into_quniv = A_subset_quniv [THEN subsetD] | 
| 13285 | 84 | |
| 85 | (*** univ(A) closure for Quine-inspired pairs and injections ***) | |
| 86 | ||
| 87 | (*Quine ordered pairs*) | |
| 46820 | 88 | lemma QPair_subset_univ: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
63040diff
changeset | 89 | "\<lbrakk>a \<subseteq> univ(A); b \<subseteq> univ(A)\<rbrakk> \<Longrightarrow> <a;b> \<subseteq> univ(A)" | 
| 13285 | 90 | by (simp add: QPair_def sum_subset_univ) | 
| 91 | ||
| 60770 | 92 | subsection\<open>Quine Disjoint Sum\<close> | 
| 13285 | 93 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
63040diff
changeset | 94 | lemma QInl_subset_univ: "a \<subseteq> univ(A) \<Longrightarrow> QInl(a) \<subseteq> univ(A)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 95 | unfolding QInl_def | 
| 13285 | 96 | apply (erule empty_subsetI [THEN QPair_subset_univ]) | 
| 97 | done | |
| 98 | ||
| 46820 | 99 | lemmas naturals_subset_nat = | 
| 45602 | 100 | Ord_nat [THEN Ord_is_Transset, unfolded Transset_def, THEN bspec] | 
| 13285 | 101 | |
| 102 | lemmas naturals_subset_univ = | |
| 103 | subset_trans [OF naturals_subset_nat nat_subset_univ] | |
| 104 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
63040diff
changeset | 105 | lemma QInr_subset_univ: "a \<subseteq> univ(A) \<Longrightarrow> QInr(a) \<subseteq> univ(A)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 106 | unfolding QInr_def | 
| 13285 | 107 | apply (erule nat_1I [THEN naturals_subset_univ, THEN QPair_subset_univ]) | 
| 108 | done | |
| 109 | ||
| 60770 | 110 | subsection\<open>Closure for Quine-Inspired Products and Sums\<close> | 
| 13285 | 111 | |
| 112 | (*Quine ordered pairs*) | |
| 46820 | 113 | lemma QPair_in_quniv: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
63040diff
changeset | 114 | "\<lbrakk>a: quniv(A); b: quniv(A)\<rbrakk> \<Longrightarrow> <a;b> \<in> quniv(A)" | 
| 46820 | 115 | by (simp add: quniv_def QPair_def sum_subset_univ) | 
| 13285 | 116 | |
| 46820 | 117 | lemma QSigma_quniv: "quniv(A) <*> quniv(A) \<subseteq> quniv(A)" | 
| 13285 | 118 | by (blast intro: QPair_in_quniv) | 
| 119 | ||
| 120 | lemmas QSigma_subset_quniv = subset_trans [OF QSigma_mono QSigma_quniv] | |
| 121 | ||
| 122 | (*The opposite inclusion*) | |
| 46820 | 123 | lemma quniv_QPair_D: | 
| 76214 | 124 | "<a;b> \<in> quniv(A) \<Longrightarrow> a: quniv(A) \<and> b: quniv(A)" | 
| 76217 | 125 | unfolding quniv_def QPair_def | 
| 46820 | 126 | apply (rule Transset_includes_summands [THEN conjE]) | 
| 13285 | 127 | apply (rule Transset_eclose [THEN Transset_univ]) | 
| 46820 | 128 | apply (erule PowD, blast) | 
| 13285 | 129 | done | 
| 130 | ||
| 45602 | 131 | lemmas quniv_QPair_E = quniv_QPair_D [THEN conjE] | 
| 13285 | 132 | |
| 76214 | 133 | lemma quniv_QPair_iff: "<a;b> \<in> quniv(A) \<longleftrightarrow> a: quniv(A) \<and> b: quniv(A)" | 
| 13285 | 134 | by (blast intro: QPair_in_quniv dest: quniv_QPair_D) | 
| 135 | ||
| 136 | ||
| 60770 | 137 | subsection\<open>Quine Disjoint Sum\<close> | 
| 13285 | 138 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
63040diff
changeset | 139 | lemma QInl_in_quniv: "a: quniv(A) \<Longrightarrow> QInl(a) \<in> quniv(A)" | 
| 13285 | 140 | by (simp add: QInl_def zero_in_quniv QPair_in_quniv) | 
| 141 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
63040diff
changeset | 142 | lemma QInr_in_quniv: "b: quniv(A) \<Longrightarrow> QInr(b) \<in> quniv(A)" | 
| 13285 | 143 | by (simp add: QInr_def one_in_quniv QPair_in_quniv) | 
| 144 | ||
| 46820 | 145 | lemma qsum_quniv: "quniv(C) <+> quniv(C) \<subseteq> quniv(C)" | 
| 13285 | 146 | by (blast intro: QInl_in_quniv QInr_in_quniv) | 
| 147 | ||
| 148 | lemmas qsum_subset_quniv = subset_trans [OF qsum_mono qsum_quniv] | |
| 149 | ||
| 150 | ||
| 60770 | 151 | subsection\<open>The Natural Numbers\<close> | 
| 13285 | 152 | |
| 153 | lemmas nat_subset_quniv = subset_trans [OF nat_subset_univ univ_subset_quniv] | |
| 154 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
63040diff
changeset | 155 | (* n:nat \<Longrightarrow> n:quniv(A) *) | 
| 45602 | 156 | lemmas nat_into_quniv = nat_subset_quniv [THEN subsetD] | 
| 13285 | 157 | |
| 158 | lemmas bool_subset_quniv = subset_trans [OF bool_subset_univ univ_subset_quniv] | |
| 159 | ||
| 45602 | 160 | lemmas bool_into_quniv = bool_subset_quniv [THEN subsetD] | 
| 13285 | 161 | |
| 162 | ||
| 13356 | 163 | (*Intersecting <a;b> with Vfrom...*) | 
| 13285 | 164 | |
| 46820 | 165 | lemma QPair_Int_Vfrom_succ_subset: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
63040diff
changeset | 166 | "Transset(X) \<Longrightarrow> | 
| 46820 | 167 | <a;b> \<inter> Vfrom(X, succ(i)) \<subseteq> <a \<inter> Vfrom(X,i); b \<inter> Vfrom(X,i)>" | 
| 13285 | 168 | by (simp add: QPair_def sum_def Int_Un_distrib2 Un_mono | 
| 169 | product_Int_Vfrom_subset [THEN subset_trans] | |
| 170 | Sigma_mono [OF Int_lower1 subset_refl]) | |
| 171 | ||
| 60770 | 172 | subsection\<open>"Take-Lemma" Rules\<close> | 
| 13356 | 173 | |
| 174 | (*for proving a=b by coinduction and c: quniv(A)*) | |
| 13285 | 175 | |
| 176 | (*Rule for level i -- preserving the level, not decreasing it*) | |
| 177 | ||
| 46820 | 178 | lemma QPair_Int_Vfrom_subset: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
63040diff
changeset | 179 | "Transset(X) \<Longrightarrow> | 
| 46820 | 180 | <a;b> \<inter> Vfrom(X,i) \<subseteq> <a \<inter> Vfrom(X,i); b \<inter> Vfrom(X,i)>" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 181 | unfolding QPair_def | 
| 13285 | 182 | apply (erule Transset_Vfrom [THEN Transset_sum_Int_subset]) | 
| 183 | done | |
| 184 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
63040diff
changeset | 185 | (*@{term"\<lbrakk>a \<inter> Vset(i) \<subseteq> c; b \<inter> Vset(i) \<subseteq> d\<rbrakk> \<Longrightarrow> <a;b> \<inter> Vset(i) \<subseteq> <c;d>"}*)
 | 
| 13285 | 186 | lemmas QPair_Int_Vset_subset_trans = | 
| 187 | subset_trans [OF Transset_0 [THEN QPair_Int_Vfrom_subset] QPair_mono] | |
| 188 | ||
| 189 | lemma QPair_Int_Vset_subset_UN: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
63040diff
changeset | 190 | "Ord(i) \<Longrightarrow> <a;b> \<inter> Vset(i) \<subseteq> (\<Union>j\<in>i. <a \<inter> Vset(j); b \<inter> Vset(j)>)" | 
| 13285 | 191 | apply (erule Ord_cases) | 
| 192 | (*0 case*) | |
| 193 | apply (simp add: Vfrom_0) | |
| 194 | (*succ(j) case*) | |
| 46820 | 195 | apply (erule ssubst) | 
| 13285 | 196 | apply (rule Transset_0 [THEN QPair_Int_Vfrom_succ_subset, THEN subset_trans]) | 
| 197 | apply (rule succI1 [THEN UN_upper]) | |
| 198 | (*Limit(i) case*) | |
| 46820 | 199 | apply (simp del: UN_simps | 
| 13285 | 200 | add: Limit_Vfrom_eq Int_UN_distrib UN_mono QPair_Int_Vset_subset_trans) | 
| 201 | done | |
| 202 | ||
| 0 | 203 | end |