src/HOL/Conditionally_Complete_Lattices.thy
author hoelzl
Tue, 05 Nov 2013 09:45:02 +0100
changeset 55715 c4159fe6fa46
parent 55714 326fd7103cb4
child 55733 b01057e72233
permissions -rw-r--r--
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
wenzelm@53402
     1
(*  Title:      HOL/Conditionally_Complete_Lattices.thy
hoelzl@52655
     2
    Author:     Amine Chaieb and L C Paulson, University of Cambridge
hoelzl@52780
     3
    Author:     Johannes Hölzl, TU München
hoelzl@55710
     4
    Author:     Luke S. Serafin, Carnegie Mellon University
hoelzl@52655
     5
*)
paulson@33269
     6
hoelzl@52910
     7
header {* Conditionally-complete Lattices *}
paulson@33269
     8
hoelzl@52910
     9
theory Conditionally_Complete_Lattices
hoelzl@55715
    10
imports Main
paulson@33269
    11
begin
paulson@33269
    12
hoelzl@55711
    13
lemma (in linorder) Sup_fin_eq_Max: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup_fin X = Max X"
hoelzl@52612
    14
  by (induct X rule: finite_ne_induct) (simp_all add: sup_max)
hoelzl@52612
    15
hoelzl@55711
    16
lemma (in linorder) Inf_fin_eq_Min: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf_fin X = Min X"
hoelzl@52612
    17
  by (induct X rule: finite_ne_induct) (simp_all add: inf_min)
hoelzl@52612
    18
hoelzl@55710
    19
context preorder
hoelzl@55710
    20
begin
hoelzl@55710
    21
hoelzl@55710
    22
definition "bdd_above A \<longleftrightarrow> (\<exists>M. \<forall>x \<in> A. x \<le> M)"
hoelzl@55710
    23
definition "bdd_below A \<longleftrightarrow> (\<exists>m. \<forall>x \<in> A. m \<le> x)"
hoelzl@55710
    24
hoelzl@55710
    25
lemma bdd_aboveI[intro]: "(\<And>x. x \<in> A \<Longrightarrow> x \<le> M) \<Longrightarrow> bdd_above A"
hoelzl@55710
    26
  by (auto simp: bdd_above_def)
hoelzl@55710
    27
hoelzl@55710
    28
lemma bdd_belowI[intro]: "(\<And>x. x \<in> A \<Longrightarrow> m \<le> x) \<Longrightarrow> bdd_below A"
hoelzl@55710
    29
  by (auto simp: bdd_below_def)
hoelzl@55710
    30
hoelzl@55715
    31
lemma bdd_aboveI2: "(\<And>x. x \<in> A \<Longrightarrow> f x \<le> M) \<Longrightarrow> bdd_above (f`A)"
hoelzl@55715
    32
  by force
hoelzl@55715
    33
hoelzl@55715
    34
lemma bdd_belowI2: "(\<And>x. x \<in> A \<Longrightarrow> m \<le> f x) \<Longrightarrow> bdd_below (f`A)"
hoelzl@55715
    35
  by force
hoelzl@55715
    36
hoelzl@55710
    37
lemma bdd_above_empty [simp, intro]: "bdd_above {}"
hoelzl@55710
    38
  unfolding bdd_above_def by auto
hoelzl@55710
    39
hoelzl@55710
    40
lemma bdd_below_empty [simp, intro]: "bdd_below {}"
hoelzl@55710
    41
  unfolding bdd_below_def by auto
hoelzl@55710
    42
hoelzl@55710
    43
lemma bdd_above_mono: "bdd_above B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> bdd_above A"
hoelzl@55710
    44
  by (metis (full_types) bdd_above_def order_class.le_neq_trans psubsetD)
hoelzl@55710
    45
hoelzl@55710
    46
lemma bdd_below_mono: "bdd_below B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> bdd_below A"
hoelzl@55710
    47
  by (metis bdd_below_def order_class.le_neq_trans psubsetD)
hoelzl@55710
    48
hoelzl@55710
    49
lemma bdd_above_Int1 [simp]: "bdd_above A \<Longrightarrow> bdd_above (A \<inter> B)"
hoelzl@55710
    50
  using bdd_above_mono by auto
hoelzl@55710
    51
hoelzl@55710
    52
lemma bdd_above_Int2 [simp]: "bdd_above B \<Longrightarrow> bdd_above (A \<inter> B)"
hoelzl@55710
    53
  using bdd_above_mono by auto
hoelzl@55710
    54
hoelzl@55710
    55
lemma bdd_below_Int1 [simp]: "bdd_below A \<Longrightarrow> bdd_below (A \<inter> B)"
hoelzl@55710
    56
  using bdd_below_mono by auto
hoelzl@55710
    57
hoelzl@55710
    58
lemma bdd_below_Int2 [simp]: "bdd_below B \<Longrightarrow> bdd_below (A \<inter> B)"
hoelzl@55710
    59
  using bdd_below_mono by auto
hoelzl@55710
    60
hoelzl@55710
    61
lemma bdd_above_Ioo [simp, intro]: "bdd_above {a <..< b}"
hoelzl@55710
    62
  by (auto simp add: bdd_above_def intro!: exI[of _ b] less_imp_le)
hoelzl@55710
    63
hoelzl@55710
    64
