src/HOL/Conditionally_Complete_Lattices.thy
changeset 73411 1f1366966296
parent 73271 05a873f90655
child 74007 df976eefcba0
equal deleted inserted replaced
73410:7b59d2945e54 73411:1f1366966296
   255 
   255 
   256 lemma cInf_superset_mono: "A \<noteq> {} \<Longrightarrow> bdd_below B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Inf B \<le> Inf A"
   256 lemma cInf_superset_mono: "A \<noteq> {} \<Longrightarrow> bdd_below B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Inf B \<le> Inf A"
   257   by (metis cInf_greatest cInf_lower subsetD)
   257   by (metis cInf_greatest cInf_lower subsetD)
   258 
   258 
   259 lemma cSup_eq_maximum: "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X = z"
   259 lemma cSup_eq_maximum: "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X = z"
   260   by (intro antisym cSup_upper[of z X] cSup_least[of X z]) auto
   260   by (intro order.antisym cSup_upper[of z X] cSup_least[of X z]) auto
   261 
   261 
   262 lemma cInf_eq_minimum: "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X = z"
   262 lemma cInf_eq_minimum: "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X = z"
   263   by (intro antisym cInf_lower[of z X] cInf_greatest[of X z]) auto
   263   by (intro order.antisym cInf_lower[of z X] cInf_greatest[of X z]) auto
   264 
   264 
   265 lemma cSup_le_iff: "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> Sup S \<le> a \<longleftrightarrow> (\<forall>x\<in>S. x \<le> a)"
   265 lemma cSup_le_iff: "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> Sup S \<le> a \<longleftrightarrow> (\<forall>x\<in>S. x \<le> a)"
   266   by (metis order_trans cSup_upper cSup_least)
   266   by (metis order_trans cSup_upper cSup_least)
   267 
   267 
   268 lemma le_cInf_iff: "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> a \<le> Inf S \<longleftrightarrow> (\<forall>x\<in>S. a \<le> x)"
   268 lemma le_cInf_iff: "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> a \<le> Inf S \<longleftrightarrow> (\<forall>x\<in>S. a \<le> x)"
   271 lemma cSup_eq_non_empty:
   271 lemma cSup_eq_non_empty:
   272   assumes 1: "X \<noteq> {}"
   272   assumes 1: "X \<noteq> {}"
   273   assumes 2: "\<And>x. x \<in> X \<Longrightarrow> x \<le> a"
   273   assumes 2: "\<And>x. x \<in> X \<Longrightarrow> x \<le> a"
   274   assumes 3: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> x \<le> y) \<Longrightarrow> a \<le> y"
   274   assumes 3: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> x \<le> y) \<Longrightarrow> a \<le> y"
   275   shows "Sup X = a"
   275   shows "Sup X = a"
   276   by (intro 3 1 antisym cSup_least) (auto intro: 2 1 cSup_upper)
   276   by (intro 3 1 order.antisym cSup_least) (auto intro: 2 1 cSup_upper)
   277 
   277 
   278 lemma cInf_eq_non_empty:
   278 lemma cInf_eq_non_empty:
   279   assumes 1: "X \<noteq> {}"
   279   assumes 1: "X \<noteq> {}"
   280   assumes 2: "\<And>x. x \<in> X \<Longrightarrow> a \<le> x"
   280   assumes 2: "\<And>x. x \<in> X \<Longrightarrow> a \<le> x"
   281   assumes 3: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> y \<le> x) \<Longrightarrow> y \<le> a"
   281   assumes 3: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> y \<le> x) \<Longrightarrow> y \<le> a"
   282   shows "Inf X = a"
   282   shows "Inf X = a"
   283   by (intro 3 1 antisym cInf_greatest) (auto intro: 2 1 cInf_lower)
   283   by (intro 3 1 order.antisym cInf_greatest) (auto intro: 2 1 cInf_lower)
   284 
   284 
   285 lemma cInf_cSup: "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> Inf S = Sup {x. \<forall>s\<in>S. x \<le> s}"
   285 lemma cInf_cSup: "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> Inf S = Sup {x. \<forall>s\<in>S. x \<le> s}"
   286   by (rule cInf_eq_non_empty) (auto intro!: cSup_upper cSup_least simp: bdd_below_def)
   286   by (rule cInf_eq_non_empty) (auto intro!: cSup_upper cSup_least simp: bdd_below_def)
   287 
   287 
   288 lemma cSup_cInf: "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> Sup S = Inf {x. \<forall>s\<in>S. s \<le> x}"
   288 lemma cSup_cInf: "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> Sup S = Inf {x. \<forall>s\<in>S. s \<le> x}"
   359 
   359 
   360 lemma cSUP_upper2: "bdd_above (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> u \<le> f x \<Longrightarrow> u \<le> \<Squnion>(f ` A)"
   360 lemma cSUP_upper2: "bdd_above (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> u \<le> f x \<Longrightarrow> u \<le> \<Squnion>(f ` A)"
   361   by (auto intro: cSUP_upper order_trans)
   361   by (auto intro: cSUP_upper order_trans)
   362 
   362 
   363 lemma cSUP_const [simp]: "A \<noteq> {} \<Longrightarrow> (\<Squnion>x\<in>A. c) = c"
   363 lemma cSUP_const [simp]: "A \<noteq> {} \<Longrightarrow> (\<Squnion>x\<in>A. c) = c"
   364   by (intro antisym cSUP_least) (auto intro: cSUP_upper)
   364   by (intro order.antisym cSUP_least) (auto intro: cSUP_upper)
   365 
   365 
   366 lemma cINF_const [simp]: "A \<noteq> {} \<Longrightarrow> (\<Sqinter>x\<in>A. c) = c"
   366 lemma cINF_const [simp]: "A \<noteq> {} \<Longrightarrow> (\<Sqinter>x\<in>A. c) = c"
   367   by (intro antisym cINF_greatest) (auto intro: cINF_lower)
   367   by (intro order.antisym cINF_greatest) (auto intro: cINF_lower)
   368 
   368 
   369 lemma le_cINF_iff: "A \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> u \<le> \<Sqinter>(f ` A) \<longleftrightarrow> (\<forall>x\<in>A. u \<le> f x)"
   369 lemma le_cINF_iff: "A \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> u \<le> \<Sqinter>(f ` A) \<longleftrightarrow> (\<forall>x\<in>A. u \<le> f x)"
   370   by (metis cINF_greatest cINF_lower order_trans)
   370   by (metis cINF_greatest cINF_lower order_trans)
   371 
   371 
   372 lemma cSUP_le_iff: "A \<noteq> {} \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> \<Squnion>(f ` A) \<le> u \<longleftrightarrow> (\<forall>x\<in>A. f x \<le> u)"
   372 lemma cSUP_le_iff: "A \<noteq> {} \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> \<Squnion>(f ` A) \<le> u \<longleftrightarrow> (\<forall>x\<in>A. f x \<le> u)"
   402 
   402 
   403 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) "
   403 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) "
   404   by (metis cSup_subset_mono lattice_class.inf_sup_ord(1) le_supI1)
   404   by (metis cSup_subset_mono lattice_class.inf_sup_ord(1) le_supI1)
   405 
   405 
   406 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)"
   406 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)"
   407   by (intro antisym le_infI cInf_greatest cInf_lower) (auto intro: le_infI1 le_infI2 cInf_lower)
   407   by (intro order.antisym le_infI cInf_greatest cInf_lower) (auto intro: le_infI1 le_infI2 cInf_lower)
   408 
   408 
   409 lemma cINF_union: "A \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_below (f ` B) \<Longrightarrow> \<Sqinter> (f ` (A \<union> B)) = \<Sqinter> (f ` A) \<sqinter> \<Sqinter> (f ` B)"
   409 lemma cINF_union: "A \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_below (f ` B) \<Longrightarrow> \<Sqinter> (f ` (A \<union> B)) = \<Sqinter> (f ` A) \<sqinter> \<Sqinter> (f ` B)"
   410   using cInf_union_distrib [of "f ` A" "f ` B"] by (simp add: image_Un)
   410   using cInf_union_distrib [of "f ` A" "f ` B"] by (simp add: image_Un)
   411 
   411 
   412 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)"
   412 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)"
   413   by (intro antisym le_supI cSup_least cSup_upper) (auto intro: le_supI1 le_supI2 cSup_upper)
   413   by (intro order.antisym le_supI cSup_least cSup_upper) (auto intro: le_supI1 le_supI2 cSup_upper)
   414 
   414 
   415 lemma cSUP_union: "A \<noteq> {} \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_above (f ` B) \<Longrightarrow> \<Squnion> (f ` (A \<union> B)) = \<Squnion> (f ` A) \<squnion> \<Squnion> (f ` B)"
   415 lemma cSUP_union: "A \<noteq> {} \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_above (f ` B) \<Longrightarrow> \<Squnion> (f ` (A \<union> B)) = \<Squnion> (f ` A) \<squnion> \<Squnion> (f ` B)"
   416   using cSup_union_distrib [of "f ` A" "f ` B"] by (simp add: image_Un)
   416   using cSup_union_distrib [of "f ` A" "f ` B"] by (simp add: image_Un)
   417 
   417 
   418 lemma cINF_inf_distrib: "A \<noteq> {} \<Longrightarrow> bdd_below (f`A) \<Longrightarrow> bdd_below (g`A) \<Longrightarrow> \<Sqinter> (f ` A) \<sqinter> \<Sqinter> (g ` A) = (\<Sqinter>a\<in>A. inf (f a) (g a))"
   418 lemma cINF_inf_distrib: "A \<noteq> {} \<Longrightarrow> bdd_below (f`A) \<Longrightarrow> bdd_below (g`A) \<Longrightarrow> \<Sqinter> (f ` A) \<sqinter> \<Sqinter> (g ` A) = (\<Sqinter>a\<in>A. inf (f a) (g a))"
   419   by (intro antisym le_infI cINF_greatest cINF_lower2)
   419   by (intro order.antisym le_infI cINF_greatest cINF_lower2)
   420      (auto intro: le_infI1 le_infI2 cINF_greatest cINF_lower le_infI)
   420      (auto intro: le_infI1 le_infI2 cINF_greatest cINF_lower le_infI)
   421 
   421 
   422 lemma SUP_sup_distrib: "A \<noteq> {} \<Longrightarrow> bdd_above (f`A) \<Longrightarrow> bdd_above (g`A) \<Longrightarrow> \<Squnion> (f ` A) \<squnion> \<Squnion> (g ` A) = (\<Squnion>a\<in>A. sup (f a) (g a))"
   422 lemma SUP_sup_distrib: "A \<noteq> {} \<Longrightarrow> bdd_above (f`A) \<Longrightarrow> bdd_above (g`A) \<Longrightarrow> \<Squnion> (f ` A) \<squnion> \<Squnion> (g ` A) = (\<Squnion>a\<in>A. sup (f a) (g a))"
   423   by (intro antisym le_supI cSUP_least cSUP_upper2)
   423   by (intro order.antisym le_supI cSUP_least cSUP_upper2)
   424      (auto intro: le_supI1 le_supI2 cSUP_least cSUP_upper le_supI)
   424      (auto intro: le_supI1 le_supI2 cSUP_least cSUP_upper le_supI)
   425 
   425 
   426 lemma cInf_le_cSup:
   426 lemma cInf_le_cSup:
   427   "A \<noteq> {} \<Longrightarrow> bdd_above A \<Longrightarrow> bdd_below A \<Longrightarrow> Inf A \<le> Sup A"
   427   "A \<noteq> {} \<Longrightarrow> bdd_above A \<Longrightarrow> bdd_below A \<Longrightarrow> Inf A \<le> Sup A"
   428   by (auto intro!: cSup_upper2[of "SOME a. a \<in> A"] intro: someI cInf_lower)
   428   by (auto intro!: cSup_upper2[of "SOME a. a \<in> A"] intro: someI cInf_lower)