merged
authorhaftmann
Thu, 21 Jul 2011 21:56:24 +0200
changeset 448133406cd754dd2
parent 44809 78a0a2ad91a3
parent 44812 481566bc20e4
child 44814 e6928fc2332a
merged
     1.1 --- a/NEWS	Thu Jul 21 08:33:57 2011 +0200
     1.2 +++ b/NEWS	Thu Jul 21 21:56:24 2011 +0200
     1.3 @@ -63,15 +63,16 @@
     1.4  * Classes bot and top require underlying partial order rather than preorder:
     1.5  uniqueness of bot and top is guaranteed.  INCOMPATIBILITY.
     1.6  
     1.7 -* Class 'complete_lattice': generalized a couple of lemmas from sets;
     1.8 -generalized theorems INF_cong and SUP_cong.  More consistent and less
     1.9 -misunderstandable names:
    1.10 +* Class complete_lattice: generalized a couple of lemmas from sets;
    1.11 +generalized theorems INF_cong and SUP_cong.  New type classes for complete
    1.12 +boolean algebras and complete linear orderes.  Lemmas Inf_less_iff,
    1.13 +less_Sup_iff, INF_less_iff, less_SUP_iff now reside in class complete_linorder.
    1.14 +More consistent and less misunderstandable names:
    1.15    INFI_def ~> INF_def
    1.16    SUPR_def ~> SUP_def
    1.17    le_SUPI ~> le_SUP_I
    1.18    le_SUPI2 ~> le_SUP_I2
    1.19    le_INFI ~> le_INF_I
    1.20 -  INF_subset ~> INF_superset_mono
    1.21    INFI_bool_eq ~> INF_bool_eq
    1.22    SUPR_bool_eq ~> SUP_bool_eq
    1.23    INFI_apply ~> INF_apply
     2.1 --- a/src/HOL/Complete_Lattice.thy	Thu Jul 21 08:33:57 2011 +0200
     2.2 +++ b/src/HOL/Complete_Lattice.thy	Thu Jul 21 21:56:24 2011 +0200
     2.3 @@ -292,12 +292,13 @@
     2.4    by (force intro!: Sup_mono simp: SUP_def)
     2.5  
     2.6  lemma INF_superset_mono:
     2.7 -  "B \<subseteq> A \<Longrightarrow> INFI A f \<sqsubseteq> INFI B f"
     2.8 -  by (rule INF_mono) auto
     2.9 +  "B \<subseteq> A \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> f x \<sqsubseteq> g x) \<Longrightarrow> (\<Sqinter>x\<in>A. f x) \<sqsubseteq> (\<Sqinter>x\<in>B. g x)"
    2.10 +  -- {* The last inclusion is POSITIVE! *}
    2.11 +  by (blast intro: INF_mono dest: subsetD)
    2.12  
    2.13  lemma SUP_subset_mono:
    2.14 -  "A \<subseteq> B \<Longrightarrow> SUPR A f \<sqsubseteq> SUPR B f"
    2.15 -  by (rule SUP_mono) auto
    2.16 +  "A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<sqsubseteq> g x) \<Longrightarrow> (\<Squnion>x\<in>A. f x) \<sqsubseteq> (\<Squnion>x\<in>B. g x)"
    2.17 +  by (blast intro: SUP_mono dest: subsetD)
    2.18  
    2.19  lemma INF_commute: "(\<Sqinter>i\<in>A. \<Sqinter>j\<in>B. f i j) = (\<Sqinter>j\<in>B. \<Sqinter>i\<in>A. f i j)"
    2.20    by (iprover intro: INF_leI le_INF_I order_trans antisym)
    2.21 @@ -371,38 +372,8 @@
    2.22    "(\<Squnion>b. A b) = A True \<squnion> A False"
    2.23    by (simp add: UNIV_bool SUP_empty SUP_insert sup_commute)
    2.24  
    2.25 -lemma INF_mono':
    2.26 -  "B \<subseteq> A \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> f x \<sqsubseteq> g x) \<Longrightarrow> (\<Sqinter>x\<in>A. f x) \<sqsubseteq> (\<Sqinter>x\<in>B. g x)"
    2.27 -  -- {* The last inclusion is POSITIVE! *}
    2.28 -  by (rule INF_mono) auto
    2.29 -
    2.30 -lemma SUP_mono':
    2.31 -  "A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<sqsubseteq> g x) \<Longrightarrow> (\<Squnion>x\<in>A. f x) \<sqsubseteq> (\<Squnion>x\<in>B. g x)"
    2.32 -  -- {* The last inclusion is POSITIVE! *}
    2.33 -  by (blast intro: SUP_mono dest: subsetD)
    2.34 -
    2.35  end
    2.36  
    2.37 -lemma Inf_less_iff:
    2.38 -  fixes a :: "'a\<Colon>{complete_lattice,linorder}"
    2.39 -  shows "\<Sqinter>S \<sqsubset> a \<longleftrightarrow> (\<exists>x\<in>S. x \<sqsubset> a)"
    2.40 -  unfolding not_le [symmetric] le_Inf_iff by auto
    2.41 -
    2.42 -lemma less_Sup_iff:
    2.43 -  fixes a :: "'a\<Colon>{complete_lattice,linorder}"
    2.44 -  shows "a \<sqsubset> \<Squnion>S \<longleftrightarrow> (\<exists>x\<in>S. a \<sqsubset> x)"
    2.45 -  unfolding not_le [symmetric] Sup_le_iff by auto
    2.46 -
    2.47 -lemma INF_less_iff:
    2.48 -  fixes a :: "'a::{complete_lattice,linorder}"
    2.49 -  shows "(\<Sqinter>i\<in>A. f i) \<sqsubset> a \<longleftrightarrow> (\<exists>x\<in>A. f x \<sqsubset> a)"
    2.50 -  unfolding INF_def Inf_less_iff by auto
    2.51 -
    2.52 -lemma less_SUP_iff:
    2.53 -  fixes a :: "'a::{complete_lattice,linorder}"
    2.54 -  shows "a \<sqsubset> (\<Squnion>i\<in>A. f i) \<longleftrightarrow> (\<exists>x\<in>A. a \<sqsubset> f x)"
    2.55 -  unfolding SUP_def less_Sup_iff by auto
    2.56 -
    2.57  class complete_boolean_algebra = boolean_algebra + complete_lattice
    2.58  begin
    2.59  
    2.60 @@ -430,6 +401,27 @@
    2.61  
    2.62  end
    2.63  
    2.64 +class complete_linorder = linorder + complete_lattice
    2.65 +begin
    2.66 +
    2.67 +lemma Inf_less_iff:
    2.68 +  "\<Sqinter>S \<sqsubset> a \<longleftrightarrow> (\<exists>x\<in>S. x \<sqsubset> a)"
    2.69 +  unfolding not_le [symmetric] le_Inf_iff by auto
    2.70 +
    2.71 +lemma less_Sup_iff:
    2.72 +  "a \<sqsubset> \<Squnion>S \<longleftrightarrow> (\<exists>x\<in>S. a \<sqsubset> x)"
    2.73 +  unfolding not_le [symmetric] Sup_le_iff by auto
    2.74 +
    2.75 +lemma INF_less_iff:
    2.76 +  "(\<Sqinter>i\<in>A. f i) \<sqsubset> a \<longleftrightarrow> (\<exists>x\<in>A. f x \<sqsubset> a)"
    2.77 +  unfolding INF_def Inf_less_iff by auto
    2.78 +
    2.79 +lemma less_SUP_iff:
    2.80 +  "a \<sqsubset> (\<Squnion>i\<in>A. f i) \<longleftrightarrow> (\<exists>x\<in>A. a \<sqsubset> f x)"
    2.81 +  unfolding SUP_def less_Sup_iff by auto
    2.82 +
    2.83 +end
    2.84 +
    2.85  
    2.86  subsection {* @{typ bool} and @{typ "_ \<Rightarrow> _"} as complete lattice *}
    2.87  
    2.88 @@ -688,7 +680,7 @@
    2.89  lemma INT_anti_mono:
    2.90    "A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<subseteq> g x) \<Longrightarrow> (\<Inter>x\<in>B. f x) \<subseteq> (\<Inter>x\<in>A. g x)"
    2.91    -- {* The last inclusion is POSITIVE! *}
    2.92 -  by (fact INF_mono')
    2.93 +  by (fact INF_superset_mono)
    2.94  
    2.95  lemma Pow_INT_eq: "Pow (\<Inter>x\<in>A. B x) = (\<Inter>x\<in>A. Pow (B x))"
    2.96    by blast
    2.97 @@ -922,7 +914,7 @@
    2.98  lemma UN_mono:
    2.99    "A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<subseteq> g x) \<Longrightarrow>
   2.100      (\<Union>x\<in>A. f x) \<subseteq> (\<Union>x\<in>B. g x)"
   2.101 -  by (fact SUP_mono')
   2.102 +  by (fact SUP_subset_mono)
   2.103  
   2.104  lemma vimage_Union: "f -` (\<Union>A) = (\<Union>X\<in>A. f -` X)"
   2.105    by blast
   2.106 @@ -1083,7 +1075,11 @@
   2.107  lemmas (in complete_lattice) le_SUPI = le_SUP_I
   2.108  lemmas (in complete_lattice) le_SUPI2 = le_SUP_I2
   2.109  lemmas (in complete_lattice) le_INFI = le_INF_I
   2.110 -lemmas (in complete_lattice) INF_subset = INF_superset_mono 
   2.111 +
   2.112 +lemma (in complete_lattice) INF_subset:
   2.113 +  "B \<subseteq> A \<Longrightarrow> INFI A f \<sqsubseteq> INFI B f"
   2.114 +  by (rule INF_superset_mono) auto
   2.115 +
   2.116  lemmas INFI_apply = INF_apply
   2.117  lemmas SUPR_apply = SUP_apply
   2.118  
     3.1 --- a/src/HOL/Library/Extended_Real.thy	Thu Jul 21 08:33:57 2011 +0200
     3.2 +++ b/src/HOL/Library/Extended_Real.thy	Thu Jul 21 21:56:24 2011 +0200
     3.3 @@ -8,7 +8,7 @@
     3.4  header {* Extended real number line *}
     3.5  
     3.6  theory Extended_Real
     3.7 -  imports Complex_Main Extended_Nat
     3.8 +imports Complex_Main Extended_Nat
     3.9  begin
    3.10  
    3.11  text {*
    3.12 @@ -1244,8 +1244,11 @@
    3.13      with Sup_le[of "uminus`A" "-x"] show "x \<le> Inf A"
    3.14        unfolding ereal_Sup_uminus_image_eq by force }
    3.15  qed
    3.16 +
    3.17  end
    3.18  
    3.19 +instance ereal :: complete_linorder ..
    3.20 +
    3.21  lemma ereal_SUPR_uminus:
    3.22    fixes f :: "'a => ereal"
    3.23    shows "(SUP i : R. -(f i)) = -(INF i : R. f i)"
    3.24 @@ -1335,12 +1338,11 @@
    3.25      by (cases e) auto
    3.26  qed
    3.27  
    3.28 -lemma Sup_eq_top_iff:
    3.29 -  fixes A :: "'a::{complete_lattice, linorder} set"
    3.30 -  shows "Sup A = top \<longleftrightarrow> (\<forall>x<top. \<exists>i\<in>A. x < i)"
    3.31 +lemma (in complete_linorder) Sup_eq_top_iff: -- "FIXME move"
    3.32 +  "Sup A = top \<longleftrightarrow> (\<forall>x<top. \<exists>i\<in>A. x < i)"
    3.33  proof
    3.34    assume *: "Sup A = top"
    3.35 -  show "(\<forall>x<top. \<exists>i\<in>A. x < i)" unfolding *[symmetric]
    3.36 +  show "(\<forall>x<top. \<exists>i\<in>A. x < i)" unfolding * [symmetric]
    3.37    proof (intro allI impI)
    3.38      fix x assume "x < Sup A" then show "\<exists>i\<in>A. x < i"
    3.39        unfolding less_Sup_iff by auto
    3.40 @@ -1350,7 +1352,7 @@
    3.41    show "Sup A = top"
    3.42    proof (rule ccontr)
    3.43      assume "Sup A \<noteq> top"
    3.44 -    with top_greatest[of "Sup A"]
    3.45 +    with top_greatest [of "Sup A"]
    3.46      have "Sup A < top" unfolding le_less by auto
    3.47      then have "Sup A < Sup A"
    3.48        using * unfolding less_Sup_iff by auto
    3.49 @@ -1358,8 +1360,8 @@
    3.50    qed
    3.51  qed
    3.52  
    3.53 -lemma SUP_eq_top_iff:
    3.54 -  fixes f :: "'a \<Rightarrow> 'b::{complete_lattice, linorder}"
    3.55 +lemma (in complete_linorder) SUP_eq_top_iff: -- "FIXME move"
    3.56 +  fixes f :: "'b \<Rightarrow> 'a"
    3.57    shows "(SUP i:A. f i) = top \<longleftrightarrow> (\<forall>x<top. \<exists>i\<in>A. x < f i)"
    3.58    unfolding SUPR_def Sup_eq_top_iff by auto
    3.59  
    3.60 @@ -2182,12 +2184,12 @@
    3.61    "Limsup net f = (LEAST l. \<forall>y>l. eventually (\<lambda>x. f x < y) net)"
    3.62  
    3.63  lemma Liminf_Sup:
    3.64 -  fixes f :: "'a => 'b::{complete_lattice, linorder}"
    3.65 +  fixes f :: "'a => 'b::complete_linorder"
    3.66    shows "Liminf net f = Sup {l. \<forall>y<l. eventually (\<lambda>x. y < f x) net}"
    3.67    by (auto intro!: Greatest_equality complete_lattice_class.Sup_upper simp: less_Sup_iff Liminf_def)
    3.68  
    3.69  lemma Limsup_Inf:
    3.70 -  fixes f :: "'a => 'b::{complete_lattice, linorder}"
    3.71 +  fixes f :: "'a => 'b::complete_linorder"
    3.72    shows "Limsup net f = Inf {l. \<forall>y>l. eventually (\<lambda>x. f x < y) net}"
    3.73    by (auto intro!: Least_equality complete_lattice_class.Inf_lower simp: Inf_less_iff Limsup_def)
    3.74  
    3.75 @@ -2208,7 +2210,7 @@
    3.76    using assms by (auto intro!: Greatest_equality)
    3.77  
    3.78  lemma Limsup_const:
    3.79 -  fixes c :: "'a::{complete_lattice, linorder}"
    3.80 +  fixes c :: "'a::complete_linorder"
    3.81    assumes ntriv: "\<not> trivial_limit net"
    3.82    shows "Limsup net (\<lambda>x. c) = c"
    3.83    unfolding Limsup_Inf
    3.84 @@ -2222,7 +2224,7 @@
    3.85  qed auto
    3.86  
    3.87  lemma Liminf_const:
    3.88 -  fixes c :: "'a::{complete_lattice, linorder}"
    3.89 +  fixes c :: "'a::complete_linorder"
    3.90    assumes ntriv: "\<not> trivial_limit net"
    3.91    shows "Liminf net (\<lambda>x. c) = c"
    3.92    unfolding Liminf_Sup
    3.93 @@ -2235,18 +2237,17 @@
    3.94    qed
    3.95  qed auto
    3.96  
    3.97 -lemma mono_set:
    3.98 -  fixes S :: "('a::order) set"
    3.99 -  shows "mono S \<longleftrightarrow> (\<forall>x y. x \<le> y \<longrightarrow> x \<in> S \<longrightarrow> y \<in> S)"
   3.100 +lemma (in order) mono_set:
   3.101 +  "mono S \<longleftrightarrow> (\<forall>x y. x \<le> y \<longrightarrow> x \<in> S \<longrightarrow> y \<in> S)"
   3.102    by (auto simp: mono_def mem_def)
   3.103  
   3.104 -lemma mono_greaterThan[intro, simp]: "mono {B<..}" unfolding mono_set by auto
   3.105 -lemma mono_atLeast[intro, simp]: "mono {B..}" unfolding mono_set by auto
   3.106 -lemma mono_UNIV[intro, simp]: "mono UNIV" unfolding mono_set by auto
   3.107 -lemma mono_empty[intro, simp]: "mono {}" unfolding mono_set by auto
   3.108 +lemma (in order) mono_greaterThan [intro, simp]: "mono {B<..}" unfolding mono_set by auto
   3.109 +lemma (in order) mono_atLeast [intro, simp]: "mono {B..}" unfolding mono_set by auto
   3.110 +lemma (in order) mono_UNIV [intro, simp]: "mono UNIV" unfolding mono_set by auto
   3.111 +lemma (in order) mono_empty [intro, simp]: "mono {}" unfolding mono_set by auto
   3.112  
   3.113 -lemma mono_set_iff:
   3.114 -  fixes S :: "'a::{linorder,complete_lattice} set"
   3.115 +lemma (in complete_linorder) mono_set_iff:
   3.116 +  fixes S :: "'a set"
   3.117    defines "a \<equiv> Inf S"
   3.118    shows "mono S \<longleftrightarrow> (S = {a <..} \<or> S = {a..})" (is "_ = ?c")
   3.119  proof
   3.120 @@ -2257,13 +2258,13 @@
   3.121      assume "a \<in> S"
   3.122      show ?c
   3.123        using mono[OF _ `a \<in> S`]
   3.124 -      by (auto intro: complete_lattice_class.Inf_lower simp: a_def)
   3.125 +      by (auto intro: Inf_lower simp: a_def)
   3.126    next
   3.127      assume "a \<notin> S"
   3.128      have "S = {a <..}"
   3.129      proof safe
   3.130        fix x assume "x \<in> S"
   3.131 -      then have "a \<le> x" unfolding a_def by (rule complete_lattice_class.Inf_lower)
   3.132 +      then have "a \<le> x" unfolding a_def by (rule Inf_lower)
   3.133        then show "a < x" using `x \<in> S` `a \<notin> S` by (cases "a = x") auto
   3.134      next
   3.135        fix x assume "a < x"
     4.1 --- a/src/HOL/Probability/Lebesgue_Integration.thy	Thu Jul 21 08:33:57 2011 +0200
     4.2 +++ b/src/HOL/Probability/Lebesgue_Integration.thy	Thu Jul 21 21:56:24 2011 +0200
     4.3 @@ -6,7 +6,7 @@
     4.4  header {*Lebesgue Integration*}
     4.5  
     4.6  theory Lebesgue_Integration
     4.7 -imports Measure Borel_Space
     4.8 +  imports Measure Borel_Space
     4.9  begin
    4.10  
    4.11  lemma real_ereal_1[simp]: "real (1::ereal) = 1"