New theorems from Constructible, and moving some Isar material from Main
authorpaulson
Fri, 17 May 2002 16:54:25 +0200
changeset 13162 660a71e712af
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
--- a/src/ZF/Main.thy	Fri May 17 16:48:11 2002 +0200
+++ b/src/ZF/Main.thy	Fri May 17 16:54:25 2002 +0200
@@ -15,21 +15,51 @@
   and wf_on_induct = wf_on_induct [consumes 2, induct set: wf_on]
   and wf_on_induct_rule = wf_on_induct [rule_format, consumes 2, induct set: wf_on]
 
-(* belongs to theory Ordinal *)
-declare Ord_Least [intro,simp,TC]
-lemmas Ord_induct = Ord_induct [consumes 2]
-  and Ord_induct_rule = Ord_induct [rule_format, consumes 2]
-  and trans_induct = trans_induct [consumes 1]
-  and trans_induct_rule = trans_induct [rule_format, consumes 1]
-  and trans_induct3 = trans_induct3 [case_names 0 succ limit, consumes 1]
-  and trans_induct3_rule = trans_induct3 [rule_format, case_names 0 succ limit, consumes 1]
-
 (* belongs to theory Nat *)
 lemmas nat_induct = nat_induct [case_names 0 succ, induct set: nat]
   and complete_induct = complete_induct [case_names less, consumes 1]
   and complete_induct_rule = complete_induct [rule_format, case_names less, consumes 1]
   and diff_induct = diff_induct [case_names 0 0_succ succ_succ, consumes 2]
 
+
+
+subsection{* Iteration of the function @{term F} *}
+
+consts  iterates :: "[i=>i,i,i] => i"   ("(_^_ '(_'))" [60,1000,1000] 60)
+
+primrec
+    "F^0 (x) = x"
+    "F^(succ(n)) (x) = F(F^n (x))"
+
+constdefs
+  iterates_omega :: "[i=>i,i] => i"
+    "iterates_omega(F,x) == \<Union>n\<in>nat. F^n (x)"
+
+syntax (xsymbols)
+  iterates_omega :: "[i=>i,i] => i"   ("(_^\<omega> '(_'))" [60,1000] 60)
+
+lemma iterates_triv:
+     "[| n\<in>nat;  F(x) = x |] ==> F^n (x) = x"  
+by (induct n rule: nat_induct, simp_all)
+
+lemma iterates_type [TC]:
+     "[| n:nat;  a: A; !!x. x:A ==> F(x) : A |] 
+      ==> F^n (a) : A"  
+by (induct n rule: nat_induct, simp_all)
+
+lemma iterates_omega_triv:
+    "F(x) = x ==> F^\<omega> (x) = x" 
+by (simp add: iterates_omega_def iterates_triv) 
+
+lemma Ord_iterates [simp]:
+     "[| n\<in>nat;  !!i. Ord(i) ==> Ord(F(i));  Ord(x) |] 
+      ==> Ord(F^n (x))"  
+by (induct n rule: nat_induct, simp_all)
+
+
+(* belongs to theory Cardinal *)
+declare Ord_Least [intro,simp,TC]
+
 (* belongs to theory Epsilon *)
 lemmas eclose_induct = eclose_induct [induct set: eclose]
   and eclose_induct_down = eclose_induct_down [consumes 1]
@@ -59,7 +89,7 @@
 
 (* belongs to theory CardinalArith *)
 
-lemma InfCard_square_eqpoll: "InfCard(K) \<Longrightarrow> K \<times> K \<approx> K"
+lemma InfCard_square_eqpoll: "InfCard(K) ==> K \<times> K \<approx> K"
 apply (rule well_ord_InfCard_square_eq)  
  apply (erule InfCard_is_Card [THEN Card_is_Ord, THEN well_ord_Memrel]) 
 apply (simp add: InfCard_is_Card [THEN Card_cardinal_eq]) 
--- a/src/ZF/OrdQuant.thy	Fri May 17 16:48:11 2002 +0200
+++ b/src/ZF/OrdQuant.thy	Fri May 17 16:54:25 2002 +0200
@@ -73,11 +73,11 @@
 lemma trans_imp_trans_on: "trans(r) ==> trans[A](r)"
 by (unfold trans_def trans_on_def, blast)
 
-lemma image_is_UN: "\<lbrakk>function(g); x <= domain(g)\<rbrakk> \<Longrightarrow> g``x = (UN k:x. {g`k})"
+lemma image_is_UN: "[| function(g); x <= domain(g) |] ==> g``x = (UN k:x. {g`k})"
 by (blast intro: function_apply_equality [THEN sym] function_apply_Pair) 
 
 lemma functionI: 