lemma bdd_above_Ico [simp, intro]: "bdd_above {a ..< b}"
hoelzl@55710
    65
  by (auto simp add: bdd_above_def intro!: exI[of _ b] less_imp_le)
hoelzl@55710
    66
hoelzl@55710
    67
lemma bdd_above_Iio [simp, intro]: "bdd_above {..< b}"
hoelzl@55710
    68
  by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le)
hoelzl@55710
    69
hoelzl@55710
    70
lemma bdd_above_Ioc [simp, intro]: "bdd_above {a <.. b}"
hoelzl@55710
    71
  by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le)
hoelzl@55710
    72
hoelzl@55710
    73
lemma bdd_above_Icc [simp, intro]: "bdd_above {a .. b}"
hoelzl@55710
    74
  by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le)
hoelzl@55710
    75
hoelzl@55710
    76
lemma bdd_above_Iic [simp, intro]: "bdd_above {.. b}"
hoelzl@55710
    77
  by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le)
hoelzl@55710
    78
hoelzl@55710
    79
lemma bdd_below_Ioo [simp, intro]: "bdd_below {a <..< b}"
hoelzl@55710
    80
  by (auto simp add: bdd_below_def intro!: exI[of _ a] less_imp_le)
hoelzl@55710
    81
hoelzl@55710
    82
lemma bdd_below_Ioc [simp, intro]: "bdd_below {a <.. b}"
hoelzl@55710
    83
  by (auto simp add: bdd_below_def intro!: exI[of _ a] less_imp_le)
hoelzl@55710
    84
hoelzl@55710
    85
lemma bdd_below_Ioi [simp, intro]: "bdd_below {a <..}"
hoelzl@55710
    86
  by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le)
hoelzl@55710
    87
hoelzl@55710
    88
lemma bdd_below_Ico [simp, intro]: "bdd_below {a ..< b}"
hoelzl@55710
    89
  by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le)
hoelzl@55710
    90
hoelzl@55710
    91
lemma bdd_below_Icc [simp, intro]: "bdd_below {a .. b}"
hoelzl@55710
    92
  by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le)
hoelzl@55710
    93
hoelzl@55710
    94
lemma bdd_below_Ici [simp, intro]: "bdd_below {a ..}"
hoelzl@55710
    95
  by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le)
hoelzl@55710
    96
hoelzl@55710
    97
end
hoelzl@55710
    98
hoelzl@55713
    99
lemma (in order_top) bdd_above_top[simp, intro!]: "bdd_above A"
hoelzl@55713
   100
  by (rule bdd_aboveI[of _ top]) simp
hoelzl@55713
   101
hoelzl@55713
   102
lemma (in order_bot) bdd_above_bot[simp, intro!]: "bdd_below A"
hoelzl@55713
   103
  by (rule bdd_belowI[of _ bot]) simp
hoelzl@55713
   104
hoelzl@55714
   105
lemma bdd_above_uminus[simp]:
hoelzl@55714
   106
  fixes X :: "'a::ordered_ab_group_add set"
hoelzl@55714
   107
  shows "bdd_above (uminus ` X) \<longleftrightarrow> bdd_below X"
hoelzl@55714
   108
  by (auto simp: bdd_above_def bdd_below_def intro: le_imp_neg_le) (metis le_imp_neg_le minus_minus)
hoelzl@55714
   109
hoelzl@55714
   110
lemma bdd_below_uminus[simp]:
hoelzl@55714
   111
  fixes X :: "'a::ordered_ab_group_add set"
hoelzl@55714
   112
  shows"bdd_below (uminus ` X) \<longleftrightarrow> bdd_above X"
hoelzl@55714
   113
  by (auto simp: bdd_above_def bdd_below_def intro: le_imp_neg_le) (metis le_imp_neg_le minus_minus)
hoelzl@55714
   114
hoelzl@55710
   115
context lattice
hoelzl@55710
   116
begin
hoelzl@55710
   117
hoelzl@55710
   118
lemma bdd_above_insert [simp]: "bdd_above (insert a A) = bdd_above A"
hoelzl@55710
   119
  by (auto simp: bdd_above_def intro: le_supI2 sup_ge1)
hoelzl@55710
   120
hoelzl@55710
   121
lemma bdd_below_insert [simp]: "bdd_below (insert a A) = bdd_below A"
hoelzl@55710
   122
  by (auto simp: bdd_below_def intro: le_infI2 inf_le1)
hoelzl@55710
   123
hoelzl@55710
   124
lemma bdd_finite [simp]:
hoelzl@55710
   125
  assumes "finite A" shows bdd_above_finite: "bdd_above A" and bdd_below_finite: "bdd_below A"
hoelzl@55710
   126
  using assms by (induct rule: finite_induct, auto)
hoelzl@55710
   127
hoelzl@55710
   128
lemma bdd_above_Un [simp]: "bdd_above (A \<union> B) = (bdd_above A \<and> bdd_above B)"
hoelzl@55710
   129
proof
hoelzl@55710
   130
  assume "bdd_above (A \<union> B)"
hoelzl@55710
   131
  thus "bdd_above A \<and> bdd_above B" unfolding bdd_above_def by auto
hoelzl@55710
   132
next
hoelzl@55710
   133
  assume "bdd_above A \<and> bdd_above B"
