# HG changeset patch # User paulson # Date 1021647265 -7200 # Node ID 660a71e712afa7c83e0bd6d1aee69a2dd811e341 # Parent a40db0418145179300a853227615348a31562280 New theorems from Constructible, and moving some Isar material from Main diff -r a40db0418145 -r 660a71e712af src/ZF/Main.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) == \n\nat. F^n (x)" + +syntax (xsymbols) + iterates_omega :: "[i=>i,i] => i" ("(_^\ '(_'))" [60,1000] 60) + +lemma iterates_triv: + "[| n\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^\ (x) = x" +by (simp add: iterates_omega_def iterates_triv) + +lemma Ord_iterates [simp]: + "[| n\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) \ K \ K \ K" +lemma InfCard_square_eqpoll: "InfCard(K) ==> K \ K \ 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]) diff -r a40db0418145 -r 660a71e712af src/ZF/OrdQuant.thy --- 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: "\function(g); x <= domain(g)\ \ 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: - "\!!x y y'. \:r; :r\ \ y=y'\ \ function(r)" + "[| !!x y y'. [| :r; :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: - "\j: J; i\j; Ord(\(J))\ \ i \ \J" + "[| j: J; i\j; Ord(\(J)) |] ==> i \ \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) \ 1 < i" +lemma Limit_has_1: "Limit(i) ==> 1 < i" by (blast intro: Limit_has_0 Limit_has_succ) -lemma Limit_Union [rule_format]: "\I \ 0; \i\I. Limit(i)\ \ Limit(\I)" +lemma Limit_Union [rule_format]: "[| I \ 0; \i\I. Limit(i) |] ==> Limit(\I)" apply (simp add: Limit_def lt_def) apply (blast intro!: equalityI) done -lemma increasing_LimitI: "\0x\l. \y\l. x \ Limit(l)" +lemma increasing_LimitI: "[| 0x\l. \y\l. x 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: - "\a\A; i < b(a); Ord(\x\A. b(x))\ \ i < (\x\A. b(x))" + "[| a\A; i < b(a); Ord(\x\A. b(x)) |] ==> i < (\x\A. b(x))" by (unfold lt_def, blast) -lemma lt_imp_0_lt: "j 0 0i\I. Ord(i) \ I=0 \ \(I) \ I \ (\(I) \ I \ Limit(\(I)))" + "\i\I. Ord(i) ==> I=0 \ \(I) \ I \ (\(I) \ I \ Limit(\(I)))" apply (clarify elim!: not_emptyE) apply (cases "\(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) \ \(succ(i)) = i" +lemma Ord_Union_succ_eq: "Ord(i) ==> \(succ(i)) = i" by (blast intro: Ord_trans) -lemma lt_Union_iff: "\i\A. Ord(i) \ (j < \(A)) <-> (\i\A. ji\A. Ord(i) ==> (j < \(A)) <-> (\i\A. j k < i Un j" @@ -153,15 +153,15 @@ by (simp add: lt_def) lemma Ord_OUN [intro,simp]: - "\!!x. x Ord(B(x))\ \ Ord(\x Ord(B(x)) |] ==> Ord(\xax \ i < (\xx i < (\xab(a); Ord(\x \ i \ (\xb(a); Ord(\x i \ (\x i 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 \ j = (if j succ(i \ j) = succ(i) \ succ(j)" +by (simp add: Ord_Un_if lt_Ord le_Ord2) + +lemma lt_Un_iff: + "[| Ord(i); Ord(j) |] ==> k < i \ 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 \ i \ j <-> k \ i | k \ 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";