src/HOL/Library/Product_Order.thy
changeset 69313 b021008c5397
parent 69260 0a9688695a1b
child 69861 62e47f06d22c
equal deleted inserted replaced
69312:e0f68a507683 69313:b021008c5397
   218 
   218 
   219 text \<open>Alternative formulations for set infima and suprema over the product
   219 text \<open>Alternative formulations for set infima and suprema over the product
   220 of two complete lattices:\<close>
   220 of two complete lattices:\<close>
   221 
   221 
   222 lemma INF_prod_alt_def:
   222 lemma INF_prod_alt_def:
   223   "INFIMUM A f = (INFIMUM A (fst \<circ> f), INFIMUM A (snd \<circ> f))"
   223   "Inf (f ` A) = (Inf ((fst \<circ> f) ` A), Inf ((snd \<circ> f) ` A))"
   224   unfolding Inf_prod_def by simp
   224   unfolding Inf_prod_def by simp
   225 
   225 
   226 lemma SUP_prod_alt_def:
   226 lemma SUP_prod_alt_def:
   227   "SUPREMUM A f = (SUPREMUM A (fst \<circ> f), SUPREMUM A (snd \<circ> f))"
   227   "Sup (f ` A) = (Sup ((fst \<circ> f) ` A), Sup((snd \<circ> f) ` A))"
   228   unfolding Sup_prod_def by simp
   228   unfolding Sup_prod_def by simp
   229 
   229 
   230 
   230 
   231 subsection \<open>Complete distributive lattices\<close>
   231 subsection \<open>Complete distributive lattices\<close>
   232 
   232 
   233 (* Contribution: Alessandro Coglio *)
   233 (* Contribution: Alessandro Coglio *)
   234 
   234 
   235 instance prod :: (complete_distrib_lattice, complete_distrib_lattice) complete_distrib_lattice
   235 instance prod :: (complete_distrib_lattice, complete_distrib_lattice) complete_distrib_lattice
   236 proof
   236 proof
   237   fix A::"('a\<times>'b) set set"
   237   fix A::"('a\<times>'b) set set"
   238   show "INFIMUM A Sup \<le> SUPREMUM {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y} Inf"
   238   show "Inf (Sup ` A) \<le> Sup (Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"
   239     by (simp add: Sup_prod_def Inf_prod_def INF_SUP_set)
   239     by (simp add: Sup_prod_def Inf_prod_def INF_SUP_set)
   240 qed
   240 qed
   241 
   241 
   242 subsection \<open>Bekic's Theorem\<close>
   242 subsection \<open>Bekic's Theorem\<close>
   243 text \<open>
   243 text \<open>