hoelzl@55710
   134
  then obtain a b where "\<forall>x\<in>A. x \<le> a" "\<forall>x\<in>B. x \<le> b" unfolding bdd_above_def by auto
hoelzl@55710
   135
  hence "\<forall>x \<in> A \<union> B. x \<le> sup a b" by (auto intro: Un_iff le_supI1 le_supI2)
hoelzl@55710
   136
  thus "bdd_above (A \<union> B)" unfolding bdd_above_def ..
hoelzl@55710
   137
qed
hoelzl@55710
   138
hoelzl@55710
   139
lemma bdd_below_Un [simp]: "bdd_below (A \<union> B) = (bdd_below A \<and> bdd_below B)"
hoelzl@55710
   140
proof
hoelzl@55710
   141
  assume "bdd_below (A \<union> B)"
hoelzl@55710
   142
  thus "bdd_below A \<and> bdd_below B" unfolding bdd_below_def by auto
hoelzl@55710
   143
next
hoelzl@55710
   144
  assume "bdd_below A \<and> bdd_below B"
hoelzl@55710
   145
  then obtain a b where "\<forall>x\<in>A. a \<le> x" "\<forall>x\<in>B. b \<le> x" unfolding bdd_below_def by auto
hoelzl@55710
   146
  hence "\<forall>x \<in> A \<union> B. inf a b \<le> x" by (auto intro: Un_iff le_infI1 le_infI2)
hoelzl@55710
   147
  thus "bdd_below (A \<union> B)" unfolding bdd_below_def ..
hoelzl@55710
   148
qed
hoelzl@55710
   149
hoelzl@55711
   150
lemma bdd_above_sup[simp]: "bdd_above ((\<lambda>x. sup (f x) (g x)) ` A) \<longleftrightarrow> bdd_above (f`A) \<and> bdd_above (g`A)"
hoelzl@55711
   151
  by (auto simp: bdd_above_def intro: le_supI1 le_supI2)
hoelzl@55711
   152
hoelzl@55711
   153
lemma bdd_below_inf[simp]: "bdd_below ((\<lambda>x. inf (f x) (g x)) ` A) \<longleftrightarrow> bdd_below (f`A) \<and> bdd_below (g`A)"
hoelzl@55711
   154
  by (auto simp: bdd_below_def intro: le_infI1 le_infI2)
hoelzl@55711
   155
hoelzl@55710
   156
end
hoelzl@55710
   157
hoelzl@55710
   158
hoelzl@52612
   159
text {*
hoelzl@52612
   160
hoelzl@52612
   161
To avoid name classes with the @{class complete_lattice}-class we prefix @{const Sup} and
hoelzl@52612
   162
@{const Inf} in theorem names with c.
hoelzl@52612
   163
hoelzl@52612
   164
*}
hoelzl@52612
   165
hoelzl@52910
   166
class conditionally_complete_lattice = lattice + Sup + Inf +
hoelzl@55710
   167
  assumes cInf_lower: "x \<in> X \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X \<le> x"
hoelzl@52612
   168
    and cInf_greatest: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> z \<le> Inf X"
hoelzl@55710
   169
  assumes cSup_upper: "x \<in> X \<Longrightarrow> bdd_above X \<Longrightarrow> x \<le> Sup X"
hoelzl@52612
   170
    and cSup_least: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X \<le> z"
paulson@33269
   171
begin
paulson@33269
   172
hoelzl@55711
   173
lemma cSup_upper2: "x \<in> X \<Longrightarrow> y \<le> x \<Longrightarrow> bdd_above X \<Longrightarrow> y \<le> Sup X"
hoelzl@55711
   174
  by (metis cSup_upper order_trans)
hoelzl@55711
   175
hoelzl@55711
   176
lemma cInf_lower2: "x \<in> X \<Longrightarrow> x \<le> y \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X \<le> y"
hoelzl@55711
   177
  by (metis cInf_lower order_trans)
hoelzl@55711
   178
hoelzl@55711
   179
lemma cSup_mono: "B \<noteq> {} \<Longrightarrow> bdd_above A \<Longrightarrow> (\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. b \<le> a) \<Longrightarrow> Sup B \<le> Sup A"
hoelzl@55711
   180
  by (metis cSup_least cSup_upper2)
hoelzl@55711
   181
hoelzl@55711
   182
lemma cInf_mono: "B \<noteq> {} \<Longrightarrow> bdd_below A \<Longrightarrow> (\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. a \<le> b) \<Longrightarrow> Inf A \<le> Inf B"
hoelzl@55711
   183
  by (metis cInf_greatest cInf_lower2)
hoelzl@55711
   184
hoelzl@55711
   185
lemma cSup_subset_mono: "A \<noteq> {} \<Longrightarrow> bdd_above B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Sup A \<le> Sup B"
hoelzl@55711
   186
  by (metis cSup_least cSup_upper subsetD)
hoelzl@55711
   187
hoelzl@55711
   188
lemma cInf_superset_mono: "A \<noteq> {} \<Longrightarrow> bdd_below B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Inf B \<le> Inf A"
hoelzl@55711
   189
  by (metis cInf_greatest cInf_lower subsetD)
hoelzl@55711
   190
hoelzl@55710
   191
lemma cSup_eq_maximum: "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X = z"
hoelzl@55710
   192
  by (intro antisym cSup_upper[of z X] cSup_least[of X z]) auto
