| author | wenzelm | 
| Sat, 30 Dec 2023 22:16:18 +0100 | |
| changeset 79398 | a9fb2bc71435 | 
| parent 76216 | 9fc34f76b4e8 | 
| child 82967 | 73af47bc277c | 
| permissions | -rw-r--r-- | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 1 | (* Title: ZF/Univ.thy | 
| 1478 | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 0 | 3 | Copyright 1992 University of Cambridge | 
| 4 | ||
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 5 | Standard notation for Vset(i) is V(i), but users might want V for a | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 6 | variable. | 
| 516 | 7 | |
| 8 | NOTE: univ(A) could be a translation; would simplify many proofs! | |
| 6093 | 9 | But Ind_Syntax.univ refers to the constant "Univ.univ" | 
| 0 | 10 | *) | 
| 11 | ||
| 60770 | 12 | section\<open>The Cumulative Hierarchy and a Small Universe for Recursive Types\<close> | 
| 13356 | 13 | |
| 16417 | 14 | theory Univ imports Epsilon Cardinal begin | 
| 3923 | 15 | |
| 24893 | 16 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 17 | Vfrom :: "[i,i]\<Rightarrow>i" where | 
| 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 18 | "Vfrom(A,i) \<equiv> transrec(i, \<lambda>x f. A \<union> (\<Union>y\<in>x. Pow(f`y)))" | 
| 0 | 19 | |
| 24892 | 20 | abbreviation | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 21 | Vset :: "i\<Rightarrow>i" where | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 22 | "Vset(x) \<equiv> Vfrom(0,x)" | 
| 0 | 23 | |
| 3923 | 24 | |
| 24893 | 25 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 26 | Vrec :: "[i, [i,i]\<Rightarrow>i] \<Rightarrow>i" where | 
| 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 27 | "Vrec(a,H) \<equiv> transrec(rank(a), \<lambda>x g. \<lambda>z\<in>Vset(succ(x)). | 
| 46820 | 28 | H(z, \<lambda>w\<in>Vset(x). g`rank(w)`w)) ` a" | 
| 13163 | 29 | |
| 24893 | 30 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 31 | Vrecursor :: "[[i,i]\<Rightarrow>i, i] \<Rightarrow>i" where | 
| 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 32 | "Vrecursor(H,a) \<equiv> transrec(rank(a), \<lambda>x g. \<lambda>z\<in>Vset(succ(x)). | 
| 46820 | 33 | H(\<lambda>w\<in>Vset(x). g`rank(w)`w, z)) ` a" | 
| 13163 | 34 | |
| 24893 | 35 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 36 | univ :: "i\<Rightarrow>i" where | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 37 | "univ(A) \<equiv> Vfrom(A,nat)" | 
| 13163 | 38 | |
| 39 | ||
| 69593 | 40 | subsection\<open>Immediate Consequences of the Definition of \<^term>\<open>Vfrom(A,i)\<close>\<close> | 
| 13356 | 41 | |
| 60770 | 42 | text\<open>NOT SUITABLE FOR REWRITING -- RECURSIVE!\<close> | 
| 46820 | 43 | lemma Vfrom: "Vfrom(A,i) = A \<union> (\<Union>j\<in>i. Pow(Vfrom(A,j)))" | 
| 13269 | 44 | by (subst Vfrom_def [THEN def_transrec], simp) | 
| 13163 | 45 | |
| 60770 | 46 | subsubsection\<open>Monotonicity\<close> | 
| 13163 | 47 | |
| 48 | lemma Vfrom_mono [rule_format]: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 49 | "A<=B \<Longrightarrow> \<forall>j. i<=j \<longrightarrow> Vfrom(A,i) \<subseteq> Vfrom(B,j)" | 
| 13163 | 50 | apply (rule_tac a=i in eps_induct) | 
| 51 | apply (rule impI [THEN allI]) | |
| 15481 | 52 | apply (subst Vfrom [of A]) | 
| 53 | apply (subst Vfrom [of B]) | |
| 13163 | 54 | apply (erule Un_mono) | 
| 46820 | 55 | apply (erule UN_mono, blast) | 
| 13163 | 56 | done | 
| 57 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 58 | lemma VfromI: "\<lbrakk>a \<in> Vfrom(A,j); j<i\<rbrakk> \<Longrightarrow> a \<in> Vfrom(A,i)" | 
| 46820 | 59 | by (blast dest: Vfrom_mono [OF subset_refl le_imp_subset [OF leI]]) | 
| 13203 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13185diff
changeset | 60 | |
| 13163 | 61 | |
| 60770 | 62 | subsubsection\<open>A fundamental equality: Vfrom does not require ordinals!\<close> | 
| 13163 | 63 | |
| 15481 | 64 | |
| 65 | ||
| 46820 | 66 | lemma Vfrom_rank_subset1: "Vfrom(A,x) \<subseteq> Vfrom(A,rank(x))" | 
| 15481 | 67 | proof (induct x rule: eps_induct) | 
| 68 | fix x | |
| 69 | assume "\<forall>y\<in>x. Vfrom(A,y) \<subseteq> Vfrom(A,rank(y))" | |
| 70 | thus "Vfrom(A, x) \<subseteq> Vfrom(A, rank(x))" | |
| 71 | by (simp add: Vfrom [of _ x] Vfrom [of _ "rank(x)"], | |
| 72 | blast intro!: rank_lt [THEN ltD]) | |
| 73 | qed | |
| 13163 | 74 | |
| 46820 | 75 | lemma Vfrom_rank_subset2: "Vfrom(A,rank(x)) \<subseteq> Vfrom(A,x)" | 
| 13163 | 76 | apply (rule_tac a=x in eps_induct) | 
| 77 | apply (subst Vfrom) | |
| 15481 | 78 | apply (subst Vfrom, rule subset_refl [THEN Un_mono]) | 
| 13163 | 79 | apply (rule UN_least) | 
| 61798 | 80 | txt\<open>expand \<open>rank(x1) = (\<Union>y\<in>x1. succ(rank(y)))\<close> in assumptions\<close> | 
| 13163 | 81 | apply (erule rank [THEN equalityD1, THEN subsetD, THEN UN_E]) | 
| 82 | apply (rule subset_trans) | |
| 83 | apply (erule_tac [2] UN_upper) | |
| 84 | apply (rule subset_refl [THEN Vfrom_mono, THEN subset_trans, THEN Pow_mono]) | |
| 85 | apply (erule ltI [THEN le_imp_subset]) | |
| 86 | apply (rule Ord_rank [THEN Ord_succ]) | |
| 87 | apply (erule bspec, assumption) | |
| 88 | done | |
| 89 | ||
| 90 | lemma Vfrom_rank_eq: "Vfrom(A,rank(x)) = Vfrom(A,x)" | |
| 91 | apply (rule equalityI) | |
| 92 | apply (rule Vfrom_rank_subset2) | |
| 93 | apply (rule Vfrom_rank_subset1) | |
| 94 | done | |
| 95 | ||
| 96 | ||
| 60770 | 97 | subsection\<open>Basic Closure Properties\<close> | 
| 13163 | 98 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 99 | lemma zero_in_Vfrom: "y:x \<Longrightarrow> 0 \<in> Vfrom(A,x)" | 
| 13163 | 100 | by (subst Vfrom, blast) | 
| 101 | ||
| 46820 | 102 | lemma i_subset_Vfrom: "i \<subseteq> Vfrom(A,i)" | 
| 13163 | 103 | apply (rule_tac a=i in eps_induct) | 
| 104 | apply (subst Vfrom, blast) | |
| 105 | done | |
| 106 | ||
| 46820 | 107 | lemma A_subset_Vfrom: "A \<subseteq> Vfrom(A,i)" | 
| 13163 | 108 | apply (subst Vfrom) | 
| 109 | apply (rule Un_upper1) | |
| 110 | done | |
| 111 | ||
| 112 | lemmas A_into_Vfrom = A_subset_Vfrom [THEN subsetD] | |
| 113 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 114 | lemma subset_mem_Vfrom: "a \<subseteq> Vfrom(A,i) \<Longrightarrow> a \<in> Vfrom(A,succ(i))" | 
| 13163 | 115 | by (subst Vfrom, blast) | 
| 116 | ||
| 60770 | 117 | subsubsection\<open>Finite sets and ordered pairs\<close> | 
| 13163 | 118 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 119 | lemma singleton_in_Vfrom: "a \<in> Vfrom(A,i) \<Longrightarrow> {a} \<in> Vfrom(A,succ(i))"
 | 
