src/ZF/OrdQuant.thy
changeset 13162 660a71e712af
parent 13149 773657d466cb
child 13169 394a6c649547
equal deleted inserted replaced
13161:a40db0418145 13162:660a71e712af
    71 by blast
    71 by blast
    72 
    72 
    73 lemma trans_imp_trans_on: "trans(r) ==> trans[A](r)"
    73 lemma trans_imp_trans_on: "trans(r) ==> trans[A](r)"
    74 by (unfold trans_def trans_on_def, blast)
    74 by (unfold trans_def trans_on_def, blast)
    75 
    75 
    76 lemma image_is_UN: "\<lbrakk>function(g); x <= domain(g)\<rbrakk> \<Longrightarrow> g``x = (UN k:x. {g`k})"
    76 lemma image_is_UN: "[| function(g); x <= domain(g) |] ==> g``x = (UN k:x. {g`k})"
    77 by (blast intro: function_apply_equality [THEN sym] function_apply_Pair) 
    77 by (blast intro: function_apply_equality [THEN sym] function_apply_Pair) 
    78 
    78 
    79 lemma functionI: 
    79 lemma functionI: 
    80      "\<lbrakk>!!x y y'. \<lbrakk><x,y>:r; <x,y'>:r\<rbrakk> \<Longrightarrow> y=y'\<rbrakk> \<Longrightarrow> function(r)"
    80      "[| !!x y y'. [| <x,y>:r; <x,y'>:r |] ==> y=y' |] ==> function(r)"
    81 by (simp add: function_def, blast) 
    81 by (simp add: function_def, blast) 
    82 
    82 
    83 lemma function_lam: "function (lam x:A. b(x))"
    83 lemma function_lam: "function (lam x:A. b(x))"
    84 by (simp add: function_def lam_def) 
    84 by (simp add: function_def lam_def) 
    85 
    85 
    90 by (simp add: restrict_def) 
    90 by (simp add: restrict_def) 
    91 
    91 
    92 (** These mostly belong to theory Ordinal **)
    92 (** These mostly belong to theory Ordinal **)
    93 
    93 
    94 lemma Union_upper_le:
    94 lemma Union_upper_le:
    95      "\<lbrakk>j: J;  i\<le>j;  Ord(\<Union>(J))\<rbrakk> \<Longrightarrow> i \<le> \<Union>J"
    95      "[| j: J;  i\<le>j;  Ord(\<Union>(J)) |] ==> i \<le> \<Union>J"
    96 apply (subst Union_eq_UN)  
    96 apply (subst Union_eq_UN)  
    97 apply (rule UN_upper_le, auto)
    97 apply (rule UN_upper_le, auto)
    98 done
    98 done
    99 
    99 
   100 lemma zero_not_Limit [iff]: "~ Limit(0)"
   100 lemma zero_not_Limit [iff]: "~ Limit(0)"
   101 by (simp add: Limit_def)
   101 by (simp add: Limit_def)
   102 
   102 
   103 lemma Limit_has_1: "Limit(i) \<Longrightarrow> 1 < i"
   103 lemma Limit_has_1: "Limit(i) ==> 1 < i"
   104 by (blast intro: Limit_has_0 Limit_has_succ)
   104 by (blast intro: Limit_has_0 Limit_has_succ)
   105 
   105 
   106 lemma Limit_Union [rule_format]: "\<lbrakk>I \<noteq> 0;  \<forall>i\<in>I. Limit(i)\<rbrakk> \<Longrightarrow> Limit(\<Union>I)"
   106 lemma Limit_Union [rule_format]: "[| I \<noteq> 0;  \<forall>i\<in>I. Limit(i) |] ==> Limit(\<Union>I)"
   107 apply (simp add: Limit_def lt_def)
   107 apply (simp add: Limit_def lt_def)
   108 apply (blast intro!: equalityI)
   108 apply (blast intro!: equalityI)
   109 done
   109 done
   110 
   110 
   111 lemma increasing_LimitI: "\<lbrakk>0<l; \<forall>x\<in>l. \<exists>y\<in>l. x<y\<rbrakk> \<Longrightarrow> Limit(l)"
   111 lemma increasing_LimitI: "[| 0<l; \<forall>x\<in>l. \<exists>y\<in>l. x<y |] ==> Limit(l)"
   112 apply (simp add: Limit_def lt_Ord2, clarify)
   112 apply (simp add: Limit_def lt_Ord2, clarify)
   113 apply (drule_tac i=y in ltD) 
   113 apply (drule_tac i=y in ltD) 
   114 apply (blast intro: lt_trans1 [OF _ ltI] lt_Ord2)
   114 apply (blast intro: lt_trans1 [OF _ ltI] lt_Ord2)
   115 done
   115 done
   116 
   116 
   117 lemma UN_upper_lt:
   117 lemma UN_upper_lt:
   118      "\<lbrakk>a\<in>A;  i < b(a);  Ord(\<Union>x\<in>A. b(x))\<rbrakk> \<Longrightarrow> i < (\<Union>x\<in>A. b(x))"
   118      "[| a\<in>A;  i < b(a);  Ord(\<Union>x\<in>A. b(x)) |] ==> i < (\<Union>x\<in>A. b(x))"
   119 by (unfold lt_def, blast) 
   119 by (unfold lt_def, blast) 
   120 
   120 
   121 lemma lt_imp_0_lt: "j<i \<Longrightarrow> 0<i"
   121 lemma lt_imp_0_lt: "j<i ==> 0<i"
   122 by (blast intro: lt_trans1 Ord_0_le [OF lt_Ord]) 
   122 by (blast intro: lt_trans1 Ord_0_le [OF lt_Ord]) 
   123 
   123 
   124 lemma Ord_set_cases:
   124 lemma Ord_set_cases:
   125    "\<forall>i\<in>I. Ord(i) \<Longrightarrow> I=0 \<or> \<Union>(I) \<in> I \<or> (\<Union>(I) \<notin> I \<and> Limit(\<Union>(I)))"
   125    "\<forall>i\<in>I. Ord(i) ==> I=0 \<or> \<Union>(I) \<in> I \<or> (\<Union>(I) \<notin> I \<and> Limit(\<Union>(I)))"
   126 apply (clarify elim!: not_emptyE) 
   126 apply (clarify elim!: not_emptyE) 
   127 apply (cases "\<Union>(I)" rule: Ord_cases) 
   127 apply (cases "\<Union>(I)" rule: Ord_cases) 
   128    apply (blast intro: Ord_Union)
   128    apply (blast intro: Ord_Union)
   129   apply (blast intro: subst_elem)
   129   apply (blast intro: subst_elem)
   130  apply auto 
   130  apply auto 
   138 
   138 
   139 lemma Ord_Union_eq_succD: "[|\<forall>x\<in>X. Ord(x);  \<Union>X = succ(j)|] ==> succ(j) \<in> X"
   139 lemma Ord_Union_eq_succD: "[|\<forall>x\<in>X. Ord(x);  \<Union>X = succ(j)|] ==> succ(j) \<in> X"
   140 by (drule Ord_set_cases, auto)
   140 by (drule Ord_set_cases, auto)
   141 
   141 
   142 (*See also Transset_iff_Union_succ*)
   142 (*See also Transset_iff_Union_succ*)
   143 lemma Ord_Union_succ_eq: "Ord(i) \<Longrightarrow> \<Union>(succ(i)) = i"
   143 lemma Ord_Union_succ_eq: "Ord(i) ==> \<Union>(succ(i)) = i"
   144 by (blast intro: Ord_trans)
   144 by (blast intro: Ord_trans)
   145 
   145 
   146 lemma lt_Union_iff: "\<forall>i\<in>A. Ord(i) \<Longrightarrow> (j < \<Union>(A)) <-> (\<exists>i\<in>A. j<i)"
   146 lemma lt_Union_iff: "\<forall>i\<in>A. Ord(i) ==> (j < \<Union>(A)) <-> (\<exists>i\<in>A. j<i)"
   147 by (auto simp: lt_def Ord_Union)
   147 by (auto simp: lt_def Ord_Union)
   148 
   148 
   149 lemma Un_upper1_lt: "[|k < i; Ord(j)|] ==> k < i Un j"
   149 lemma Un_upper1_lt: "[|k < i; Ord(j)|] ==> k < i Un j"
   150 by (simp add: lt_def) 
   150 by (simp add: lt_def) 
   151 
   151 
   152 lemma Un_upper2_lt: "[|k < j; Ord(i)|] ==> k < i Un j"
   152 lemma Un_upper2_lt: "[|k < j; Ord(i)|] ==> k < i Un j"
   153 by (simp add: lt_def) 
   153 by (simp add: lt_def) 
   154 
   154 
   155 lemma Ord_OUN [intro,simp]:
   155 lemma Ord_OUN [intro,simp]:
   156      "\<lbrakk>!!x. x<A \<Longrightarrow> Ord(B(x))\<rbrakk> \<Longrightarrow> Ord(\<Union>x<A. B(x))"
   156      "[| !!x. x<A ==> Ord(B(x)) |] ==> Ord(\<Union>x<A. B(x))"
   157 by (simp add: OUnion_def ltI Ord_UN) 
   157 by (simp add: OUnion_def ltI Ord_UN) 
   158 
   158 
   159 lemma OUN_upper_lt:
   159 lemma OUN_upper_lt:
   160      "\<lbrakk>a<A;  i < b(a);  Ord(\<Union>x<A. b(x))\<rbrakk> \<Longrightarrow> i < (\<Union>x<A. b(x))"
   160      "[| a<A;  i < b(a);  Ord(\<Union>x<A. b(x)) |] ==> i < (\<Union>x<A. b(x))"
   161 by (unfold OUnion_def lt_def, blast )
   161 by (unfold OUnion_def lt_def, blast )
   162 
   162 
   163 lemma OUN_upper_le:
   163 lemma OUN_upper_le:
   164      "\<lbrakk>a<A;  i\<le>b(a);  Ord(\<Union>x<A. b(x))\<rbrakk> \<Longrightarrow> i \<le> (\<Union>x<A. b(x))"
   164      "[| a<A;  i\<le>b(a);  Ord(\<Union>x<A. b(x)) |] ==> i \<le> (\<Union>x<A. b(x))"
   165 apply (unfold OUnion_def, auto)
   165 apply (unfold OUnion_def, auto)
   166 apply (rule UN_upper_le )
   166 apply (rule UN_upper_le )
   167 apply (auto simp add: lt_def) 
   167 apply (auto simp add: lt_def) 
   168 done
   168 done
   169 
   169