hoelzl@52612
   193
hoelzl@55710
   194
lemma cInf_eq_minimum: "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X = z"
hoelzl@55710
   195
  by (intro antisym cInf_lower[of z X] cInf_greatest[of X z]) auto
hoelzl@52612
   196
hoelzl@55710
   197
lemma cSup_le_iff: "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> Sup S \<le> a \<longleftrightarrow> (\<forall>x\<in>S. x \<le> a)"
hoelzl@52612
   198
  by (metis order_trans cSup_upper cSup_least)
hoelzl@52612
   199
hoelzl@55710
   200
lemma le_cInf_iff: "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> a \<le> Inf S \<longleftrightarrow> (\<forall>x\<in>S. a \<le> x)"
hoelzl@52612
   201
  by (metis order_trans cInf_lower cInf_greatest)
hoelzl@52612
   202
hoelzl@52612
   203
lemma cSup_eq_non_empty:
hoelzl@52612
   204
  assumes 1: "X \<noteq> {}"
hoelzl@52612
   205
  assumes 2: "\<And>x. x \<in> X \<Longrightarrow> x \<le> a"
hoelzl@52612
   206
  assumes 3: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> x \<le> y) \<Longrightarrow> a \<le> y"
hoelzl@52612
   207
  shows "Sup X = a"
hoelzl@52612
   208
  by (intro 3 1 antisym cSup_least) (auto intro: 2 1 cSup_upper)
hoelzl@52612
   209
hoelzl@52612
   210
lemma cInf_eq_non_empty:
hoelzl@52612
   211
  assumes 1: "X \<noteq> {}"
hoelzl@52612
   212
  assumes 2: "\<And>x. x \<in> X \<Longrightarrow> a \<le> x"
hoelzl@52612
   213
  assumes 3: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> y \<le> x) \<Longrightarrow> y \<le> a"
hoelzl@52612
   214
  shows "Inf X = a"
hoelzl@52612
   215
  by (intro 3 1 antisym cInf_greatest) (auto intro: 2 1 cInf_lower)
hoelzl@52612
   216
hoelzl@55710
   217
lemma cInf_cSup: "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> Inf S = Sup {x. \<forall>s\<in>S. x \<le> s}"
hoelzl@55710
   218
  by (rule cInf_eq_non_empty) (auto intro!: cSup_upper cSup_least simp: bdd_below_def)
hoelzl@52655
   219
hoelzl@55710
   220
lemma cSup_cInf: "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> Sup S = Inf {x. \<forall>s\<in>S. s \<le> x}"
hoelzl@55710
   221
  by (rule cSup_eq_non_empty) (auto intro!: cInf_lower cInf_greatest simp: bdd_above_def)
hoelzl@52655
   222
hoelzl@55710
   223
lemma cSup_insert: "X \<noteq> {} \<Longrightarrow> bdd_above X \<Longrightarrow> Sup (insert a X) = sup a (Sup X)"
hoelzl@55710
   224
  by (intro cSup_eq_non_empty) (auto intro: le_supI2 cSup_upper cSup_least)
hoelzl@52612
   225
hoelzl@55710
   226
lemma cInf_insert: "X \<noteq> {} \<Longrightarrow> bdd_below X \<Longrightarrow> Inf (insert a X) = inf a (Inf X)"
hoelzl@55710
   227
  by (intro cInf_eq_non_empty) (auto intro: le_infI2 cInf_lower cInf_greatest)
hoelzl@52612
   228
hoelzl@55711
   229
lemma cSup_singleton [simp]: "Sup {x} = x"
hoelzl@55711
   230
  by (intro cSup_eq_maximum) auto
hoelzl@55711
   231
hoelzl@55711
   232
lemma cInf_singleton [simp]: "Inf {x} = x"
hoelzl@55711
   233
  by (intro cInf_eq_minimum) auto
hoelzl@55711
   234
hoelzl@55710
   235
lemma cSup_insert_If:  "bdd_above X \<Longrightarrow> Sup (insert a X) = (if X = {} then a else sup a (Sup X))"
hoelzl@55710
   236
  using cSup_insert[of X] by simp
hoelzl@52612
   237
hoelzl@55711
   238
lemma cInf_insert_If: "bdd_below X \<Longrightarrow> Inf (insert a X) = (if X = {} then a else inf a (Inf X))"
hoelzl@55710
   239
  using cInf_insert[of X] by simp
hoelzl@52612
   240
hoelzl@52612
   241
lemma le_cSup_finite: "finite X \<Longrightarrow> x \<in> X \<Longrightarrow> x \<le> Sup X"
hoelzl@52612
   242
proof (induct X arbitrary: x rule: finite_induct)
hoelzl@52612
   243
  case (insert x X y) then show ?case
hoelzl@55710
   244
    by (cases "X = {}") (auto simp: cSup_insert intro: le_supI2)
hoelzl@52612
   245
qed simp
hoelzl@52612
   246
hoelzl@52612
   247
lemma cInf_le_finite: "finite X \<Longrightarrow> x \<in> X \<Longrightarrow> Inf X \<le> x"
hoelzl@52612
   248
proof (induct X arbitrary: x rule: finite_induct)
hoelzl@52612
   249
  case (insert x X y) then show ?case