| 13163 | 120 | by (rule subset_mem_Vfrom, safe) | 
| 121 | ||
| 122 | lemma doubleton_in_Vfrom: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 123 |      "\<lbrakk>a \<in> Vfrom(A,i);  b \<in> Vfrom(A,i)\<rbrakk> \<Longrightarrow> {a,b} \<in> Vfrom(A,succ(i))"
 | 
| 13163 | 124 | by (rule subset_mem_Vfrom, safe) | 
| 125 | ||
| 126 | lemma Pair_in_Vfrom: | |
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 127 | "\<lbrakk>a \<in> Vfrom(A,i); b \<in> Vfrom(A,i)\<rbrakk> \<Longrightarrow> \<langle>a,b\<rangle> \<in> Vfrom(A,succ(succ(i)))" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 128 | unfolding Pair_def | 
| 46820 | 129 | apply (blast intro: doubleton_in_Vfrom) | 
| 13163 | 130 | done | 
| 131 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 132 | lemma succ_in_Vfrom: "a \<subseteq> Vfrom(A,i) \<Longrightarrow> succ(a) \<in> Vfrom(A,succ(succ(i)))" | 
| 13163 | 133 | apply (intro subset_mem_Vfrom succ_subsetI, assumption) | 
| 46820 | 134 | apply (erule subset_trans) | 
| 135 | apply (rule Vfrom_mono [OF subset_refl subset_succI]) | |
| 13163 | 136 | done | 
| 137 | ||
| 69593 | 138 | subsection\<open>0, Successor and Limit Equations for \<^term>\<open>Vfrom\<close>\<close> | 
| 13163 | 139 | |
| 140 | lemma Vfrom_0: "Vfrom(A,0) = A" | |
| 141 | by (subst Vfrom, blast) | |
| 142 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 143 | lemma Vfrom_succ_lemma: "Ord(i) \<Longrightarrow> Vfrom(A,succ(i)) = A \<union> Pow(Vfrom(A,i))" | 
| 13163 | 144 | apply (rule Vfrom [THEN trans]) | 
| 46820 | 145 | apply (rule equalityI [THEN subst_context, | 
| 13163 | 146 | OF _ succI1 [THEN RepFunI, THEN Union_upper]]) | 
| 147 | apply (rule UN_least) | |
| 148 | apply (rule subset_refl [THEN Vfrom_mono, THEN Pow_mono]) | |
| 149 | apply (erule ltI [THEN le_imp_subset]) | |
| 150 | apply (erule Ord_succ) | |
| 151 | done | |
| 152 | ||
| 46820 | 153 | lemma Vfrom_succ: "Vfrom(A,succ(i)) = A \<union> Pow(Vfrom(A,i))" | 
| 13163 | 154 | apply (rule_tac x1 = "succ (i)" in Vfrom_rank_eq [THEN subst]) | 
| 13784 | 155 | apply (rule_tac x1 = i in Vfrom_rank_eq [THEN subst]) | 
| 13163 | 156 | apply (subst rank_succ) | 
| 157 | apply (rule Ord_rank [THEN Vfrom_succ_lemma]) | |
| 158 | done | |
| 159 | ||
| 160 | (*The premise distinguishes this from Vfrom(A,0); allowing X=0 forces | |
| 46820 | 161 | the conclusion to be Vfrom(A,\<Union>(X)) = A \<union> (\<Union>y\<in>X. Vfrom(A,y)) *) | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 162 | lemma Vfrom_Union: "y:X \<Longrightarrow> Vfrom(A,\<Union>(X)) = (\<Union>y\<in>X. Vfrom(A,y))" | 
| 13163 | 163 | apply (subst Vfrom) | 
| 164 | apply (rule equalityI) | |
| 60770 | 165 | txt\<open>first inclusion\<close> | 
| 13163 | 166 | apply (rule Un_least) | 
| 167 | apply (rule A_subset_Vfrom [THEN subset_trans]) | |
| 168 | apply (rule UN_upper, assumption) | |
| 169 | apply (rule UN_least) | |
| 170 | apply (erule UnionE) | |
| 171 | apply (rule subset_trans) | |
| 172 | apply (erule_tac [2] UN_upper, | |
| 173 | subst Vfrom, erule subset_trans [OF UN_upper Un_upper2]) | |
| 60770 | 174 | txt\<open>opposite inclusion\<close> | 
| 13163 | 175 | apply (rule UN_least) | 
| 176 | apply (subst Vfrom, blast) | |
| 177 | done | |
| 178 | ||
| 69593 | 179 | subsection\<open>\<^term>\<open>Vfrom\<close> applied to Limit Ordinals\<close> | 
| 13163 | 180 | |
| 181 | (*NB. limit ordinals are non-empty: | |
| 46820 | 182 | Vfrom(A,0) = A = A \<union> (\<Union>y\<in>0. Vfrom(A,y)) *) | 
| 13163 | 183 | lemma Limit_Vfrom_eq: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 184 | "Limit(i) \<Longrightarrow> Vfrom(A,i) = (\<Union>y\<in>i. Vfrom(A,y))" | 
| 13163 | 185 | apply (rule Limit_has_0 [THEN ltD, THEN Vfrom_Union, THEN subst], assumption) | 
| 46820 | 186 | apply (simp add: Limit_Union_eq) | 
| 13163 | 187 | done | 
| 188 | ||
| 189 | lemma Limit_VfromE: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 190 | "\<lbrakk>a \<in> Vfrom(A,i); \<not>R \<Longrightarrow> Limit(i); | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 191 | \<And>x. \<lbrakk>x<i; a \<in> Vfrom(A,x)\<rbrakk> \<Longrightarrow> R | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 192 | \<rbrakk> \<Longrightarrow> R" | 
| 13163 | 193 | apply (rule classical) | 
| 194 | apply (rule Limit_Vfrom_eq [THEN equalityD1, THEN subsetD, THEN UN_E]) | |
| 13203 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13185diff
changeset | 195 | prefer 2 apply assumption | 
| 46820 | 196 | apply blast | 
| 13203 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13185diff
changeset | 197 | apply (blast intro: ltI Limit_is_Ord) | 
| 13163 | 198 | done | 
| 199 | ||
| 200 | lemma singleton_in_VLimit: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 201 |     "\<lbrakk>a \<in> Vfrom(A,i);  Limit(i)\<rbrakk> \<Longrightarrow> {a} \<in> Vfrom(A,i)"
 | 
