author | nipkow |
Fri Mar 06 17:38:47 2009 +0100 (2009-03-06) | |
changeset 30313 | b2441b0c8d38 |
parent 29654 | 24e73987bfe2 |
child 30738 | 0842e906300c |
permissions | -rw-r--r-- |
paulson@5078 | 1 |
(* Title : Lubs.thy |
paulson@5078 | 2 |
Author : Jacques D. Fleuriot |
paulson@5078 | 3 |
Copyright : 1998 University of Cambridge |
paulson@14368 | 4 |
*) |
paulson@5078 | 5 |
|
paulson@14368 | 6 |
header{*Definitions of Upper Bounds and Least Upper Bounds*} |
paulson@5078 | 7 |
|
nipkow@15131 | 8 |
theory Lubs |
haftmann@29654 | 9 |
imports Plain Main |
nipkow@15131 | 10 |
begin |
paulson@5078 | 11 |
|
paulson@14368 | 12 |
text{*Thanks to suggestions by James Margetson*} |
paulson@5078 | 13 |
|
wenzelm@19765 | 14 |
definition |
wenzelm@21404 | 15 |
setle :: "['a set, 'a::ord] => bool" (infixl "*<=" 70) where |
wenzelm@19765 | 16 |
"S *<= x = (ALL y: S. y <= x)" |
paulson@14368 | 17 |
|
wenzelm@21404 | 18 |
definition |
wenzelm@21404 | 19 |
setge :: "['a::ord, 'a set] => bool" (infixl "<=*" 70) where |
wenzelm@19765 | 20 |
"x <=* S = (ALL y: S. x <= y)" |
paulson@14368 | 21 |
|
wenzelm@21404 | 22 |
definition |
wenzelm@21404 | 23 |
leastP :: "['a =>bool,'a::ord] => bool" where |
wenzelm@19765 | 24 |
"leastP P x = (P x & x <=* Collect P)" |
paulson@5078 | 25 |
|
wenzelm@21404 | 26 |
definition |
wenzelm@21404 | 27 |
isUb :: "['a set, 'a set, 'a::ord] => bool" where |
wenzelm@19765 | 28 |
"isUb R S x = (S *<= x & x: R)" |
wenzelm@14653 | 29 |
|
wenzelm@21404 | 30 |
definition |
wenzelm@21404 | 31 |
isLub :: "['a set, 'a set, 'a::ord] => bool" where |
wenzelm@19765 | 32 |
"isLub R S x = leastP (isUb R S) x" |
paulson@5078 | 33 |
|
wenzelm@21404 | 34 |
definition |
wenzelm@21404 | 35 |
ubs :: "['a set, 'a::ord set] => 'a set" where |
wenzelm@19765 | 36 |
"ubs R S = Collect (isUb R S)" |
paulson@5078 | 37 |
|
paulson@14368 | 38 |
|
paulson@14368 | 39 |
|
paulson@14368 | 40 |
subsection{*Rules for the Relations @{text "*<="} and @{text "<=*"}*} |
paulson@14368 | 41 |
|
paulson@14368 | 42 |
lemma setleI: "ALL y: S. y <= x ==> S *<= x" |
paulson@14368 | 43 |
by (simp add: setle_def) |
paulson@14368 | 44 |
|
paulson@14368 | 45 |
lemma setleD: "[| S *<= x; y: S |] ==> y <= x" |
paulson@14368 | 46 |
by (simp add: setle_def) |
paulson@14368 | 47 |
|
paulson@14368 | 48 |
lemma setgeI: "ALL y: S. x<= y ==> x <=* S" |
paulson@14368 | 49 |
by (simp add: setge_def) |
paulson@14368 | 50 |
|
paulson@14368 | 51 |
lemma setgeD: "[| x <=* S; y: S |] ==> x <= y" |
paulson@14368 | 52 |
by (simp add: setge_def) |
paulson@14368 | 53 |
|
paulson@14368 | 54 |
|
paulson@14368 | 55 |
subsection{*Rules about the Operators @{term leastP}, @{term ub} |
paulson@14368 | 56 |
and @{term lub}*} |
paulson@14368 | 57 |
|
paulson@14368 | 58 |
lemma leastPD1: "leastP P x ==> P x" |
paulson@14368 | 59 |
by (simp add: leastP_def) |
paulson@14368 | 60 |
|
paulson@14368 | 61 |
lemma leastPD2: "leastP P x ==> x <=* Collect P" |
paulson@14368 | 62 |
by (simp add: leastP_def) |
paulson@14368 | 63 |
|
paulson@14368 | 64 |
lemma leastPD3: "[| leastP P x; y: Collect P |] ==> x <= y" |
paulson@14368 | 65 |
by (blast dest!: leastPD2 setgeD) |
paulson@14368 | 66 |
|
paulson@14368 | 67 |
lemma isLubD1: "isLub R S x ==> S *<= x" |
paulson@14368 | 68 |
by (simp add: isLub_def isUb_def leastP_def) |
paulson@14368 | 69 |
|
paulson@14368 | 70 |
lemma isLubD1a: "isLub R S x ==> x: R" |
paulson@14368 | 71 |
by (simp add: isLub_def isUb_def leastP_def) |
paulson@14368 | 72 |
|
paulson@14368 | 73 |
lemma isLub_isUb: "isLub R S x ==> isUb R S x" |
paulson@14368 | 74 |
apply (simp add: isUb_def) |
paulson@14368 | 75 |
apply (blast dest: isLubD1 isLubD1a) |
paulson@14368 | 76 |
done |
paulson@14368 | 77 |
|
paulson@14368 | 78 |
lemma isLubD2: "[| isLub R S x; y : S |] ==> y <= x" |
paulson@14368 | 79 |
by (blast dest!: isLubD1 setleD) |
paulson@14368 | 80 |
|
paulson@14368 | 81 |
lemma isLubD3: "isLub R S x ==> leastP(isUb R S) x" |
paulson@14368 | 82 |
by (simp add: isLub_def) |
paulson@14368 | 83 |
|
paulson@14368 | 84 |
lemma isLubI1: "leastP(isUb R S) x ==> isLub R S x" |
paulson@14368 | 85 |
by (simp add: isLub_def) |
paulson@14368 | 86 |
|
paulson@14368 | 87 |
lemma isLubI2: "[| isUb R S x; x <=* Collect (isUb R S) |] ==> isLub R S x" |
paulson@14368 | 88 |
by (simp add: isLub_def leastP_def) |
paulson@5078 | 89 |
|
paulson@14368 | 90 |
lemma isUbD: "[| isUb R S x; y : S |] ==> y <= x" |
paulson@14368 | 91 |
by (simp add: isUb_def setle_def) |
paulson@14368 | 92 |
|
paulson@14368 | 93 |
lemma isUbD2: "isUb R S x ==> S *<= x" |
paulson@14368 | 94 |
by (simp add: isUb_def) |
paulson@14368 | 95 |
|
paulson@14368 | 96 |
lemma isUbD2a: "isUb R S x ==> x: R" |
paulson@14368 | 97 |
by (simp add: isUb_def) |
paulson@14368 | 98 |
|
paulson@14368 | 99 |
lemma isUbI: "[| S *<= x; x: R |] ==> isUb R S x" |
paulson@14368 | 100 |
by (simp add: isUb_def) |
paulson@14368 | 101 |
|
paulson@14368 | 102 |
lemma isLub_le_isUb: "[| isLub R S x; isUb R S y |] ==> x <= y" |
paulson@14368 | 103 |
apply (simp add: isLub_def) |
paulson@14368 | 104 |
apply (blast intro!: leastPD3) |
paulson@14368 | 105 |
done |
paulson@14368 | 106 |
|
paulson@14368 | 107 |
lemma isLub_ubs: "isLub R S x ==> x <=* ubs R S" |
paulson@14368 | 108 |
apply (simp add: ubs_def isLub_def) |
paulson@14368 | 109 |
apply (erule leastPD2) |
paulson@14368 | 110 |
done |
paulson@5078 | 111 |
|
paulson@14368 | 112 |
end |