hoelzl@55710
   250
    by (cases "X = {}") (auto simp: cInf_insert intro: le_infI2)
hoelzl@52612
   251
qed simp
hoelzl@52612
   252
hoelzl@52612
   253
lemma cSup_eq_Sup_fin: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Sup_fin X"
hoelzl@55710
   254
  by (induct X rule: finite_ne_induct) (simp_all add: cSup_insert)
hoelzl@52612
   255
hoelzl@52612
   256
lemma cInf_eq_Inf_fin: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Inf_fin X"
hoelzl@55710
   257
  by (induct X rule: finite_ne_induct) (simp_all add: cInf_insert)
hoelzl@52612
   258
hoelzl@52612
   259
lemma cSup_atMost[simp]: "Sup {..x} = x"
hoelzl@52612
   260
  by (auto intro!: cSup_eq_maximum)
hoelzl@52612
   261
hoelzl@52612
   262
lemma cSup_greaterThanAtMost[simp]: "y < x \<Longrightarrow> Sup {y<..x} = x"
hoelzl@52612
   263
  by (auto intro!: cSup_eq_maximum)
hoelzl@52612
   264
hoelzl@52612
   265
lemma cSup_atLeastAtMost[simp]: "y \<le> x \<Longrightarrow> Sup {y..x} = x"
hoelzl@52612
   266
  by (auto intro!: cSup_eq_maximum)
hoelzl@52612
   267
hoelzl@52612
   268
lemma cInf_atLeast[simp]: "Inf {x..} = x"
hoelzl@52612
   269
  by (auto intro!: cInf_eq_minimum)
hoelzl@52612
   270
hoelzl@52612
   271
lemma cInf_atLeastLessThan[simp]: "y < x \<Longrightarrow> Inf {y..<x} = y"
hoelzl@52612
   272
  by (auto intro!: cInf_eq_minimum)
hoelzl@52612
   273
hoelzl@52612
   274
lemma cInf_atLeastAtMost[simp]: "y \<le> x \<Longrightarrow> Inf {y..x} = y"
hoelzl@52612
   275
  by (auto intro!: cInf_eq_minimum)
hoelzl@52612
   276
hoelzl@55711
   277
lemma cINF_lower: "bdd_below (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> INFI A f \<le> f x"
hoelzl@55711
   278
  unfolding INF_def by (rule cInf_lower) auto
hoelzl@55711
   279
hoelzl@55711
   280
lemma cINF_greatest: "A \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> m \<le> f x) \<Longrightarrow> m \<le> INFI A f"
hoelzl@55711
   281
  unfolding INF_def by (rule cInf_greatest) auto
hoelzl@55711
   282
hoelzl@55711
   283
lemma cSUP_upper: "x \<in> A \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> f x \<le> SUPR A f"
hoelzl@55711
   284
  unfolding SUP_def by (rule cSup_upper) auto
hoelzl@55711
   285
hoelzl@55711
   286
lemma cSUP_least: "A \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<le> M) \<Longrightarrow> SUPR A f \<le> M"
hoelzl@55711
   287
  unfolding SUP_def by (rule cSup_least) auto
hoelzl@55711
   288
hoelzl@55711
   289
lemma cINF_lower2: "bdd_below (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> f x \<le> u \<Longrightarrow> INFI A f \<le> u"
hoelzl@55711
   290
  by (auto intro: cINF_lower assms order_trans)
hoelzl@55711
   291
hoelzl@55711
   292
lemma cSUP_upper2: "bdd_above (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> u \<le> f x \<Longrightarrow> u \<le> SUPR A f"
hoelzl@55711
   293
  by (auto intro: cSUP_upper assms order_trans)
hoelzl@55711
   294
hoelzl@55713
   295
lemma cSUP_const: "A \<noteq> {} \<Longrightarrow> (SUP x:A. c) = c"
hoelzl@55713
   296
  by (intro antisym cSUP_least) (auto intro: cSUP_upper)
hoelzl@55713
   297
hoelzl@55713
   298
lemma cINF_const: "A \<noteq> {} \<Longrightarrow> (INF x:A. c) = c"
hoelzl@55713
   299
  by (intro antisym cINF_greatest) (auto intro: cINF_lower)
hoelzl@55713
   300
hoelzl@55711
   301
lemma le_cINF_iff: "A \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> u \<le> INFI A f \<longleftrightarrow> (\<forall>x\<in>A. u \<le> f x)"
hoelzl@55711
   302
  by (metis cINF_greatest cINF_lower assms order_trans)
hoelzl@55711
   303
hoelzl@55711
   304
lemma cSUP_le_iff: "A \<noteq> {} \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> SUPR A f \<le> u \<longleftrightarrow> (\<forall>x\<in>A. f x \<le> u)"
hoelzl@55711
   305
  by (metis cSUP_least cSUP_upper assms order_trans)
hoelzl@55711
   306
hoelzl@55715
   307
lemma less_cINF_D: "bdd_below (f`A) \<Longrightarrow> y < (INF i:A. f i) \<Longrightarrow> i \<in> A \<Longrightarrow> y < f i"
hoelzl@55715
   308
  by (metis cINF_lower less_le_trans)
hoelzl@55715
   309
hoelzl@55715
   310