| 13163 | 202 | apply (erule Limit_VfromE, assumption) | 
| 13203 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13185diff
changeset | 203 | apply (erule singleton_in_Vfrom [THEN VfromI]) | 
| 46820 | 204 | apply (blast intro: Limit_has_succ) | 
| 13163 | 205 | done | 
| 206 | ||
| 46820 | 207 | lemmas Vfrom_UnI1 = | 
| 45602 | 208 | Un_upper1 [THEN subset_refl [THEN Vfrom_mono, THEN subsetD]] | 
| 46820 | 209 | lemmas Vfrom_UnI2 = | 
| 45602 | 210 | Un_upper2 [THEN subset_refl [THEN Vfrom_mono, THEN subsetD]] | 
| 13163 | 211 | |
| 60770 | 212 | text\<open>Hard work is finding a single j:i such that {a,b}<=Vfrom(A,j)\<close>
 | 
| 13163 | 213 | lemma doubleton_in_VLimit: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 214 |     "\<lbrakk>a \<in> Vfrom(A,i);  b \<in> Vfrom(A,i);  Limit(i)\<rbrakk> \<Longrightarrow> {a,b} \<in> Vfrom(A,i)"
 | 
| 13163 | 215 | apply (erule Limit_VfromE, assumption) | 
| 216 | apply (erule Limit_VfromE, assumption) | |
| 13203 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13185diff
changeset | 217 | apply (blast intro: VfromI [OF doubleton_in_Vfrom] | 
| 13163 | 218 | Vfrom_UnI1 Vfrom_UnI2 Limit_has_succ Un_least_lt) | 
| 219 | done | |
| 220 | ||
| 221 | lemma Pair_in_VLimit: | |
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 222 | "\<lbrakk>a \<in> Vfrom(A,i); b \<in> Vfrom(A,i); Limit(i)\<rbrakk> \<Longrightarrow> \<langle>a,b\<rangle> \<in> Vfrom(A,i)" | 
| 60770 | 223 | txt\<open>Infer that a, b occur at ordinals x,xa < i.\<close> | 
| 13163 | 224 | apply (erule Limit_VfromE, assumption) | 
| 225 | apply (erule Limit_VfromE, assumption) | |
| 69593 | 226 | txt\<open>Infer that \<^term>\<open>succ(succ(x \<union> xa)) < i\<close>\<close> | 
| 13203 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13185diff
changeset | 227 | apply (blast intro: VfromI [OF Pair_in_Vfrom] | 
| 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13185diff
changeset | 228 | Vfrom_UnI1 Vfrom_UnI2 Limit_has_succ Un_least_lt) | 
| 13163 | 229 | done | 
| 230 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 231 | lemma product_VLimit: "Limit(i) \<Longrightarrow> Vfrom(A,i) * Vfrom(A,i) \<subseteq> Vfrom(A,i)" | 
| 13163 | 232 | by (blast intro: Pair_in_VLimit) | 
| 233 | ||
| 234 | lemmas Sigma_subset_VLimit = | |
| 235 | subset_trans [OF Sigma_mono product_VLimit] | |
| 236 | ||
| 237 | lemmas nat_subset_VLimit = | |
| 238 | subset_trans [OF nat_le_Limit [THEN le_imp_subset] i_subset_Vfrom] | |
| 239 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 240 | lemma nat_into_VLimit: "\<lbrakk>n: nat; Limit(i)\<rbrakk> \<Longrightarrow> n \<in> Vfrom(A,i)" | 
| 13163 | 241 | by (blast intro: nat_subset_VLimit [THEN subsetD]) | 
| 242 | ||
| 60770 | 243 | subsubsection\<open>Closure under Disjoint Union\<close> | 
| 13163 | 244 | |
| 45602 | 245 | lemmas zero_in_VLimit = Limit_has_0 [THEN ltD, THEN zero_in_Vfrom] | 
| 13163 | 246 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 247 | lemma one_in_VLimit: "Limit(i) \<Longrightarrow> 1 \<in> Vfrom(A,i)" | 
| 13163 | 248 | by (blast intro: nat_into_VLimit) | 
| 249 | ||
| 250 | lemma Inl_in_VLimit: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 251 | "\<lbrakk>a \<in> Vfrom(A,i); Limit(i)\<rbrakk> \<Longrightarrow> Inl(a) \<in> Vfrom(A,i)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 252 | unfolding Inl_def | 
| 13163 | 253 | apply (blast intro: zero_in_VLimit Pair_in_VLimit) | 
| 254 | done | |
| 255 | ||
| 256 | lemma Inr_in_VLimit: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 257 | "\<lbrakk>b \<in> Vfrom(A,i); Limit(i)\<rbrakk> \<Longrightarrow> Inr(b) \<in> Vfrom(A,i)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 258 | unfolding Inr_def | 
| 13163 | 259 | apply (blast intro: one_in_VLimit Pair_in_VLimit) | 
| 260 | done | |
| 261 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 262 | lemma sum_VLimit: "Limit(i) \<Longrightarrow> Vfrom(C,i)+Vfrom(C,i) \<subseteq> Vfrom(C,i)" | 
| 13163 | 263 | by (blast intro!: Inl_in_VLimit Inr_in_VLimit) | 
| 264 | ||
| 265 | lemmas sum_subset_VLimit = subset_trans [OF sum_mono sum_VLimit] | |
| 266 | ||
| 267 | ||
| 268 | ||
| 69593 | 269 | subsection\<open>Properties assuming \<^term>\<open>Transset(A)\<close>\<close> | 
| 13163 | 270 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 271 | lemma Transset_Vfrom: "Transset(A) \<Longrightarrow> Transset(Vfrom(A,i))" | 
| 13163 | 272 | apply (rule_tac a=i in eps_induct) | 
| 273 | apply (subst Vfrom) | |
| 274 | apply (blast intro!: Transset_Union_family Transset_Un Transset_Pow) | |
| 275 | done | |
| 276 | ||
| 277 | lemma Transset_Vfrom_succ: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 278 | "Transset(A) \<Longrightarrow> Vfrom(A, succ(i)) = Pow(Vfrom(A,i))" | 
| 13163 | 279 | apply (rule Vfrom_succ [THEN trans]) | 
| 280 | apply (rule equalityI [OF _ Un_upper2]) | |
| 281 | apply (rule Un_least [OF _ subset_refl]) | |
| 282 | apply (rule A_subset_Vfrom [THEN subset_trans]) | |
| 283 | apply (erule Transset_Vfrom [THEN Transset_iff_Pow [THEN iffD1]]) | |
| 284 | done | |
| 285 | ||
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 286 | lemma Transset_Pair_subset: "\<lbrakk>\<langle>a,b\<rangle> \<subseteq> C; Transset(C)\<rbrakk> \<Longrightarrow> a: C \<and> b: C" | 
| 13163 | 287 | by (unfold Pair_def Transset_def, blast) | 
| 288 | ||
| 289 | lemma Transset_Pair_subset_VLimit: | |
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 290 | "\<lbrakk>\<langle>a,b\<rangle> \<subseteq> Vfrom(A,i); Transset(A); Limit(i)\<rbrakk> | 
| 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 291 | \<Longrightarrow> \<langle>a,b\<rangle> \<in> Vfrom(A,i)" | 
| 13163 | 292 | apply (erule Transset_Pair_subset [THEN conjE]) | 
| 293 | apply (erule Transset_Vfrom) | |
| 294 | apply (blast intro: Pair_in_VLimit) | |
| 295 | done | |
| 296 | ||
| 297 | lemma Union_in_Vfrom: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 298 | "\<lbrakk>X \<in> Vfrom(A,j); Transset(A)\<rbrakk> \<Longrightarrow> \<Union>(X) \<in> Vfrom(A, succ(j))" | 
| 13163 | 299 | apply (drule Transset_Vfrom) | 
| 300 | apply (rule subset_mem_Vfrom) | |
| 301 | apply (unfold Transset_def, blast) | |
| 302 | done | |
| 303 | ||
| 304 | lemma Union_in_VLimit: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 305 | "\<lbrakk>X \<in> Vfrom(A,i); Limit(i); Transset(A)\<rbrakk> \<Longrightarrow> \<Union>(X) \<in> Vfrom(A,i)" | 
| 13163 | 306 | apply (rule Limit_VfromE, assumption+) | 
| 13203 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13185diff
changeset | 307 | apply (blast intro: Limit_has_succ VfromI Union_in_Vfrom) | 
| 13163 | 308 | done | 
| 309 | ||
| 310 | ||
| 311 | (*** Closure under product/sum applied to elements -- thus Vfrom(A,i) | |
| 312 | is a model of simple type theory provided A is a transitive set | |
| 313 | and i is a limit ordinal | |
| 314 | ***) | |
| 315 | ||
| 60770 | 316 | text\<open>General theorem for membership in Vfrom(A,i) when i is a limit ordinal\<close> | 
| 13163 | 317 | lemma in_VLimit: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 318 | "\<lbrakk>a \<in> Vfrom(A,i); b \<in> Vfrom(A,i); Limit(i); | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 319 | \<And>x y j. \<lbrakk>j<i; 1:j; x \<in> Vfrom(A,j); y \<in> Vfrom(A,j)\<rbrakk> | 
| 76214 | 320 | \<Longrightarrow> \<exists>k. h(x,y) \<in> Vfrom(A,k) \<and> k<i\<rbrakk> | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 321 | \<Longrightarrow> h(a,b) \<in> Vfrom(A,i)" | 
| 60770 | 322 | txt\<open>Infer that a, b occur at ordinals x,xa < i.\<close> | 
| 13163 | 323 | apply (erule Limit_VfromE, assumption) | 
| 324 | apply (erule Limit_VfromE, assumption, atomize) | |
| 46820 | 325 | apply (drule_tac x=a in spec) | 
| 326 | apply (drule_tac x=b in spec) | |
| 327 | apply (drule_tac x="x \<union> xa \<union> 2" in spec) | |
| 328 | apply (simp add: Un_least_lt_iff lt_Ord Vfrom_UnI1 Vfrom_UnI2) | |
| 13203 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13185diff
changeset | 329 | apply (blast intro: Limit_has_0 Limit_has_succ VfromI) | 
| 13163 | 330 | done | 
| 331 | ||
| 60770 | 332 | subsubsection\<open>Products\<close> | 
| 13163 | 333 | |
| 334 | lemma prod_in_Vfrom: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 335 | "\<lbrakk>a \<in> Vfrom(A,j); b \<in> Vfrom(A,j); Transset(A)\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 336 | \<Longrightarrow> a*b \<in> Vfrom(A, succ(succ(succ(j))))" | 
| 13163 | 337 | apply (drule Transset_Vfrom) | 
| 338 | apply (rule subset_mem_Vfrom) | |
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 339 | unfolding Transset_def | 
| 13163 | 340 | apply (blast intro: Pair_in_Vfrom) | 
| 341 | done | |
| 342 | ||
| 343 | lemma prod_in_VLimit: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 344 | "\<lbrakk>a \<in> Vfrom(A,i); b \<in> Vfrom(A,i); Limit(i); Transset(A)\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 345 | \<Longrightarrow> a*b \<in> Vfrom(A,i)" | 
| 13163 | 346 | apply (erule in_VLimit, assumption+) | 
| 347 | apply (blast intro: prod_in_Vfrom Limit_has_succ) | |
| 348 | done | |
| 349 | ||
| 60770 | 350 | subsubsection\<open>Disjoint Sums, or Quine Ordered Pairs\<close> | 
| 13163 | 351 | |
| 352 | lemma sum_in_Vfrom: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 353 | "\<lbrakk>a \<in> Vfrom(A,j); b \<in> Vfrom(A,j); Transset(A); 1:j\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 354 | \<Longrightarrow> a+b \<in> Vfrom(A, succ(succ(succ(j))))" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 355 | unfolding sum_def | 
| 13163 | 356 | apply (drule Transset_Vfrom) | 
| 357 | apply (rule subset_mem_Vfrom) | |
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 358 | unfolding Transset_def | 
| 13163 | 359 | apply (blast intro: zero_in_Vfrom Pair_in_Vfrom i_subset_Vfrom [THEN subsetD]) | 
| 360 | done | |
| 361 | ||
| 362 | lemma sum_in_VLimit: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 363 | "\<lbrakk>a \<in> Vfrom(A,i); b \<in> Vfrom(A,i); Limit(i); Transset(A)\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 364 | \<Longrightarrow> a+b \<in> Vfrom(A,i)" | 
| 13163 | 365 | apply (erule in_VLimit, assumption+) | 
| 366 | apply (blast intro: sum_in_Vfrom Limit_has_succ) | |
| 367 | done | |
| 368 | ||
| 60770 | 369 | subsubsection\<open>Function Space!\<close> | 
| 13163 | 370 | |
| 371 | lemma fun_in_Vfrom: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 372 | "\<lbrakk>a \<in> Vfrom(A,j); b \<in> Vfrom(A,j); Transset(A)\<rbrakk> \<Longrightarrow> | 
| 13220 | 373 | a->b \<in> Vfrom(A, succ(succ(succ(succ(j)))))" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 374 | unfolding Pi_def | 
| 13163 | 375 | apply (drule Transset_Vfrom) | 
| 376 | apply (rule subset_mem_Vfrom) | |
| 377 | apply (rule Collect_subset [THEN subset_trans]) | |
| 378 | apply (subst Vfrom) | |
| 379 | apply (rule subset_trans [THEN subset_trans]) | |
| 380 | apply (rule_tac [3] Un_upper2) | |
| 381 | apply (rule_tac [2] succI1 [THEN UN_upper]) | |
| 382 | apply (rule Pow_mono) | |
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 383 | unfolding Transset_def | 
| 13163 | 384 | apply (blast intro: Pair_in_Vfrom) | 
| 385 | done | |
| 386 | ||
| 387 | lemma fun_in_VLimit: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 388 | "\<lbrakk>a \<in> Vfrom(A,i); b \<in> Vfrom(A,i); Limit(i); Transset(A)\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 389 | \<Longrightarrow> a->b \<in> Vfrom(A,i)" | 
| 13163 | 390 | apply (erule in_VLimit, assumption+) | 
| 391 | apply (blast intro: fun_in_Vfrom Limit_has_succ) | |
| 392 | done | |
| 393 | ||
| 394 | lemma Pow_in_Vfrom: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 395 | "\<lbrakk>a \<in> Vfrom(A,j); Transset(A)\<rbrakk> \<Longrightarrow> Pow(a) \<in> Vfrom(A, succ(succ(j)))" | 
| 13163 | 396 | apply (drule Transset_Vfrom) | 
| 397 | apply (rule subset_mem_Vfrom) | |
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 398 | unfolding Transset_def | 
| 13163 | 399 | apply (subst Vfrom, blast) | 
| 400 | done | |
| 401 | ||
| 402 | lemma Pow_in_VLimit: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 403 | "\<lbrakk>a \<in> Vfrom(A,i); Limit(i); Transset(A)\<rbrakk> \<Longrightarrow> Pow(a) \<in> Vfrom(A,i)" | 
| 13203 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 paulson parents: 
13185diff
changeset | 404 | by (blast elim: Limit_VfromE intro: Limit_has_succ Pow_in_Vfrom VfromI) | 
| 13163 | 405 | |
| 406 | ||
| 69593 | 407 | subsection\<open>The Set \<^term>\<open>Vset(i)\<close>\<close> | 
| 13163 | 408 | |
| 13220 | 409 | lemma Vset: "Vset(i) = (\<Union>j\<in>i. Pow(Vset(j)))" | 
| 13163 | 410 | by (subst Vfrom, blast) | 
| 411 | ||
| 45602 | 412 | lemmas Vset_succ = Transset_0 [THEN Transset_Vfrom_succ] | 
| 413 | lemmas Transset_Vset = Transset_0 [THEN Transset_Vfrom] | |
| 13163 | 414 | |
| 69593 | 415 | subsubsection\<open>Characterisation of the elements of \<^term>\<open>Vset(i)\<close>\<close> | 
| 13163 | 416 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 417 | lemma VsetD [rule_format]: "Ord(i) \<Longrightarrow> \<forall>b. b \<in> Vset(i) \<longrightarrow> rank(b) < i" | 
| 13163 | 418 | apply (erule trans_induct) | 
| 419 | apply (subst Vset, safe) | |
| 420 | apply (subst rank) | |
| 46820 | 421 | apply (blast intro: ltI UN_succ_least_lt) | 
| 13163 | 422 | done | 
| 423 | ||
| 424 | lemma VsetI_lemma [rule_format]: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 425 | "Ord(i) \<Longrightarrow> \<forall>b. rank(b) \<in> i \<longrightarrow> b \<in> Vset(i)" | 
| 13163 | 426 | apply (erule trans_induct) | 
| 427 | apply (rule allI) | |
| 428 | apply (subst Vset) | |
| 429 | apply (blast intro!: rank_lt [THEN ltD]) | |
| 430 | done | |
| 431 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 432 | lemma VsetI: "rank(x)<i \<Longrightarrow> x \<in> Vset(i)" | 
| 13163 | 433 | by (blast intro: VsetI_lemma elim: ltE) | 
| 434 | ||
| 60770 | 435 | text\<open>Merely a lemma for the next result\<close> | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 436 | lemma Vset_Ord_rank_iff: "Ord(i) \<Longrightarrow> b \<in> Vset(i) \<longleftrightarrow> rank(b) < i" | 
| 13163 | 437 | by (blast intro: VsetD VsetI) | 
| 438 | ||
| 46821 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
46820diff
changeset | 439 | lemma Vset_rank_iff [simp]: "b \<in> Vset(a) \<longleftrightarrow> rank(b) < rank(a)" | 
| 13163 | 440 | apply (rule Vfrom_rank_eq [THEN subst]) | 
| 441 | apply (rule Ord_rank [THEN Vset_Ord_rank_iff]) | |
| 442 | done | |
| 443 | ||
| 60770 | 444 | text\<open>This is rank(rank(a)) = rank(a)\<close> | 
| 13163 | 445 | declare Ord_rank [THEN rank_of_Ord, simp] | 
| 446 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 447 | lemma rank_Vset: "Ord(i) \<Longrightarrow> rank(Vset(i)) = i" | 
| 13163 | 448 | apply (subst rank) | 
| 449 | apply (rule equalityI, safe) | |
| 46820 | 450 | apply (blast intro: VsetD [THEN ltD]) | 
| 451 | apply (blast intro: VsetD [THEN ltD] Ord_trans) | |
| 13163 | 452 | apply (blast intro: i_subset_Vfrom [THEN subsetD] | 
| 453 | Ord_in_Ord [THEN rank_of_Ord, THEN ssubst]) | |
| 454 | done | |
| 455 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 456 | lemma Finite_Vset: "i \<in> nat \<Longrightarrow> Finite(Vset(i))" | 
| 13269 | 457 | apply (erule nat_induct) | 
| 46820 | 458 | apply (simp add: Vfrom_0) | 
| 459 | apply (simp add: Vset_succ) | |
| 13269 | 460 | done | 
| 461 | ||
| 60770 | 462 | subsubsection\<open>Reasoning about Sets in Terms of Their Elements' Ranks\<close> | 
| 0 | 463 | |
| 46820 | 464 | lemma arg_subset_Vset_rank: "a \<subseteq> Vset(rank(a))" | 
| 13163 | 465 | apply (rule subsetI) | 
| 466 | apply (erule rank_lt [THEN VsetI]) | |
| 467 | done | |
| 468 | ||
| 469 | lemma Int_Vset_subset: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 470 | "\<lbrakk>\<And>i. Ord(i) \<Longrightarrow> a \<inter> Vset(i) \<subseteq> b\<rbrakk> \<Longrightarrow> a \<subseteq> b" | 
| 46820 | 471 | apply (rule subset_trans) | 
| 13163 | 472 | apply (rule Int_greatest [OF subset_refl arg_subset_Vset_rank]) | 
| 46820 | 473 | apply (blast intro: Ord_rank) | 
| 13163 | 474 | done | 
| 475 | ||
| 60770 | 476 | subsubsection\<open>Set Up an Environment for Simplification\<close> | 
| 13163 | 477 | |
| 478 | lemma rank_Inl: "rank(a) < rank(Inl(a))" | |
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 479 | unfolding Inl_def | 
| 13163 | 480 | apply (rule rank_pair2) | 
| 481 | done | |
| 482 | ||
| 483 | lemma rank_Inr: "rank(a) < rank(Inr(a))" | |
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 484 | unfolding Inr_def | 
| 13163 | 485 | apply (rule rank_pair2) | 
| 486 | done | |
| 487 | ||
| 488 | lemmas rank_rls = rank_Inl rank_Inr rank_pair1 rank_pair2 | |
| 489 | ||
| 60770 | 490 | subsubsection\<open>Recursion over Vset Levels!\<close> | 
| 13163 | 491 | |
| 60770 | 492 | text\<open>NOT SUITABLE FOR REWRITING: recursive!\<close> | 
| 46820 | 493 | lemma Vrec: "Vrec(a,H) = H(a, \<lambda>x\<in>Vset(rank(a)). Vrec(x,H))" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 494 | unfolding Vrec_def | 
| 13269 | 495 | apply (subst transrec, simp) | 
| 13175 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 paulson parents: 
13163diff
changeset | 496 | apply (rule refl [THEN lam_cong, THEN subst_context], simp add: lt_def) | 
| 13163 | 497 | done | 
| 498 | ||
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 499 | text\<open>This form avoids giant explosions in proofs. NOTE the form of the premise!\<close> | 
| 13163 | 500 | lemma def_Vrec: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 501 | "\<lbrakk>\<And>x. h(x)\<equiv>Vrec(x,H)\<rbrakk> \<Longrightarrow> | 
| 46820 | 502 | h(a) = H(a, \<lambda>x\<in>Vset(rank(a)). h(x))" | 
| 503 | apply simp | |
| 13163 | 504 | apply (rule Vrec) | 
| 505 | done | |
| 506 | ||
| 60770 | 507 | text\<open>NOT SUITABLE FOR REWRITING: recursive!\<close> | 
| 13163 | 508 | lemma Vrecursor: | 
| 46820 | 509 | "Vrecursor(H,a) = H(\<lambda>x\<in>Vset(rank(a)). Vrecursor(H,x), a)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 510 | unfolding Vrecursor_def | 
| 13163 | 511 | apply (subst transrec, simp) | 
| 13175 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 paulson parents: 
13163diff
changeset | 512 | apply (rule refl [THEN lam_cong, THEN subst_context], simp add: lt_def) | 
| 13163 | 513 | done | 
| 514 | ||
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 515 | text\<open>This form avoids giant explosions in proofs. NOTE the form of the premise!\<close> | 
| 13163 | 516 | lemma def_Vrecursor: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 517 | "h \<equiv> Vrecursor(H) \<Longrightarrow> h(a) = H(\<lambda>x\<in>Vset(rank(a)). h(x), a)" | 
| 13163 | 518 | apply simp | 
| 519 | apply (rule Vrecursor) | |
| 520 | done | |
| 521 | ||
| 522 | ||
| 69593 | 523 | subsection\<open>The Datatype Universe: \<^term>\<open>univ(A)\<close>\<close> | 
| 13163 | 524 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 525 | lemma univ_mono: "A<=B \<Longrightarrow> univ(A) \<subseteq> univ(B)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 526 | unfolding univ_def | 
| 13163 | 527 | apply (erule Vfrom_mono) | 
| 528 | apply (rule subset_refl) | |
| 529 | done | |
| 530 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 531 | lemma Transset_univ: "Transset(A) \<Longrightarrow> Transset(univ(A))" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 532 | unfolding univ_def | 
| 13163 | 533 | apply (erule Transset_Vfrom) | 
| 534 | done | |
| 535 | ||
| 69593 | 536 | subsubsection\<open>The Set \<^term>\<open>univ(A)\<close> as a Limit\<close> | 
| 13163 | 537 | |
| 13220 | 538 | lemma univ_eq_UN: "univ(A) = (\<Union>i\<in>nat. Vfrom(A,i))" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 539 | unfolding univ_def | 
| 13163 | 540 | apply (rule Limit_nat [THEN Limit_Vfrom_eq]) | 
| 541 | done | |
| 542 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 543 | lemma subset_univ_eq_Int: "c \<subseteq> univ(A) \<Longrightarrow> c = (\<Union>i\<in>nat. c \<inter> Vfrom(A,i))" | 
| 13163 | 544 | apply (rule subset_UN_iff_eq [THEN iffD1]) | 
| 545 | apply (erule univ_eq_UN [THEN subst]) | |
| 546 | done | |
| 547 | ||
| 548 | lemma univ_Int_Vfrom_subset: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 549 | "\<lbrakk>a \<subseteq> univ(X); | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 550 | \<And>i. i:nat \<Longrightarrow> a \<inter> Vfrom(X,i) \<subseteq> b\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 551 | \<Longrightarrow> a \<subseteq> b" | 
| 13163 | 552 | apply (subst subset_univ_eq_Int, assumption) | 
| 46820 | 553 | apply (rule UN_least, simp) | 
| 13163 | 554 | done | 
| 555 | ||
| 556 | lemma univ_Int_Vfrom_eq: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 557 | "\<lbrakk>a \<subseteq> univ(X); b \<subseteq> univ(X); | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 558 | \<And>i. i:nat \<Longrightarrow> a \<inter> Vfrom(X,i) = b \<inter> Vfrom(X,i) | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 559 | \<rbrakk> \<Longrightarrow> a = b" | 
| 13163 | 560 | apply (rule equalityI) | 
| 561 | apply (rule univ_Int_Vfrom_subset, assumption) | |
| 46820 | 562 | apply (blast elim: equalityCE) | 
| 13163 | 563 | apply (rule univ_Int_Vfrom_subset, assumption) | 
| 46820 | 564 | apply (blast elim: equalityCE) | 
| 13163 | 565 | done | 
| 566 | ||
| 69593 | 567 | subsection\<open>Closure Properties for \<^term>\<open>univ(A)\<close>\<close> | 
| 13163 | 568 | |
| 13220 | 569 | lemma zero_in_univ: "0 \<in> univ(A)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 570 | unfolding univ_def | 
| 13163 | 571 | apply (rule nat_0I [THEN zero_in_Vfrom]) | 
| 572 | done | |
| 573 | ||
| 46820 | 574 | lemma zero_subset_univ: "{0} \<subseteq> univ(A)"
 | 
