| 30661 |      1 | (* Author: Amine Chaieb, University of Cambridge *)
 | 
| 29838 |      2 | 
 | 
| 30661 |      3 | header {* Definitions of Lower Bounds and Greatest Lower Bounds, analogous to Lubs *}
 | 
| 29838 |      4 | 
 | 
|  |      5 | theory Glbs
 | 
|  |      6 | imports Lubs
 | 
|  |      7 | begin
 | 
|  |      8 | 
 | 
| 46509 |      9 | definition greatestP :: "('a \<Rightarrow> bool) \<Rightarrow> 'a::ord \<Rightarrow> bool"
 | 
|  |     10 |   where "greatestP P x = (P x \<and> Collect P *<=  x)"
 | 
| 29838 |     11 | 
 | 
| 46509 |     12 | definition isLb :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"
 | 
|  |     13 |   where "isLb R S x = (x <=* S \<and> x: R)"
 | 
| 29838 |     14 | 
 | 
| 46509 |     15 | definition isGlb :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"
 | 
|  |     16 |   where "isGlb R S x = greatestP (isLb R S) x"
 | 
| 29838 |     17 | 
 | 
| 46509 |     18 | definition lbs :: "'a set \<Rightarrow> 'a::ord set \<Rightarrow> 'a set"
 | 
|  |     19 |   where "lbs R S = Collect (isLb R S)"
 | 
|  |     20 | 
 | 
| 29838 |     21 | 
 | 
| 46509 |     22 | subsection {* Rules about the Operators @{term greatestP}, @{term isLb}
 | 
|  |     23 |   and @{term isGlb} *}
 | 
| 29838 |     24 | 
 | 
| 46509 |     25 | lemma leastPD1: "greatestP P x \<Longrightarrow> P x"
 | 
|  |     26 |   by (simp add: greatestP_def)
 | 
| 29838 |     27 | 
 | 
| 46509 |     28 | lemma greatestPD2: "greatestP P x \<Longrightarrow> Collect P *<= x"
 | 
|  |     29 |   by (simp add: greatestP_def)
 | 
| 29838 |     30 | 
 | 
| 46509 |     31 | lemma greatestPD3: "greatestP P x \<Longrightarrow> y: Collect P \<Longrightarrow> x \<ge> y"
 | 
|  |     32 |   by (blast dest!: greatestPD2 setleD)
 | 
| 29838 |     33 | 
 | 
| 46509 |     34 | lemma isGlbD1: "isGlb R S x \<Longrightarrow> x <=* S"
 | 
|  |     35 |   by (simp add: isGlb_def isLb_def greatestP_def)
 | 
| 29838 |     36 | 
 | 
| 46509 |     37 | lemma isGlbD1a: "isGlb R S x \<Longrightarrow> x: R"
 | 
|  |     38 |   by (simp add: isGlb_def isLb_def greatestP_def)
 | 
| 29838 |     39 | 
 | 
| 46509 |     40 | lemma isGlb_isLb: "isGlb R S x \<Longrightarrow> isLb R S x"
 | 
|  |     41 |   unfolding isLb_def by (blast dest: isGlbD1 isGlbD1a)
 | 
| 29838 |     42 | 
 | 
| 46509 |     43 | lemma isGlbD2: "isGlb R S x \<Longrightarrow> y : S \<Longrightarrow> y \<ge> x"
 | 
|  |     44 |   by (blast dest!: isGlbD1 setgeD)
 | 
| 29838 |     45 | 
 | 
| 46509 |     46 | lemma isGlbD3: "isGlb R S x \<Longrightarrow> greatestP (isLb R S) x"
 | 
|  |     47 |   by (simp add: isGlb_def)
 | 
| 29838 |     48 | 
 | 
| 46509 |     49 | lemma isGlbI1: "greatestP (isLb R S) x \<Longrightarrow> isGlb R S x"
 | 
|  |     50 |   by (simp add: isGlb_def)
 | 
| 29838 |     51 | 
 | 
| 46509 |     52 | lemma isGlbI2: "isLb R S x \<Longrightarrow> Collect (isLb R S) *<= x \<Longrightarrow> isGlb R S x"
 | 
|  |     53 |   by (simp add: isGlb_def greatestP_def)
 | 
| 29838 |     54 | 
 | 
| 46509 |     55 | lemma isLbD: "isLb R S x \<Longrightarrow> y : S \<Longrightarrow> y \<ge> x"
 | 
|  |     56 |   by (simp add: isLb_def setge_def)
 | 
| 29838 |     57 | 
 | 
| 46509 |     58 | lemma isLbD2: "isLb R S x \<Longrightarrow> x <=* S "
 | 
|  |     59 |   by (simp add: isLb_def)
 | 
| 29838 |     60 | 
 | 
| 46509 |     61 | lemma isLbD2a: "isLb R S x \<Longrightarrow> x: R"
 | 
|  |     62 |   by (simp add: isLb_def)
 | 
| 29838 |     63 | 
 | 
| 46509 |     64 | lemma isLbI: "x <=* S \<Longrightarrow> x: R \<Longrightarrow> isLb R S x"
 | 
|  |     65 |   by (simp add: isLb_def)
 | 
| 29838 |     66 | 
 | 
| 46509 |     67 | lemma isGlb_le_isLb: "isGlb R S x \<Longrightarrow> isLb R S y \<Longrightarrow> x \<ge> y"
 | 
|  |     68 |   unfolding isGlb_def by (blast intro!: greatestPD3)
 | 
| 29838 |     69 | 
 | 
| 46509 |     70 | lemma isGlb_ubs: "isGlb R S x \<Longrightarrow> lbs R S *<= x"
 | 
|  |     71 |   unfolding lbs_def isGlb_def by (rule greatestPD2)
 | 
| 29838 |     72 | 
 | 
| 51342 |     73 | lemma isGlb_unique: "[| isGlb R S x; isGlb R S y |] ==> x = (y::'a::linorder)"
 | 
|  |     74 |   apply (frule isGlb_isLb)
 | 
|  |     75 |   apply (frule_tac x = y in isGlb_isLb)
 | 
|  |     76 |   apply (blast intro!: order_antisym dest!: isGlb_le_isLb)
 | 
|  |     77 |   done
 | 
|  |     78 | 
 | 
| 29838 |     79 | end
 |