lemma cSUP_lessD: "bdd_above (f`A) \<Longrightarrow> (SUP i:A. f i) < y \<Longrightarrow> i \<in> A \<Longrightarrow> f i < y"
hoelzl@55715
   311
  by (metis cSUP_upper le_less_trans)
hoelzl@55715
   312
hoelzl@55711
   313
lemma cINF_insert: "A \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> INFI (insert a A) f = inf (f a) (INFI A f)"
hoelzl@55711
   314
  by (metis INF_def cInf_insert assms empty_is_image image_insert)
hoelzl@55711
   315
hoelzl@55711
   316
lemma cSUP_insert: "A \<noteq> {} \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> SUPR (insert a A) f = sup (f a) (SUPR A f)"
hoelzl@55711
   317
  by (metis SUP_def cSup_insert assms empty_is_image image_insert)
hoelzl@55711
   318
hoelzl@55711
   319
lemma cINF_mono: "B \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> (\<And>m. m \<in> B \<Longrightarrow> \<exists>n\<in>A. f n \<le> g m) \<Longrightarrow> INFI A f \<le> INFI B g"
hoelzl@55711
   320
  unfolding INF_def by (auto intro: cInf_mono)
hoelzl@55711
   321
hoelzl@55711
   322
lemma cSUP_mono: "A \<noteq> {} \<Longrightarrow> bdd_above (g ` B) \<Longrightarrow> (\<And>n. n \<in> A \<Longrightarrow> \<exists>m\<in>B. f n \<le> g m) \<Longrightarrow> SUPR A f \<le> SUPR B g"
hoelzl@55711
   323
  unfolding SUP_def by (auto intro: cSup_mono)
hoelzl@55711
   324
hoelzl@55711
   325
lemma cINF_superset_mono: "A \<noteq> {} \<Longrightarrow> bdd_below (g ` B) \<Longrightarrow> A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> g x \<le> f x) \<Longrightarrow> INFI B g \<le> INFI A f"
hoelzl@55711
   326
  by (rule cINF_mono) auto
hoelzl@55711
   327
hoelzl@55711
   328
lemma cSUP_subset_mono: "A \<noteq> {} \<Longrightarrow> bdd_above (g ` B) \<Longrightarrow> A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> f x \<le> g x) \<Longrightarrow> SUPR A f \<le> SUPR B g"
hoelzl@55711
   329
  by (rule cSUP_mono) auto
hoelzl@55711
   330
hoelzl@55711
   331
lemma less_eq_cInf_inter: "bdd_below A \<Longrightarrow> bdd_below B \<Longrightarrow> A \<inter> B \<noteq> {} \<Longrightarrow> inf (Inf A) (Inf B) \<le> Inf (A \<inter> B)"
hoelzl@55711
   332
  by (metis cInf_superset_mono lattice_class.inf_sup_ord(1) le_infI1)
hoelzl@55711
   333
hoelzl@55711
   334
lemma cSup_inter_less_eq: "bdd_above A \<Longrightarrow> bdd_above B \<Longrightarrow> A \<inter> B \<noteq> {} \<Longrightarrow> Sup (A \<inter> B) \<le> sup (Sup A) (Sup B) "
hoelzl@55711
   335
  by (metis cSup_subset_mono lattice_class.inf_sup_ord(1) le_supI1)
hoelzl@55711
   336
hoelzl@55711
   337
lemma cInf_union_distrib: "A \<noteq> {} \<Longrightarrow> bdd_below A \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_below B \<Longrightarrow> Inf (A \<union> B) = inf (Inf A) (Inf B)"
hoelzl@55711
   338
  by (intro antisym le_infI cInf_greatest cInf_lower) (auto intro: le_infI1 le_infI2 cInf_lower)
hoelzl@55711
   339
hoelzl@55711
   340
lemma cINF_union: "A \<noteq> {} \<Longrightarrow> bdd_below (f`A) \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_below (f`B) \<Longrightarrow> INFI (A \<union> B) f = inf (INFI A f) (INFI B f)"
hoelzl@55711
   341
  unfolding INF_def using assms by (auto simp add: image_Un intro: cInf_union_distrib)
hoelzl@55711
   342
hoelzl@55711
   343
lemma cSup_union_distrib: "A \<noteq> {} \<Longrightarrow> bdd_above A \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_above B \<Longrightarrow> Sup (A \<union> B) = sup (Sup A) (Sup B)"
hoelzl@55711
   344
  by (intro antisym le_supI cSup_least cSup_upper) (auto intro: le_supI1 le_supI2 cSup_upper)
hoelzl@55711
   345
hoelzl@55711
   346
lemma cSUP_union: "A \<noteq> {} \<Longrightarrow> bdd_above (f`A) \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_above (f`B) \<Longrightarrow> SUPR (A \<union> B) f = sup (SUPR A f) (SUPR B f)"
hoelzl@55711
   347
  unfolding SUP_def by (auto simp add: image_Un intro: cSup_union_distrib)
hoelzl@55711
   348
hoelzl@55711
   349
lemma cINF_inf_distrib: "A \<noteq> {} \<Longrightarrow> bdd_below (f`A) \<Longrightarrow> bdd_below (g`A) \<Longrightarrow> inf (INFI A f) (INFI A g) = (INF a:A. inf (f a) (g a))"
hoelzl@55711
   350
  by (intro antisym le_infI cINF_greatest cINF_lower2)