| 13255 | 575 | by (blast intro: zero_in_univ) | 
| 576 | ||
| 46820 | 577 | lemma A_subset_univ: "A \<subseteq> univ(A)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 578 | unfolding univ_def | 
| 13163 | 579 | apply (rule A_subset_Vfrom) | 
| 580 | done | |
| 581 | ||
| 45602 | 582 | lemmas A_into_univ = A_subset_univ [THEN subsetD] | 
| 13163 | 583 | |
| 60770 | 584 | subsubsection\<open>Closure under Unordered and Ordered Pairs\<close> | 
| 13163 | 585 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 586 | lemma singleton_in_univ: "a: univ(A) \<Longrightarrow> {a} \<in> univ(A)"
 | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 587 | unfolding univ_def | 
| 13163 | 588 | apply (blast intro: singleton_in_VLimit Limit_nat) | 
| 589 | done | |
| 590 | ||
| 591 | lemma doubleton_in_univ: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 592 |     "\<lbrakk>a: univ(A);  b: univ(A)\<rbrakk> \<Longrightarrow> {a,b} \<in> univ(A)"
 | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 593 | unfolding univ_def | 
| 13163 | 594 | apply (blast intro: doubleton_in_VLimit Limit_nat) | 
| 595 | done | |
| 596 | ||
| 597 | lemma Pair_in_univ: | |
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 598 | "\<lbrakk>a: univ(A); b: univ(A)\<rbrakk> \<Longrightarrow> \<langle>a,b\<rangle> \<in> univ(A)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 599 | unfolding univ_def | 
| 13163 | 600 | apply (blast intro: Pair_in_VLimit Limit_nat) | 
| 601 | done | |
| 602 | ||
| 603 | lemma Union_in_univ: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 604 | "\<lbrakk>X: univ(A); Transset(A)\<rbrakk> \<Longrightarrow> \<Union>(X) \<in> univ(A)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 605 | unfolding univ_def | 
| 13163 | 606 | apply (blast intro: Union_in_VLimit Limit_nat) | 
| 607 | done | |
| 608 | ||
| 46820 | 609 | lemma product_univ: "univ(A)*univ(A) \<subseteq> univ(A)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 610 | unfolding univ_def | 
| 13163 | 611 | apply (rule Limit_nat [THEN product_VLimit]) | 
| 612 | done | |
| 613 | ||
| 614 | ||
| 60770 | 615 | subsubsection\<open>The Natural Numbers\<close> | 
| 13163 | 616 | |
| 46820 | 617 | lemma nat_subset_univ: "nat \<subseteq> univ(A)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 618 | unfolding univ_def | 
| 13163 | 619 | apply (rule i_subset_Vfrom) | 
| 620 | done | |
| 621 | ||
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 622 | lemma nat_into_univ: "n \<in> nat \<Longrightarrow> n \<in> univ(A)" | 
| 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 623 | by (rule nat_subset_univ [THEN subsetD]) | 
| 13163 | 624 | |
| 60770 | 625 | subsubsection\<open>Instances for 1 and 2\<close> | 
| 13163 | 626 | |
| 13220 | 627 | lemma one_in_univ: "1 \<in> univ(A)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 628 | unfolding univ_def | 
| 13163 | 629 | apply (rule Limit_nat [THEN one_in_VLimit]) | 
| 630 | done | |
| 631 | ||
| 60770 | 632 | text\<open>unused!\<close> | 
| 13220 | 633 | lemma two_in_univ: "2 \<in> univ(A)" | 
| 13163 | 634 | by (blast intro: nat_into_univ) | 
| 635 | ||
| 46820 | 636 | lemma bool_subset_univ: "bool \<subseteq> univ(A)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 637 | unfolding bool_def | 
| 13163 | 638 | apply (blast intro!: zero_in_univ one_in_univ) | 
| 639 | done | |
| 640 | ||
| 45602 | 641 | lemmas bool_into_univ = bool_subset_univ [THEN subsetD] | 
| 13163 | 642 | |
| 643 | ||
| 60770 | 644 | subsubsection\<open>Closure under Disjoint Union\<close> | 
| 13163 | 645 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 646 | lemma Inl_in_univ: "a: univ(A) \<Longrightarrow> Inl(a) \<in> univ(A)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 647 | unfolding univ_def | 
| 13163 | 648 | apply (erule Inl_in_VLimit [OF _ Limit_nat]) | 
| 649 | done | |
| 650 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 651 | lemma Inr_in_univ: "b: univ(A) \<Longrightarrow> Inr(b) \<in> univ(A)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 652 | unfolding univ_def | 
| 13163 | 653 | apply (erule Inr_in_VLimit [OF _ Limit_nat]) | 
| 654 | done | |
| 655 | ||
| 46820 | 656 | lemma sum_univ: "univ(C)+univ(C) \<subseteq> univ(C)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 657 | unfolding univ_def | 
| 13163 | 658 | apply (rule Limit_nat [THEN sum_VLimit]) | 
| 659 | done | |
| 660 | ||
| 661 | lemmas sum_subset_univ = subset_trans [OF sum_mono sum_univ] | |
| 662 | ||
| 13255 | 663 | lemma Sigma_subset_univ: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 664 | "\<lbrakk>A \<subseteq> univ(D); \<And>x. x \<in> A \<Longrightarrow> B(x) \<subseteq> univ(D)\<rbrakk> \<Longrightarrow> Sigma(A,B) \<subseteq> univ(D)" | 
| 13255 | 665 | apply (simp add: univ_def) | 
| 46820 | 666 | apply (blast intro: Sigma_subset_VLimit del: subsetI) | 
| 13255 | 667 | done | 
| 13163 | 668 | |
| 13255 | 669 | |
| 670 | (*Closure under binary union -- use Un_least | |
| 671 | Closure under Collect -- use Collect_subset [THEN subset_trans] | |
| 672 | Closure under RepFun -- use RepFun_subset *) | |
| 13163 | 673 | |
| 674 | ||
| 60770 | 675 | subsection\<open>Finite Branching Closure Properties\<close> | 
| 13163 | 676 | |
| 60770 | 677 | subsubsection\<open>Closure under Finite Powerset\<close> | 
| 13163 | 678 | |
| 679 | lemma Fin_Vfrom_lemma: | |
| 76214 | 680 | "\<lbrakk>b: Fin(Vfrom(A,i)); Limit(i)\<rbrakk> \<Longrightarrow> \<exists>j. b \<subseteq> Vfrom(A,j) \<and> j<i" | 
| 13163 | 681 | apply (erule Fin_induct) | 
| 682 | apply (blast dest!: Limit_has_0, safe) | |
| 683 | apply (erule Limit_VfromE, assumption) | |
| 684 | apply (blast intro!: Un_least_lt intro: Vfrom_UnI1 Vfrom_UnI2) | |
| 685 | done | |
| 0 | 686 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 687 | lemma Fin_VLimit: "Limit(i) \<Longrightarrow> Fin(Vfrom(A,i)) \<subseteq> Vfrom(A,i)" | 
| 13163 | 688 | apply (rule subsetI) | 
| 689 | apply (drule Fin_Vfrom_lemma, safe) | |
| 690 | apply (rule Vfrom [THEN ssubst]) | |
| 691 | apply (blast dest!: ltD) | |
| 692 | done | |
| 693 | ||
| 694 | lemmas Fin_subset_VLimit = subset_trans [OF Fin_mono Fin_VLimit] | |
| 695 | ||
| 46820 | 696 | lemma Fin_univ: "Fin(univ(A)) \<subseteq> univ(A)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 697 | unfolding univ_def | 
| 13163 | 698 | apply (rule Limit_nat [THEN Fin_VLimit]) | 
| 699 | done | |
| 700 | ||
| 60770 | 701 | subsubsection\<open>Closure under Finite Powers: Functions from a Natural Number\<close> | 
| 13163 | 702 | |
| 703 | lemma nat_fun_VLimit: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 704 | "\<lbrakk>n: nat; Limit(i)\<rbrakk> \<Longrightarrow> n -> Vfrom(A,i) \<subseteq> Vfrom(A,i)" | 
| 13163 | 705 | apply (erule nat_fun_subset_Fin [THEN subset_trans]) | 
| 706 | apply (blast del: subsetI | |
| 707 | intro: subset_refl Fin_subset_VLimit Sigma_subset_VLimit nat_subset_VLimit) | |
| 708 | done | |
| 709 | ||
| 710 | lemmas nat_fun_subset_VLimit = subset_trans [OF Pi_mono nat_fun_VLimit] | |
| 711 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 712 | lemma nat_fun_univ: "n: nat \<Longrightarrow> n -> univ(A) \<subseteq> univ(A)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 713 | unfolding univ_def | 
| 13163 | 714 | apply (erule nat_fun_VLimit [OF _ Limit_nat]) | 
| 715 | done | |
| 716 | ||
| 717 | ||
| 60770 | 718 | subsubsection\<open>Closure under Finite Function Space\<close> | 
| 13163 | 719 | |
| 60770 | 720 | text\<open>General but seldom-used version; normally the domain is fixed\<close> | 
| 13163 | 721 | lemma FiniteFun_VLimit1: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 722 | "Limit(i) \<Longrightarrow> Vfrom(A,i) -||> Vfrom(A,i) \<subseteq> Vfrom(A,i)" | 
| 13163 | 723 | apply (rule FiniteFun.dom_subset [THEN subset_trans]) | 
| 724 | apply (blast del: subsetI | |
| 725 | intro: Fin_subset_VLimit Sigma_subset_VLimit subset_refl) | |
| 726 | done | |
| 727 | ||
| 46820 | 728 | lemma FiniteFun_univ1: "univ(A) -||> univ(A) \<subseteq> univ(A)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 729 | unfolding univ_def | 
| 13163 | 730 | apply (rule Limit_nat [THEN FiniteFun_VLimit1]) | 
| 731 | done | |
| 732 | ||
| 60770 | 733 | text\<open>Version for a fixed domain\<close> | 
| 13163 | 734 | lemma FiniteFun_VLimit: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 735 | "\<lbrakk>W \<subseteq> Vfrom(A,i); Limit(i)\<rbrakk> \<Longrightarrow> W -||> Vfrom(A,i) \<subseteq> Vfrom(A,i)" | 
| 46820 | 736 | apply (rule subset_trans) | 
| 13163 | 737 | apply (erule FiniteFun_mono [OF _ subset_refl]) | 
| 738 | apply (erule FiniteFun_VLimit1) | |
| 739 | done | |
| 740 | ||
| 741 | lemma FiniteFun_univ: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 742 | "W \<subseteq> univ(A) \<Longrightarrow> W -||> univ(A) \<subseteq> univ(A)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 743 | unfolding univ_def | 
| 13163 | 744 | apply (erule FiniteFun_VLimit [OF _ Limit_nat]) | 
| 745 | done | |
| 746 | ||
| 747 | lemma FiniteFun_in_univ: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 748 | "\<lbrakk>f: W -||> univ(A); W \<subseteq> univ(A)\<rbrakk> \<Longrightarrow> f \<in> univ(A)" | 
| 13163 | 749 | by (erule FiniteFun_univ [THEN subsetD], assumption) | 
| 750 | ||
| 61798 | 751 | text\<open>Remove \<open>\<subseteq>\<close> from the rule above\<close> | 
| 13163 | 752 | lemmas FiniteFun_in_univ' = FiniteFun_in_univ [OF _ subsetI] | 
| 753 | ||
| 754 | ||
| 60770 | 755 | subsection\<open>* For QUniv. Properties of Vfrom analogous to the "take-lemma" *\<close> | 
| 13163 | 756 | |
| 60770 | 757 | text\<open>Intersecting a*b with Vfrom...\<close> | 
| 13163 | 758 | |
| 60770 | 759 | text\<open>This version says a, b exist one level down, in the smaller set Vfrom(X,i)\<close> | 
| 13163 | 760 | lemma doubleton_in_Vfrom_D: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 761 |      "\<lbrakk>{a,b} \<in> Vfrom(X,succ(i));  Transset(X)\<rbrakk>
 | 