-     "\<lbrakk>!!x y y'. \<lbrakk><x,y>:r; <x,y'>:r\<rbrakk> \<Longrightarrow> y=y'\<rbrakk> \<Longrightarrow> function(r)"
+     "[| !!x y y'. [| <x,y>:r; <x,y'>:r |] ==> y=y' |] ==> function(r)"
 by (simp add: function_def, blast) 
 
 lemma function_lam: "function (lam x:A. b(x))"
@@ -92,7 +92,7 @@
 (** These mostly belong to theory Ordinal **)
 
 lemma Union_upper_le:
-     "\<lbrakk>j: J;  i\<le>j;  Ord(\<Union>(J))\<rbrakk> \<Longrightarrow> i \<le> \<Union>J"
+     "[| j: J;  i\<le>j;  Ord(\<Union>(J)) |] ==> i \<le> \<Union>J"
 apply (subst Union_eq_UN)  
 apply (rule UN_upper_le, auto)
 done
@@ -100,29 +100,29 @@
 lemma zero_not_Limit [iff]: "~ Limit(0)"
 by (simp add: Limit_def)
 
-lemma Limit_has_1: "Limit(i) \<Longrightarrow> 1 < i"
+lemma Limit_has_1: "Limit(i) ==> 1 < i"
 by (blast intro: Limit_has_0 Limit_has_succ)
 
-lemma Limit_Union [rule_format]: "\<lbrakk>I \<noteq> 0;  \<forall>i\<in>I. Limit(i)\<rbrakk> \<Longrightarrow> Limit(\<Union>I)"
+lemma Limit_Union [rule_format]: "[| I \<noteq> 0;  \<forall>i\<in>I. Limit(i) |] ==> Limit(\<Union>I)"
 apply (simp add: Limit_def lt_def)
 apply (blast intro!: equalityI)
 done
 
-lemma increasing_LimitI: "\<lbrakk>0<l; \<forall>x\<in>l. \<exists>y\<in>l. x<y\<rbrakk> \<Longrightarrow> Limit(l)"
+lemma increasing_LimitI: "[| 0<l; \<forall>x\<in>l. \<exists>y\<in>l. x<y |] ==> Limit(l)"
 apply (simp add: Limit_def lt_Ord2, clarify)
 apply (drule_tac i=y in ltD) 
 apply (blast intro: lt_trans1 [OF _ ltI] lt_Ord2)
 done
 
 lemma UN_upper_lt:
-     "\<lbrakk>a\<in>A;  i < b(a);  Ord(\<Union>x\<in>A. b(x))\<rbrakk> \<Longrightarrow> i < (\<Union>x\<in>A. b(x))"
+     "[| a\<in>A;  i < b(a);  Ord(\<Union>x\<in>A. b(x)) |] ==> i < (\<Union>x\<in>A. b(x))"
 by (unfold lt_def, blast) 
 
-lemma lt_imp_0_lt: "j<i \<Longrightarrow> 0<i"
+lemma lt_imp_0_lt: "j<i ==> 0<i"
 by (blast intro: lt_trans1 Ord_0_le [OF lt_Ord]) 
 
 lemma Ord_set_cases:
-   "\<forall>i\<in>I. Ord(i) \<Longrightarrow> I=0 \<or> \<Union>(I) \<in> I \<or> (\<Union>(I) \<notin> I \<and> Limit(\<Union>(I)))"
+   "\<forall>i\<in>I. Ord(i) ==> I=0 \<or> \<Union>(I) \<in> I \<or> (\<Union>(I) \<notin> I \<and> Limit(\<Union>(I)))"
 apply (clarify elim!: not_emptyE) 
 apply (cases "\<Union>(I)" rule: Ord_cases) 
    apply (blast intro: Ord_Union)
@@ -140,10 +140,10 @@
 by (drule Ord_set_cases, auto)
 
 (*See also Transset_iff_Union_succ*)
-lemma Ord_Union_succ_eq: "Ord(i) \<Longrightarrow> \<Union>(succ(i)) = i"
+lemma Ord_Union_succ_eq: "Ord(i) ==> \<Union>(succ(i)) = i"
 by (blast intro: Ord_trans)
 
-lemma lt_Union_iff: "\<forall>i\<in>A. Ord(i) \<Longrightarrow> (j < \<Union>(A)) <-> (\<exists>i\<in>A. j<i)"
+lemma lt_Union_iff: "\<forall>i\<in>A. Ord(i) ==> (j < \<Union>(A)) <-> (\<exists>i\<in>A. j<i)"
 by (auto simp: lt_def Ord_Union)
 
 lemma Un_upper1_lt: "[|k < i; Ord(j)|] ==> k < i Un j"
