tidying up
authorpaulson
Wed, 22 May 2002 18:11:57 +0200
changeset 13172 03a5afa7b888
parent 13171 3208b614dc71
child 13173 8f4680be79cc
tidying up
src/ZF/OrdQuant.thy
src/ZF/Ordinal.thy
src/ZF/equalities.thy
src/ZF/func.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'. [| <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))"
-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 \<in> restrict(r,A) \<longleftrightarrow> z \<in> r & (\<exists>x\<in>A. \<exists>y. z = \<langle>x, y\<rangle>)"
-by (simp add: restrict_def) 
-
-(** These mostly belong to theory Ordinal **)
-
-lemma Union_upper_le:
-     "[| j: J;  i\<le>j;  Ord(\<Union>(J)) |] ==> i \<le> \<Union>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 \<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: "[| 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:
-     "[| 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 ==> 0<i"
-by (blast intro: lt_trans1 Ord_0_le [OF lt_Ord]) 
-
-lemma Ord_set_cases:
-   "\<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)
-  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: "[|\<forall>x\<in>X. Ord(x);  \<Union>X = succ(j)|] ==> succ(j) \<in> X"
-by (drule Ord_set_cases, auto)
-
-(*See also Transset_iff_Union_succ*)
-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) ==> (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"
-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<A ==> Ord(B(x)) |] ==> Ord(\<Union>x<A. B(x))"
 by (simp add: OUnion_def ltI Ord_UN) 
@@ -272,7 +188,6 @@
 
 declare ltD [THEN beta, simp]
 
-
 lemma lt_induct: 
     "[| i<k;  !!x.[| x<k;  ALL y<x. P(y) |] ==> P(x) |]  ==>  P(i)"
 apply (simp add: lt_def oall_def)
--- 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<i ==> 0<i"
+by (blast intro: lt_trans1 Ord_0_le [OF lt_Ord]) 
+
 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) 
@@ -518,17 +521,28 @@
      "[| 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]) 
 
+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) ==> \<Union>(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\<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 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: "\<forall>i\<in>A. Ord(i) ==> (j < \<Union>(A)) <-> (\<exists>i\<in>A. j<i)"
+by (auto simp: lt_def Ord_Union)
+
+lemma Union_upper_le:
+     "[| j: J;  i\<le>j;  Ord(\<Union>(J)) |] ==> i \<le> \<Union>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<i |] ==> 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: "[| 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 non_succ_LimitI: 
     "[| 0<i;  ALL y. succ(y) ~= i |] ==> 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:
+   "\<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)
+  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: "[|\<forall>x\<in>X. Ord(x);  \<Union>X = succ(j)|] ==> succ(j) \<in> X"
+by (drule Ord_set_cases, auto)
+
+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
+
 (*special induction rules for the "induct" method*)
 lemmas Ord_induct = Ord_induct [consumes 2]
   and Ord_induct_rule = Ord_induct [rule_format, consumes 2]
--- 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 **)
 
--- 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'. [| <x,y>:r; <x,y'>: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 \<in> restrict(r,A) \<longleftrightarrow> z \<in> r & (\<exists>x\<in>A. \<exists>y. z = \<langle>x, y\<rangle>)"
+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)