| 76214 | 762 | \<Longrightarrow> a \<in> Vfrom(X,i) \<and> b \<in> Vfrom(X,i)" | 
| 46820 | 763 | by (drule Transset_Vfrom_succ [THEN equalityD1, THEN subsetD, THEN PowD], | 
| 13163 | 764 | assumption, fast) | 
| 765 | ||
| 60770 | 766 | text\<open>This weaker version says a, b exist at the same level\<close> | 
| 45602 | 767 | lemmas Vfrom_doubleton_D = Transset_Vfrom [THEN Transset_doubleton_D] | 
| 13163 | 768 | |
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 769 | (** Using only the weaker theorem would prove \<langle>a,b\<rangle> \<in> Vfrom(X,i) | 
| 46821 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
46820diff
changeset | 770 | implies a, b \<in> Vfrom(X,i), which is useless for induction. | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 771 | Using only the stronger theorem would prove \<langle>a,b\<rangle> \<in> Vfrom(X,succ(succ(i))) | 
| 46821 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
46820diff
changeset | 772 | implies a, b \<in> Vfrom(X,i), leaving the succ(i) case untreated. | 
| 13163 | 773 | The combination gives a reduction by precisely one level, which is | 
| 774 | most convenient for proofs. | |
| 775 | **) | |
| 776 | ||
| 777 | lemma Pair_in_Vfrom_D: | |
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 778 | "\<lbrakk>\<langle>a,b\<rangle> \<in> Vfrom(X,succ(i)); Transset(X)\<rbrakk> | 
| 76214 | 779 | \<Longrightarrow> a \<in> Vfrom(X,i) \<and> b \<in> Vfrom(X,i)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 780 | unfolding Pair_def | 
| 13163 | 781 | apply (blast dest!: doubleton_in_Vfrom_D Vfrom_doubleton_D) | 
| 782 | done | |
| 783 | ||
| 784 | lemma product_Int_Vfrom_subset: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 785 | "Transset(X) \<Longrightarrow> | 
| 46820 | 786 | (a*b) \<inter> Vfrom(X, succ(i)) \<subseteq> (a \<inter> Vfrom(X,i)) * (b \<inter> Vfrom(X,i))" | 
| 13163 | 787 | by (blast dest!: Pair_in_Vfrom_D) | 
| 788 | ||
| 789 | ||
| 790 | ML | |
| 60770 | 791 | \<open> | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46821diff
changeset | 792 | val rank_ss = | 
| 69593 | 793 |   simpset_of (\<^context> addsimps [@{thm VsetI}]
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46821diff
changeset | 794 |     addsimps @{thms rank_rls} @ (@{thms rank_rls} RLN (2, [@{thm lt_trans}])));
 | 
| 60770 | 795 | \<close> | 
| 0 | 796 | |
| 797 | end |