hoelzl@55711
   351
     (auto intro: le_infI1 le_infI2 cINF_greatest cINF_lower le_infI)
hoelzl@55711
   352
hoelzl@55711
   353
lemma SUP_sup_distrib: "A \<noteq> {} \<Longrightarrow> bdd_above (f`A) \<Longrightarrow> bdd_above (g`A) \<Longrightarrow> sup (SUPR A f) (SUPR A g) = (SUP a:A. sup (f a) (g a))"
hoelzl@55711
   354
  by (intro antisym le_supI cSUP_least cSUP_upper2)
hoelzl@55711
   355
     (auto intro: le_supI1 le_supI2 cSUP_least cSUP_upper le_supI)
hoelzl@55711
   356
paulson@33269
   357
end
paulson@33269
   358
hoelzl@52910
   359
instance complete_lattice \<subseteq> conditionally_complete_lattice
hoelzl@52612
   360
  by default (auto intro: Sup_upper Sup_least Inf_lower Inf_greatest)
hoelzl@52612
   361
hoelzl@52612
   362
lemma cSup_eq:
hoelzl@52910
   363
  fixes a :: "'a :: {conditionally_complete_lattice, no_bot}"
hoelzl@52612
   364
  assumes upper: "\<And>x. x \<in> X \<Longrightarrow> x \<le> a"
hoelzl@52612
   365
  assumes least: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> x \<le> y) \<Longrightarrow> a \<le> y"
hoelzl@52612
   366
  shows "Sup X = a"
hoelzl@52612
   367
proof cases
hoelzl@52612
   368
  assume "X = {}" with lt_ex[of a] least show ?thesis by (auto simp: less_le_not_le)
hoelzl@52612
   369
qed (intro cSup_eq_non_empty assms)
hoelzl@52612
   370
hoelzl@52612
   371
lemma cInf_eq:
hoelzl@52910
   372
  fixes a :: "'a :: {conditionally_complete_lattice, no_top}"
hoelzl@52612
   373
  assumes upper: "\<And>x. x \<in> X \<Longrightarrow> a \<le> x"
hoelzl@52612
   374
  assumes least: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> y \<le> x) \<Longrightarrow> y \<le> a"
hoelzl@52612
   375
  shows "Inf X = a"
hoelzl@52612
   376
proof cases
hoelzl@52612
   377
  assume "X = {}" with gt_ex[of a] least show ?thesis by (auto simp: less_le_not_le)
hoelzl@52612
   378
qed (intro cInf_eq_non_empty assms)
hoelzl@52612
   379
hoelzl@52910
   380
class conditionally_complete_linorder = conditionally_complete_lattice + linorder
paulson@33269
   381
begin
hoelzl@52612
   382
hoelzl@52612
   383
lemma less_cSup_iff : (*REAL_SUP_LE in HOL4*)
hoelzl@55710
   384
  "X \<noteq> {} \<Longrightarrow> bdd_above X \<Longrightarrow> y < Sup X \<longleftrightarrow> (\<exists>x\<in>X. y < x)"
hoelzl@52612
   385
  by (rule iffI) (metis cSup_least not_less, metis cSup_upper less_le_trans)
hoelzl@52612
   386
hoelzl@55710
   387
lemma cInf_less_iff: "X \<noteq> {} \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X < y \<longleftrightarrow> (\<exists>x\<in>X. x < y)"
hoelzl@52612
   388
  by (rule iffI) (metis cInf_greatest not_less, metis cInf_lower le_less_trans)
hoelzl@52612
   389
hoelzl@55715
   390
lemma cINF_less_iff: "A \<noteq> {} \<Longrightarrow> bdd_below (f`A) \<Longrightarrow> (INF i:A. f i) < a \<longleftrightarrow> (\<exists>x\<in>A. f x < a)"
hoelzl@55715
   391
  unfolding INF_def using cInf_less_iff[of "f`A"] by auto
hoelzl@55715
   392
hoelzl@55715
   393
lemma less_cSUP_iff: "A \<noteq> {} \<Longrightarrow> bdd_above (f`A) \<Longrightarrow> a < (SUP i:A. f i) \<longleftrightarrow> (\<exists>x\<in>A. a < f x)"
hoelzl@55715
   394
  unfolding SUP_def using less_cSup_iff[of "f`A"] by auto
hoelzl@55715
   395
hoelzl@52612
   396
lemma less_cSupE:
hoelzl@52612
   397
  assumes "y < Sup X" "X \<noteq> {}" obtains x where "x \<in> X" "y < x"
hoelzl@52612
   398
  by (metis cSup_least assms not_le that)
hoelzl@52612
   399
hoelzl@52655
   400
lemma less_cSupD:
hoelzl@52655
   401
  "X \<noteq> {} \<Longrightarrow> z < Sup X \<Longrightarrow> \<exists>x\<in>X. z < x"
hoelzl@55710
   402
  by (metis less_cSup_iff not_leE bdd_above_def)
hoelzl@52655
   403
hoelzl@52655
   404
lemma cInf_lessD:
hoelzl@52655
   405
  "X \<noteq> {} \<Longrightarrow> Inf X < z \<Longrightarrow> \<exists>x\<in>X. x < z"