@@ -153,15 +153,15 @@
 by (simp add: lt_def) 
 
 lemma Ord_OUN [intro,simp]:
-     "\<lbrakk>!!x. x<A \<Longrightarrow> Ord(B(x))\<rbrakk> \<Longrightarrow> Ord(\<Union>x<A. B(x))"
+     "[| !!x. x<A ==> Ord(B(x)) |] ==> Ord(\<Union>x<A. B(x))"
 by (simp add: OUnion_def ltI Ord_UN) 
 
 lemma OUN_upper_lt:
-     "\<lbrakk>a<A;  i < b(a);  Ord(\<Union>x<A. b(x))\<rbrakk> \<Longrightarrow> i < (\<Union>x<A. b(x))"
+     "[| a<A;  i < b(a);  Ord(\<Union>x<A. b(x)) |] ==> i < (\<Union>x<A. b(x))"
 by (unfold OUnion_def lt_def, blast )
 
 lemma OUN_upper_le:
-     "\<lbrakk>a<A;  i\<le>b(a);  Ord(\<Union>x<A. b(x))\<rbrakk> \<Longrightarrow> i \<le> (\<Union>x<A. b(x))"
+     "[| a<A;  i\<le>b(a);  Ord(\<Union>x<A. b(x)) |] ==> i \<le> (\<Union>x<A. b(x))"
 apply (unfold OUnion_def, auto)
 apply (rule UN_upper_le )
 apply (auto simp add: lt_def) 
--- a/src/ZF/Ordinal.thy	Fri May 17 16:48:11 2002 +0200
+++ b/src/ZF/Ordinal.thy	Fri May 17 16:54:25 2002 +0200
@@ -456,6 +456,17 @@
 apply (blast intro: elim: ltE) +
 done
 
+lemma succ_lt_iff: "succ(i) < j \<longleftrightarrow> i<j & succ(i) \<noteq> j"
+apply auto 
+apply (blast intro: lt_trans le_refl dest: lt_Ord) 
+apply (frule lt_Ord) 
+apply (rule not_le_iff_lt [THEN iffD1]) 
+  apply (blast intro: lt_Ord2)
+ apply blast  
+apply (simp add: lt_Ord lt_Ord2 le_iff) 
+apply (blast dest: lt_asym) 
+done
+
 (** Union and Intersection **)
 
 lemma Un_upper1_le: "[| Ord(i); Ord(j) |] ==> i le i Un j"
@@ -488,6 +499,26 @@
 apply (auto simp add: Int_commute le_subset_iff subset_Int_iff lt_Ord) 
 done
 
+lemma Ord_Un_if:
+     "[| Ord(i); Ord(j) |] ==> i \<union> j = (if j<i then i else j)"
+by (simp add: not_lt_iff_le le_imp_subset leI
+              subset_Un_iff [symmetric]  subset_Un_iff2 [symmetric]) 
+
+lemma succ_Un_distrib:
+     "[| Ord(i); Ord(j) |] ==> succ(i \<union> j) = succ(i) \<union> succ(j)"
+by (simp add: Ord_Un_if lt_Ord le_Ord2) 
+
+lemma lt_Un_iff:
+     "[| Ord(i); Ord(j) |] ==> k < i \<union> j <-> k < i | k < j";
+apply (simp add: Ord_Un_if not_lt_iff_le) 
+apply (blast intro: leI lt_trans2)+ 
+done
+
+lemma le_Un_iff:
+     "[| Ord(i); Ord(j) |] ==> k \<le> i \<union> j <-> k \<le> i | k \<le> j";
+by (simp add: succ_Un_distrib lt_Un_iff [symmetric]) 
+
+
 (*FIXME: the Intersection duals are missing!*)
 
 (*** Results about limits ***)
@@ -600,6 +631,14 @@
 apply (erule Ord_cases, blast+)
 done
 
+(*special induction rules for the "induct" method*)
+lemmas Ord_induct = Ord_induct [consumes 2]
+  and Ord_induct_rule = Ord_induct [rule_format, consumes 2]
+  and trans_induct = trans_induct [consumes 1]
+  and trans_induct_rule = trans_induct [rule_format, consumes 1]
+  and trans_induct3 = trans_induct3 [case_names 0 succ limit, consumes 1]
+  and trans_induct3_rule = trans_induct3 [rule_format, case_names 0 succ limit, consumes 1]
+
 ML 
 {*
 val Memrel_def = thm "Memrel_def";