src/HOL/Library/Glbs.thy
author bulwahn
Tue, 07 Dec 2010 11:50:16 +0100
changeset 41044 1c0eefa8d02a
parent 30661 54858c8ad226
child 46509 c4b2ec379fdd
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
30661
54858c8ad226 tuned header
haftmann
parents: 30267
diff changeset
     1
(* Author: Amine Chaieb, University of Cambridge *)
29838
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
     2
30661
54858c8ad226 tuned header
haftmann
parents: 30267
diff changeset
     3
header {* Definitions of Lower Bounds and Greatest Lower Bounds, analogous to Lubs *}
29838
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
     4
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
     5
theory Glbs
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
     6
imports Lubs
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
     7
begin
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
     8
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
     9
definition
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    10
  greatestP      :: "['a =>bool,'a::ord] => bool" where
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    11
  "greatestP P x = (P x & Collect P *<=  x)"
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    12
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    13
definition
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    14
  isLb        :: "['a set, 'a set, 'a::ord] => bool" where
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    15
  "isLb R S x = (x <=* S & x: R)"
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    16
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    17
definition
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    18
  isGlb       :: "['a set, 'a set, 'a::ord] => bool" where
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    19
  "isGlb R S x = greatestP (isLb R S) x"
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    20
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    21
definition
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    22
  lbs         :: "['a set, 'a::ord set] => 'a set" where
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    23
  "lbs R S = Collect (isLb R S)"
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    24
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    25
subsection{*Rules about the Operators @{term greatestP}, @{term isLb}
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    26
    and @{term isGlb}*}
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    27
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    28
lemma leastPD1: "greatestP P x ==> P x"
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    29
by (simp add: greatestP_def)
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    30
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    31
lemma greatestPD2: "greatestP P x ==> Collect P *<= x"
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    32
by (simp add: greatestP_def)
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    33
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    34
lemma greatestPD3: "[| greatestP P x; y: Collect P |] ==> x >= y"
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    35
by (blast dest!: greatestPD2 setleD)
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    36
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    37
lemma isGlbD1: "isGlb R S x ==> x <=* S"
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    38
by (simp add: isGlb_def isLb_def greatestP_def)
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    39
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    40
lemma isGlbD1a: "isGlb R S x ==> x: R"
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    41
by (simp add: isGlb_def isLb_def greatestP_def)
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    42
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    43
lemma isGlb_isLb: "isGlb R S x ==> isLb R S x"
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    44
apply (simp add: isLb_def)
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    45
apply (blast dest: isGlbD1 isGlbD1a)
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    46
done
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    47
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    48
lemma isGlbD2: "[| isGlb R S x; y : S |] ==> y >= x"
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    49
by (blast dest!: isGlbD1 setgeD)
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    50
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    51
lemma isGlbD3: "isGlb R S x ==> greatestP(isLb R S) x"
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    52
by (simp add: isGlb_def)
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    53
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    54
lemma isGlbI1: "greatestP(isLb R S) x ==> isGlb R S x"
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    55
by (simp add: isGlb_def)
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    56
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    57
lemma isGlbI2: "[| isLb R S x; Collect (isLb R S) *<= x |] ==> isGlb R S x"
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    58
by (simp add: isGlb_def greatestP_def)
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    59
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    60
lemma isLbD: "[| isLb R S x; y : S |] ==> y >= x"
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    61
by (simp add: isLb_def setge_def)
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    62
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    63
lemma isLbD2: "isLb R S x ==> x <=* S "
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    64
by (simp add: isLb_def)
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    65
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    66
lemma isLbD2a: "isLb R S x ==> x: R"
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    67
by (simp add: isLb_def)
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    68
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    69
lemma isLbI: "[| x <=* S ; x: R |] ==> isLb R S x"
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    70
by (simp add: isLb_def)
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    71
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    72
lemma isGlb_le_isLb: "[| isGlb R S x; isLb R S y |] ==> x >= y"
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    73
apply (simp add: isGlb_def)
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    74
apply (blast intro!: greatestPD3)
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    75
done
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    76
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    77
lemma isGlb_ubs: "isGlb R S x ==> lbs R S *<= x"
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    78
apply (simp add: lbs_def isGlb_def)
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    79
apply (erule greatestPD2)
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    80
done
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    81
a562ca0c408d A theory of greatest lower bounds
chaieb
parents:
diff changeset
    82
end