hoelzl@55710
   406
  by (metis cInf_less_iff not_leE bdd_below_def)
hoelzl@52655
   407
hoelzl@52612
   408
lemma complete_interval:
hoelzl@52612
   409
  assumes "a < b" and "P a" and "\<not> P b"
hoelzl@52612
   410
  shows "\<exists>c. a \<le> c \<and> c \<le> b \<and> (\<forall>x. a \<le> x \<and> x < c \<longrightarrow> P x) \<and>
hoelzl@52612
   411
             (\<forall>d. (\<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x) \<longrightarrow> d \<le> c)"
hoelzl@52612
   412
proof (rule exI [where x = "Sup {d. \<forall>x. a \<le> x & x < d --> P x}"], auto)
hoelzl@52612
   413
  show "a \<le> Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
hoelzl@55710
   414
    by (rule cSup_upper, auto simp: bdd_above_def)
hoelzl@52612
   415
       (metis `a < b` `\<not> P b` linear less_le)
hoelzl@52612
   416
next
hoelzl@52612
   417
  show "Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c} \<le> b"
hoelzl@52612
   418
    apply (rule cSup_least) 
hoelzl@52612
   419
    apply auto
hoelzl@52612
   420
    apply (metis less_le_not_le)
hoelzl@52612
   421
    apply (metis `a<b` `~ P b` linear less_le)
hoelzl@52612
   422
    done
hoelzl@52612
   423
next
hoelzl@52612
   424
  fix x
hoelzl@52612
   425
  assume x: "a \<le> x" and lt: "x < Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
hoelzl@52612
   426
  show "P x"
hoelzl@52612
   427
    apply (rule less_cSupE [OF lt], auto)
hoelzl@52612
   428
    apply (metis less_le_not_le)
hoelzl@52612
   429
    apply (metis x) 
hoelzl@52612
   430
    done
hoelzl@52612
   431
next
hoelzl@52612
   432
  fix d
hoelzl@52612
   433
    assume 0: "\<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x"
hoelzl@52612
   434
    thus "d \<le> Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
hoelzl@55710
   435
      by (rule_tac cSup_upper, auto simp: bdd_above_def)
hoelzl@52612
   436
         (metis `a<b` `~ P b` linear less_le)
hoelzl@52612
   437
qed
hoelzl@52612
   438
hoelzl@52612
   439
end
hoelzl@52612
   440
hoelzl@55711
   441
lemma cSup_eq_Max: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Max X"
hoelzl@55711
   442
  using cSup_eq_Sup_fin[of X] Sup_fin_eq_Max[of X] by simp
hoelzl@52912
   443
hoelzl@55711
   444
lemma cInf_eq_Min: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Min X"
hoelzl@55711
   445
  using cInf_eq_Inf_fin[of X] Inf_fin_eq_Min[of X] by simp
hoelzl@52912
   446
hoelzl@55709
   447
lemma cSup_lessThan[simp]: "Sup {..<x::'a::{conditionally_complete_linorder, no_bot, dense_linorder}} = x"
hoelzl@52612
   448
  by (auto intro!: cSup_eq_non_empty intro: dense_le)
hoelzl@52612
   449
hoelzl@55709
   450
lemma cSup_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Sup {y<..<x::'a::{conditionally_complete_linorder, no_bot, dense_linorder}} = x"
hoelzl@52612
   451
  by (auto intro!: cSup_eq intro: dense_le_bounded)
hoelzl@52612
   452
hoelzl@55709
   453
lemma cSup_atLeastLessThan[simp]: "y < x \<Longrightarrow> Sup {y..<x::'a::{conditionally_complete_linorder, no_bot, dense_linorder}} = x"
hoelzl@52612
   454
  by (auto intro!: cSup_eq intro: dense_le_bounded)
hoelzl@52612
   455
hoelzl@55709
   456
lemma cInf_greaterThan[simp]: "Inf {x::'a::{conditionally_complete_linorder, no_top, dense_linorder} <..} = x"
hoelzl@52612
   457
  by (auto intro!: cInf_eq intro: dense_ge)
hoelzl@52612
   458
hoelzl@55709
   459
lemma cInf_greaterThanAtMost[simp]: "y < x \<Longrightarrow> Inf {y<..x::'a::{conditionally_complete_linorder, no_top, dense_linorder}} = y"
hoelzl@52612
   460
  by (auto intro!: cInf_eq intro: dense_ge_bounded)
hoelzl@52612
   461
hoelzl@55709
   462
lemma cInf_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Inf {y<..<x::'a::{conditionally_complete_linorder, no_top, dense_linorder}} = y"
hoelzl@52612
   463
  by (auto intro!: cInf_eq intro: dense_ge_bounded)
hoelzl@52612
   464
hoelzl@55711
   465
class linear_continuum = conditionally_complete_linorder + dense_linorder +
hoelzl@55711
   466
  assumes UNIV_not_singleton: "\<exists>a b::'a. a \<noteq> b"
hoelzl@55711
   467
begin
hoelzl@55711
   468
hoelzl@55711
   469
lemma ex_gt_or_lt: "\<exists>b. a < b \<or> b < a"
hoelzl@55711
   470
  by (metis UNIV_not_singleton neq_iff)
hoelzl@55711
   471
paulson@33269
   472
end
hoelzl@55711
   473
hoelzl@55711
   474
end