# HG changeset patch # User paulson # Date 1022083917 -7200 # Node ID 03a5afa7b888aa17fdce9443e9181c9152a4f3f2 # Parent 3208b614dc716ea09f63ac42906659775abf2c81 tidying up diff -r 3208b614dc71 -r 03a5afa7b888 src/ZF/OrdQuant.thy --- a/src/ZF/OrdQuant.thy Wed May 22 17:26:34 2002 +0200 +++ b/src/ZF/OrdQuant.thy Wed May 22 18:11:57 2002 +0200 @@ -58,100 +58,16 @@ apply (blast intro: lt_Ord2) done -declare Ord_Un [intro,simp,TC] -declare Ord_UN [intro,simp,TC] -declare Ord_Union [intro,simp,TC] - (** Now some very basic ZF theorems **) +(*FIXME: move to ZF.thy or even to FOL.thy??*) lemma [simp]: "((P-->Q) <-> (P-->R)) <-> (P --> (Q<->R))" by blast -lemma [simp]: "cons(a,cons(a,B)) = cons(a,B)" -by blast - +(*FIXME: move to Rel.thy*) 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})" -by (blast intro: function_apply_equality [THEN sym] function_apply_Pair) - -lemma functionI: - "[| !!x y y'. [| :r; :r |] ==> y=y' |] ==> function(r)" -by (simp add: function_def, blast) - -lemma function_lam: "function (lam x:A. b(x))" -by (simp add: function_def lam_def) - -lemma relation_lam: "relation (lam x:A. b(x))" -by (simp add: relation_def lam_def) - -lemma restrict_iff: "z \ restrict(r,A) \ z \ r & (\x\A. \y. z = \x, y\)" -by (simp add: restrict_def) - -(** These mostly belong to theory Ordinal **) - -lemma Union_upper_le: - "[| j: J; i\j; Ord(\(J)) |] ==> i \ \J" -apply (subst Union_eq_UN) -apply (rule UN_upper_le, auto) -done - -lemma zero_not_Limit [iff]: "~ Limit(0)" -by (simp add: Limit_def) - -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)" -apply (simp add: Limit_def lt_def) -apply (blast intro!: equalityI) -done - -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))" -by (unfold lt_def, blast) - -lemma lt_imp_0_lt: "j 0i\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) - apply (blast intro: subst_elem) - apply auto -apply (clarify elim!: equalityE succ_subsetE) -apply (simp add: Union_subset_iff) -apply (subgoal_tac "B = succ(j)", blast ) -apply (rule le_anti_sym) - apply (simp add: le_subset_iff) -apply (simp add: ltI) -done - -lemma Ord_Union_eq_succD: "[|\x\X. Ord(x); \X = succ(j)|] ==> succ(j) \ X" -by (drule Ord_set_cases, auto) - -(*See also Transset_iff_Union_succ*) -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. j k < i Un j" -by (simp add: lt_def) - -lemma Un_upper2_lt: "[|k < j; Ord(i)|] ==> k < i Un j" -by (simp add: lt_def) - lemma Ord_OUN [intro,simp]: "[| !!x. x Ord(B(x)) |] ==> Ord(\x P(x) |] ==> P(i)" apply (simp add: lt_def oall_def) diff -r 3208b614dc71 -r 03a5afa7b888 src/ZF/Ordinal.thy --- a/src/ZF/Ordinal.thy Wed May 22 17:26:34 2002 +0200 +++ b/src/ZF/Ordinal.thy Wed May 22 18:11:57 2002 +0200 @@ -147,7 +147,7 @@ lemma Ord_succ_iff [iff]: "Ord(succ(i)) <-> Ord(i)" by (blast intro: Ord_succ dest!: Ord_succD) -lemma Ord_Un [TC]: "[| Ord(i); Ord(j) |] ==> Ord(i Un j)" +lemma Ord_Un [intro,simp,TC]: "[| Ord(i); Ord(j) |] ==> Ord(i Un j)" apply (unfold Ord_def) apply (blast intro!: Transset_Un) done @@ -456,6 +456,9 @@ apply (blast intro: elim: ltE) + done +lemma lt_imp_0_lt: "j 0 i j" apply auto apply (blast intro: lt_trans le_refl dest: lt_Ord) @@ -518,17 +521,28 @@ "[| Ord(i); Ord(j) |] ==> k \ i \ j <-> k \ i | k \ j"; by (simp add: succ_Un_distrib lt_Un_iff [symmetric]) +lemma Un_upper1_lt: "[|k < i; Ord(j)|] ==> k < i Un j" +by (simp add: lt_Un_iff lt_Ord2) + +lemma Un_upper2_lt: "[|k < j; Ord(i)|] ==> k < i Un j" +by (simp add: lt_Un_iff lt_Ord2) + +(*See also Transset_iff_Union_succ*) +lemma Ord_Union_succ_eq: "Ord(i) ==> \(succ(i)) = i" +by (blast intro: Ord_trans) + (*FIXME: the Intersection duals are missing!*) (*** Results about limits ***) -lemma Ord_Union: "[| !!i. i:A ==> Ord(i) |] ==> Ord(Union(A))" +lemma Ord_Union [intro,simp,TC]: "[| !!i. i:A ==> Ord(i) |] ==> Ord(Union(A))" apply (rule Ord_is_Transset [THEN Transset_Union_family, THEN OrdI]) apply (blast intro: Ord_contains_Transset)+ done -lemma Ord_UN: "[| !!x. x:A ==> Ord(B(x)) |] ==> Ord(UN x:A. B(x))" +lemma Ord_UN [intro,simp,TC]: + "[| !!x. x:A ==> Ord(B(x)) |] ==> Ord(UN x:A. B(x))" by (rule Ord_Union, blast) (* No < version; consider (UN i:nat.i)=nat *) @@ -545,6 +559,10 @@ apply (blast intro: succ_leI)+ done +lemma UN_upper_lt: + "[| a\A; i < b(a); Ord(\x\A. b(x)) |] ==> i < (\x\A. b(x))" +by (unfold lt_def, blast) + lemma UN_upper_le: "[| a: A; i le b(a); Ord(UN x:A. b(x)) |] ==> i le (UN x:A. b(x))" apply (frule ltD) @@ -552,6 +570,15 @@ apply (blast intro: lt_Ord UN_upper)+ done +lemma lt_Union_iff: "\i\A. Ord(i) ==> (j < \(A)) <-> (\i\A. jj; Ord(\(J)) |] ==> i \ \J" +apply (subst Union_eq_UN) +apply (rule UN_upper_le, auto) +done + lemma le_implies_UN_le_UN: "[| !!x. x:A ==> c(x) le d(x) |] ==> (UN x:A. c(x)) le (UN x:A. d(x))" apply (rule UN_least_le) @@ -587,6 +614,18 @@ lemma Limit_has_succ: "[| Limit(i); j succ(j) < i" by (unfold Limit_def, blast) +lemma zero_not_Limit [iff]: "~ Limit(0)" +by (simp add: Limit_def) + +lemma Limit_has_1: "Limit(i) ==> 1 < i" +by (blast intro: Limit_has_0 Limit_has_succ) + +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 non_succ_LimitI: "[| 0 Limit(i)" apply (unfold Limit_def) @@ -608,6 +647,7 @@ lemma Limit_le_succD: "[| Limit(i); i le succ(j) |] ==> i le j" by (blast elim!: leE) + (** Traditional 3-way case analysis on ordinals **) lemma Ord_cases_disj: "Ord(i) ==> i=0 | (EX j. Ord(j) & i=succ(j)) | Limit(i)" @@ -631,6 +671,33 @@ apply (erule Ord_cases, blast+) done +text{*A set of ordinals is either empty, contains its own union, or its +union is a limit ordinal.*} +lemma Ord_set_cases: + "\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) + apply (blast intro: subst_elem) + apply auto +apply (clarify elim!: equalityE succ_subsetE) +apply (simp add: Union_subset_iff) +apply (subgoal_tac "B = succ(j)", blast) +apply (rule le_anti_sym) + apply (simp add: le_subset_iff) +apply (simp add: ltI) +done + +text{*If the union of a set of ordinals is a successor, then it is +an element of that set.*} +lemma Ord_Union_eq_succD: "[|\x\X. Ord(x); \X = succ(j)|] ==> succ(j) \ X" +by (drule Ord_set_cases, auto) + +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 + (*special induction rules for the "induct" method*) lemmas Ord_induct = Ord_induct [consumes 2] and Ord_induct_rule = Ord_induct [rule_format, consumes 2] diff -r 3208b614dc71 -r 03a5afa7b888 src/ZF/equalities.thy --- a/src/ZF/equalities.thy Wed May 22 17:26:34 2002 +0200 +++ b/src/ZF/equalities.thy Wed May 22 18:11:57 2002 +0200 @@ -198,6 +198,8 @@ lemma equal_singleton [rule_format]: "[| a: C; ALL y:C. y=b |] ==> C = {b}" by blast +lemma [simp]: "cons(a,cons(a,B)) = cons(a,B)" +by blast (** Binary Intersection **) diff -r 3208b614dc71 -r 03a5afa7b888 src/ZF/func.thy --- a/src/ZF/func.thy Wed May 22 17:26:34 2002 +0200 +++ b/src/ZF/func.thy Wed May 22 18:11:57 2002 +0200 @@ -22,6 +22,10 @@ lemma fun_is_function: "f: Pi(A,B) ==> function(f)" by (simp only: Pi_iff) +lemma functionI: + "[| !!x y y'. [| :r; :r |] ==> y=y' |] ==> function(r)" +by (simp add: function_def, blast) + (*Functions are relations*) lemma fun_is_rel: "f: Pi(A,B) ==> f <= Sigma(A,B)" by (unfold Pi_def, blast) @@ -135,6 +139,12 @@ lemma lam_funtype: "(lam x:A. b(x)) : A -> {b(x). x:A}" by (blast intro: lam_type); +lemma function_lam: "function (lam x:A. b(x))" +by (simp add: function_def lam_def) + +lemma relation_lam: "relation (lam x:A. b(x))" +by (simp add: relation_def lam_def) + lemma beta [simp]: "a : A ==> (lam x:A. b(x)) ` a = b(a)" by (blast intro: apply_equality lam_funtype lamI) @@ -256,6 +266,12 @@ apply (simp (no_asm)) done +lemma restrict_iff: "z \ restrict(r,A) \ z \ r & (\x\A. \y. z = \x, y\)" +by (simp add: restrict_def) + +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 domain_restrict_lam [simp]: "domain(restrict(Lambda(A,f),C)) = A Int C" apply (unfold restrict_def lam_def) apply (rule equalityI)