New theorems from Constructible, and moving some Isar material from Main
authorpaulson
Fri, 17 May 2002 16:54:25 +0200
changeset 13162660a71e712af
parent 13161 a40db0418145
child 13163 e320a52ff711
New theorems from Constructible, and moving some Isar material from Main
src/ZF/Main.thy
src/ZF/OrdQuant.thy
src/ZF/Ordinal.thy
     1.1 --- a/src/ZF/Main.thy	Fri May 17 16:48:11 2002 +0200
     1.2 +++ b/src/ZF/Main.thy	Fri May 17 16:54:25 2002 +0200
     1.3 @@ -15,21 +15,51 @@
     1.4    and wf_on_induct = wf_on_induct [consumes 2, induct set: wf_on]
     1.5    and wf_on_induct_rule = wf_on_induct [rule_format, consumes 2, induct set: wf_on]
     1.6  
     1.7 -(* belongs to theory Ordinal *)
     1.8 -declare Ord_Least [intro,simp,TC]
     1.9 -lemmas Ord_induct = Ord_induct [consumes 2]
    1.10 -  and Ord_induct_rule = Ord_induct [rule_format, consumes 2]
    1.11 -  and trans_induct = trans_induct [consumes 1]
    1.12 -  and trans_induct_rule = trans_induct [rule_format, consumes 1]
    1.13 -  and trans_induct3 = trans_induct3 [case_names 0 succ limit, consumes 1]
    1.14 -  and trans_induct3_rule = trans_induct3 [rule_format, case_names 0 succ limit, consumes 1]
    1.15 -
    1.16  (* belongs to theory Nat *)
    1.17  lemmas nat_induct = nat_induct [case_names 0 succ, induct set: nat]
    1.18    and complete_induct = complete_induct [case_names less, consumes 1]
    1.19    and complete_induct_rule = complete_induct [rule_format, case_names less, consumes 1]
    1.20    and diff_induct = diff_induct [case_names 0 0_succ succ_succ, consumes 2]
    1.21  
    1.22 +
    1.23 +
    1.24 +subsection{* Iteration of the function @{term F} *}
    1.25 +
    1.26 +consts  iterates :: "[i=>i,i,i] => i"   ("(_^_ '(_'))" [60,1000,1000] 60)
    1.27 +
    1.28 +primrec
    1.29 +    "F^0 (x) = x"
    1.30 +    "F^(succ(n)) (x) = F(F^n (x))"
    1.31 +
    1.32 +constdefs
    1.33 +  iterates_omega :: "[i=>i,i] => i"
    1.34 +    "iterates_omega(F,x) == \<Union>n\<in>nat. F^n (x)"
    1.35 +
    1.36 +syntax (xsymbols)
    1.37 +  iterates_omega :: "[i=>i,i] => i"   ("(_^\<omega> '(_'))" [60,1000] 60)
    1.38 +
    1.39 +lemma iterates_triv:
    1.40 +     "[| n\<in>nat;  F(x) = x |] ==> F^n (x) = x"  
    1.41 +by (induct n rule: nat_induct, simp_all)
    1.42 +
    1.43 +lemma iterates_type [TC]:
    1.44 +     "[| n:nat;  a: A; !!x. x:A ==> F(x) : A |] 
    1.45 +      ==> F^n (a) : A"  
    1.46 +by (induct n rule: nat_induct, simp_all)
    1.47 +
    1.48 +lemma iterates_omega_triv:
    1.49 +    "F(x) = x ==> F^\<omega> (x) = x" 
    1.50 +by (simp add: iterates_omega_def iterates_triv) 
    1.51 +
    1.52 +lemma Ord_iterates [simp]:
    1.53 +     "[| n\<in>nat;  !!i. Ord(i) ==> Ord(F(i));  Ord(x) |] 
    1.54 +      ==> Ord(F^n (x))"  
    1.55 +by (induct n rule: nat_induct, simp_all)
    1.56 +
    1.57 +
    1.58 +(* belongs to theory Cardinal *)
    1.59 +declare Ord_Least [intro,simp,TC]
    1.60 +
    1.61  (* belongs to theory Epsilon *)
    1.62  lemmas eclose_induct = eclose_induct [induct set: eclose]
    1.63    and eclose_induct_down = eclose_induct_down [consumes 1]
    1.64 @@ -59,7 +89,7 @@
    1.65  
    1.66  (* belongs to theory CardinalArith *)
    1.67  
    1.68 -lemma InfCard_square_eqpoll: "InfCard(K) \<Longrightarrow> K \<times> K \<approx> K"
    1.69 +lemma InfCard_square_eqpoll: "InfCard(K) ==> K \<times> K \<approx> K"
    1.70  apply (rule well_ord_InfCard_square_eq)  
    1.71   apply (erule InfCard_is_Card [THEN Card_is_Ord, THEN well_ord_Memrel]) 
    1.72  apply (simp add: InfCard_is_Card [THEN Card_cardinal_eq]) 
     2.1 --- a/src/ZF/OrdQuant.thy	Fri May 17 16:48:11 2002 +0200
     2.2 +++ b/src/ZF/OrdQuant.thy	Fri May 17 16:54:25 2002 +0200
     2.3 @@ -73,11 +73,11 @@
     2.4  lemma trans_imp_trans_on: "trans(r) ==> trans[A](r)"
     2.5  by (unfold trans_def trans_on_def, blast)
     2.6  
     2.7 -lemma image_is_UN: "\<lbrakk>function(g); x <= domain(g)\<rbrakk> \<Longrightarrow> g``x = (UN k:x. {g`k})"
     2.8 +lemma image_is_UN: "[| function(g); x <= domain(g) |] ==> g``x = (UN k:x. {g`k})"
     2.9  by (blast intro: function_apply_equality [THEN sym] function_apply_Pair) 
    2.10  
    2.11  lemma functionI: 
    2.12 -     "\<lbrakk>!!x y y'. \<lbrakk><x,y>:r; <x,y'>:r\<rbrakk> \<Longrightarrow> y=y'\<rbrakk> \<Longrightarrow> function(r)"
    2.13 +     "[| !!x y y'. [| <x,y>:r; <x,y'>:r |] ==> y=y' |] ==> function(r)"
    2.14  by (simp add: function_def, blast) 
    2.15  
    2.16  lemma function_lam: "function (lam x:A. b(x))"
    2.17 @@ -92,7 +92,7 @@
    2.18  (** These mostly belong to theory Ordinal **)
    2.19  
    2.20  lemma Union_upper_le:
    2.21 -     "\<lbrakk>j: J;  i\<le>j;  Ord(\<Union>(J))\<rbrakk> \<Longrightarrow> i \<le> \<Union>J"
    2.22 +     "[| j: J;  i\<le>j;  Ord(\<Union>(J)) |] ==> i \<le> \<Union>J"
    2.23  apply (subst Union_eq_UN)  
    2.24  apply (rule UN_upper_le, auto)
    2.25  done
    2.26 @@ -100,29 +100,29 @@
    2.27  lemma zero_not_Limit [iff]: "~ Limit(0)"
    2.28  by (simp add: Limit_def)
    2.29  
    2.30 -lemma Limit_has_1: "Limit(i) \<Longrightarrow> 1 < i"
    2.31 +lemma Limit_has_1: "Limit(i) ==> 1 < i"
    2.32  by (blast intro: Limit_has_0 Limit_has_succ)
    2.33  
    2.34 -lemma Limit_Union [rule_format]: "\<lbrakk>I \<noteq> 0;  \<forall>i\<in>I. Limit(i)\<rbrakk> \<Longrightarrow> Limit(\<Union>I)"
    2.35 +lemma Limit_Union [rule_format]: "[| I \<noteq> 0;  \<forall>i\<in>I. Limit(i) |] ==> Limit(\<Union>I)"
    2.36  apply (simp add: Limit_def lt_def)
    2.37  apply (blast intro!: equalityI)
    2.38  done
    2.39  
    2.40 -lemma increasing_LimitI: "\<lbrakk>0<l; \<forall>x\<in>l. \<exists>y\<in>l. x<y\<rbrakk> \<Longrightarrow> Limit(l)"
    2.41 +lemma increasing_LimitI: "[| 0<l; \<forall>x\<in>l. \<exists>y\<in>l. x<y |] ==> Limit(l)"
    2.42  apply (simp add: Limit_def lt_Ord2, clarify)
    2.43  apply (drule_tac i=y in ltD) 
    2.44  apply (blast intro: lt_trans1 [OF _ ltI] lt_Ord2)
    2.45  done
    2.46  
    2.47  lemma UN_upper_lt:
    2.48 -     "\<lbrakk>a\<in>A;  i < b(a);  Ord(\<Union>x\<in>A. b(x))\<rbrakk> \<Longrightarrow> i < (\<Union>x\<in>A. b(x))"
    2.49 +     "[| a\<in>A;  i < b(a);  Ord(\<Union>x\<in>A. b(x)) |] ==> i < (\<Union>x\<in>A. b(x))"
    2.50  by (unfold lt_def, blast) 
    2.51  
    2.52 -lemma lt_imp_0_lt: "j<i \<Longrightarrow> 0<i"
    2.53 +lemma lt_imp_0_lt: "j<i ==> 0<i"
    2.54  by (blast intro: lt_trans1 Ord_0_le [OF lt_Ord]) 
    2.55  
    2.56  lemma Ord_set_cases:
    2.57 -   "\<forall>i\<in>I. Ord(i) \<Longrightarrow> I=0 \<or> \<Union>(I) \<in> I \<or> (\<Union>(I) \<notin> I \<and> Limit(\<Union>(I)))"
    2.58 +   "\<forall>i\<in>I. Ord(i) ==> I=0 \<or> \<Union>(I) \<in> I \<or> (\<Union>(I) \<notin> I \<and> Limit(\<Union>(I)))"
    2.59  apply (clarify elim!: not_emptyE) 
    2.60  apply (cases "\<Union>(I)" rule: Ord_cases) 
    2.61     apply (blast intro: Ord_Union)
    2.62 @@ -140,10 +140,10 @@
    2.63  by (drule Ord_set_cases, auto)
    2.64  
    2.65  (*See also Transset_iff_Union_succ*)
    2.66 -lemma Ord_Union_succ_eq: "Ord(i) \<Longrightarrow> \<Union>(succ(i)) = i"
    2.67 +lemma Ord_Union_succ_eq: "Ord(i) ==> \<Union>(succ(i)) = i"
    2.68  by (blast intro: Ord_trans)
    2.69  
    2.70 -lemma lt_Union_iff: "\<forall>i\<in>A. Ord(i) \<Longrightarrow> (j < \<Union>(A)) <-> (\<exists>i\<in>A. j<i)"
    2.71 +lemma lt_Union_iff: "\<forall>i\<in>A. Ord(i) ==> (j < \<Union>(A)) <-> (\<exists>i\<in>A. j<i)"
    2.72  by (auto simp: lt_def Ord_Union)
    2.73  
    2.74  lemma Un_upper1_lt: "[|k < i; Ord(j)|] ==> k < i Un j"
    2.75 @@ -153,15 +153,15 @@
    2.76  by (simp add: lt_def) 
    2.77  
    2.78  lemma Ord_OUN [intro,simp]:
    2.79 -     "\<lbrakk>!!x. x<A \<Longrightarrow> Ord(B(x))\<rbrakk> \<Longrightarrow> Ord(\<Union>x<A. B(x))"
    2.80 +     "[| !!x. x<A ==> Ord(B(x)) |] ==> Ord(\<Union>x<A. B(x))"
    2.81  by (simp add: OUnion_def ltI Ord_UN) 
    2.82  
    2.83  lemma OUN_upper_lt:
    2.84 -     "\<lbrakk>a<A;  i < b(a);  Ord(\<Union>x<A. b(x))\<rbrakk> \<Longrightarrow> i < (\<Union>x<A. b(x))"
    2.85 +     "[| a<A;  i < b(a);  Ord(\<Union>x<A. b(x)) |] ==> i < (\<Union>x<A. b(x))"
    2.86  by (unfold OUnion_def lt_def, blast )
    2.87  
    2.88  lemma OUN_upper_le:
    2.89 -     "\<lbrakk>a<A;  i\<le>b(a);  Ord(\<Union>x<A. b(x))\<rbrakk> \<Longrightarrow> i \<le> (\<Union>x<A. b(x))"
    2.90 +     "[| a<A;  i\<le>b(a);  Ord(\<Union>x<A. b(x)) |] ==> i \<le> (\<Union>x<A. b(x))"
    2.91  apply (unfold OUnion_def, auto)
    2.92  apply (rule UN_upper_le )
    2.93  apply (auto simp add: lt_def) 
     3.1 --- a/src/ZF/Ordinal.thy	Fri May 17 16:48:11 2002 +0200
     3.2 +++ b/src/ZF/Ordinal.thy	Fri May 17 16:54:25 2002 +0200
     3.3 @@ -456,6 +456,17 @@
     3.4  apply (blast intro: elim: ltE) +
     3.5  done
     3.6  
     3.7 +lemma succ_lt_iff: "succ(i) < j \<longleftrightarrow> i<j & succ(i) \<noteq> j"
     3.8 +apply auto 
     3.9 +apply (blast intro: lt_trans le_refl dest: lt_Ord) 
    3.10 +apply (frule lt_Ord) 
    3.11 +apply (rule not_le_iff_lt [THEN iffD1]) 
    3.12 +  apply (blast intro: lt_Ord2)
    3.13 + apply blast  
    3.14 +apply (simp add: lt_Ord lt_Ord2 le_iff) 
    3.15 +apply (blast dest: lt_asym) 
    3.16 +done
    3.17 +
    3.18  (** Union and Intersection **)
    3.19  
    3.20  lemma Un_upper1_le: "[| Ord(i); Ord(j) |] ==> i le i Un j"
    3.21 @@ -488,6 +499,26 @@
    3.22  apply (auto simp add: Int_commute le_subset_iff subset_Int_iff lt_Ord) 
    3.23  done
    3.24  
    3.25 +lemma Ord_Un_if:
    3.26 +     "[| Ord(i); Ord(j) |] ==> i \<union> j = (if j<i then i else j)"
    3.27 +by (simp add: not_lt_iff_le le_imp_subset leI
    3.28 +              subset_Un_iff [symmetric]  subset_Un_iff2 [symmetric]) 
    3.29 +
    3.30 +lemma succ_Un_distrib:
    3.31 +     "[| Ord(i); Ord(j) |] ==> succ(i \<union> j) = succ(i) \<union> succ(j)"
    3.32 +by (simp add: Ord_Un_if lt_Ord le_Ord2) 
    3.33 +
    3.34 +lemma lt_Un_iff:
    3.35 +     "[| Ord(i); Ord(j) |] ==> k < i \<union> j <-> k < i | k < j";
    3.36 +apply (simp add: Ord_Un_if not_lt_iff_le) 
    3.37 +apply (blast intro: leI lt_trans2)+ 
    3.38 +done
    3.39 +
    3.40 +lemma le_Un_iff:
    3.41 +     "[| Ord(i); Ord(j) |] ==> k \<le> i \<union> j <-> k \<le> i | k \<le> j";
    3.42 +by (simp add: succ_Un_distrib lt_Un_iff [symmetric]) 
    3.43 +
    3.44 +
    3.45  (*FIXME: the Intersection duals are missing!*)
    3.46  
    3.47  (*** Results about limits ***)
    3.48 @@ -600,6 +631,14 @@
    3.49  apply (erule Ord_cases, blast+)
    3.50  done
    3.51  
    3.52 +(*special induction rules for the "induct" method*)
    3.53 +lemmas Ord_induct = Ord_induct [consumes 2]
    3.54 +  and Ord_induct_rule = Ord_induct [rule_format, consumes 2]
    3.55 +  and trans_induct = trans_induct [consumes 1]
    3.56 +  and trans_induct_rule = trans_induct [rule_format, consumes 1]
    3.57 +  and trans_induct3 = trans_induct3 [case_names 0 succ limit, consumes 1]
    3.58 +  and trans_induct3_rule = trans_induct3 [rule_format, case_names 0 succ limit, consumes 1]
    3.59 +
    3.60  ML 
    3.61  {*
    3.62  val Memrel_def = thm "Memrel_def";