src/HOL/Algebra/Lattice.thy
changeset 21049 379542c9d951
parent 21041 60e418260b4d
child 21404 eb85850d3eb7
equal deleted inserted replaced
21048:e57e91f72831 21049:379542c9d951
    16   carrier :: "'a set"
    16   carrier :: "'a set"
    17 
    17 
    18 
    18 
    19 subsection {* Partial Orders *}
    19 subsection {* Partial Orders *}
    20 
    20 
    21 (*
       
    22 record 'a order = "'a partial_object" +
       
    23   le :: "['a, 'a] => bool" (infixl "\<sqsubseteq>\<index>" 50)
       
    24 *)
       
    25 
       
    26 text {* Locale @{text order_syntax} is required since we want to refer
    21 text {* Locale @{text order_syntax} is required since we want to refer
    27   to definitions (and their derived theorems) outside of @{text partial_order}.
    22   to definitions (and their derived theorems) outside of @{text partial_order}.
    28   *}
    23   *}
    29 
    24 
    30 locale order_syntax =
    25 locale order_syntax =
    31   fixes carrier :: "'a set" and le :: "['a, 'a] => bool" (infix "\<sqsubseteq>" 50)
    26   fixes L :: "'a set" and le :: "['a, 'a] => bool" (infix "\<sqsubseteq>" 50)
    32 
    27 
    33 text {* Note that the type constraints above are necessary, because the
    28 text {* Note that the type constraints above are necessary, because the
    34   definition command cannot specialise the types. *}
    29   definition command cannot specialise the types. *}
    35 
    30 
    36 definition (in order_syntax)
    31 definition (in order_syntax)
    38 
    33 
    39 text {* Upper and lower bounds of a set. *}
    34 text {* Upper and lower bounds of a set. *}
    40 
    35 
    41 definition (in order_syntax)
    36 definition (in order_syntax)
    42   Upper where
    37   Upper where
    43   "Upper A == {u. (ALL x. x \<in> A \<inter> carrier --> x \<sqsubseteq> u)} \<inter>
    38   "Upper A == {u. (ALL x. x \<in> A \<inter> L --> x \<sqsubseteq> u)} \<inter> L"
    44               carrier"
       
    45 
    39 
    46 definition (in order_syntax)
    40 definition (in order_syntax)
    47   Lower :: "'a set => 'a set"
    41   Lower :: "'a set => 'a set"
    48   "Lower A == {l. (ALL x. x \<in> A \<inter> carrier --> l \<sqsubseteq> x)} \<inter>
    42   "Lower A == {l. (ALL x. x \<in> A \<inter> L --> l \<sqsubseteq> x)} \<inter> L"
    49               carrier"
       
    50 
    43 
    51 text {* Least and greatest, as predicate. *}
    44 text {* Least and greatest, as predicate. *}
    52 
    45 
    53 definition (in order_syntax)
    46 definition (in order_syntax)
    54   least :: "['a, 'a set] => bool"
    47   least :: "['a, 'a set] => bool"
    55   "least l A == A \<subseteq> carrier & l \<in> A & (ALL x : A. l \<sqsubseteq> x)"
    48   "least l A == A \<subseteq> L & l \<in> A & (ALL x : A. l \<sqsubseteq> x)"
    56 
    49 
    57 definition (in order_syntax)
    50 definition (in order_syntax)
    58   greatest :: "['a, 'a set] => bool"
    51   greatest :: "['a, 'a set] => bool"
    59   "greatest g A == A \<subseteq> carrier & g \<in> A & (ALL x : A. x \<sqsubseteq> g)"
    52   "greatest g A == A \<subseteq> L & g \<in> A & (ALL x : A. x \<sqsubseteq> g)"
    60 
    53 
    61 text {* Supremum and infimum *}
    54 text {* Supremum and infimum *}
    62 
    55 
    63 definition (in order_syntax)
    56 definition (in order_syntax)
    64   sup :: "'a set => 'a" ("\<Squnion>")  (* FIXME precedence [90] 90 *)
    57   sup :: "'a set => 'a" ("\<Squnion>_" [90] 90)
    65   "\<Squnion>A == THE x. least x (Upper A)"
    58   "\<Squnion>A == THE x. least x (Upper A)"
    66 
    59 
    67 definition (in order_syntax)
    60 definition (in order_syntax)
    68   inf :: "'a set => 'a" ("\<Sqinter>") (* FIXME precedence *)
    61   inf :: "'a set => 'a" ("\<Sqinter>_" [90] 90)
    69   "\<Sqinter>A == THE x. greatest x (Lower A)"
    62   "\<Sqinter>A == THE x. greatest x (Lower A)"
    70 
    63 
    71 definition (in order_syntax)
    64 definition (in order_syntax)
    72   join :: "['a, 'a] => 'a" (infixl "\<squnion>" 65)
    65   join :: "['a, 'a] => 'a" (infixl "\<squnion>" 65)
    73   "x \<squnion> y == sup {x, y}"
    66   "x \<squnion> y == sup {x, y}"
    76   meet :: "['a, 'a] => 'a" (infixl "\<sqinter>" 70)
    69   meet :: "['a, 'a] => 'a" (infixl "\<sqinter>" 70)
    77   "x \<sqinter> y == inf {x, y}"
    70   "x \<sqinter> y == inf {x, y}"
    78 
    71 
    79 locale partial_order = order_syntax +
    72 locale partial_order = order_syntax +
    80   assumes refl [intro, simp]:
    73   assumes refl [intro, simp]:
    81                   "x \<in> carrier ==> x \<sqsubseteq> x"
    74                   "x \<in> L ==> x \<sqsubseteq> x"
    82     and anti_sym [intro]:
    75     and anti_sym [intro]:
    83                   "[| x \<sqsubseteq> y; y \<sqsubseteq> x; x \<in> carrier; y \<in> carrier |] ==> x = y"
    76                   "[| x \<sqsubseteq> y; y \<sqsubseteq> x; x \<in> L; y \<in> L |] ==> x = y"
    84     and trans [trans]:
    77     and trans [trans]:
    85                   "[| x \<sqsubseteq> y; y \<sqsubseteq> z;
    78                   "[| x \<sqsubseteq> y; y \<sqsubseteq> z;
    86                    x \<in> carrier; y \<in> carrier; z \<in> carrier |] ==> x \<sqsubseteq> z"
    79                    x \<in> L; y \<in> L; z \<in> L |] ==> x \<sqsubseteq> z"
    87 
    80 
    88 abbreviation (in partial_order)
    81 abbreviation (in partial_order)
    89   less (infixl "\<sqsubset>" 50) "less == order_syntax.less le"
    82   less (infixl "\<sqsubset>" 50) "less == order_syntax.less le"
    90 abbreviation (in partial_order)
    83 abbreviation (in partial_order)
    91   Upper where "Upper == order_syntax.Upper carrier le"
    84   Upper where "Upper == order_syntax.Upper L le"
    92 abbreviation (in partial_order)
    85 abbreviation (in partial_order)
    93   Lower where "Lower == order_syntax.Lower carrier le"
    86   Lower where "Lower == order_syntax.Lower L le"
    94 abbreviation (in partial_order)
    87 abbreviation (in partial_order)
    95   least where "least == order_syntax.least carrier le"
    88   least where "least == order_syntax.least L le"
    96 abbreviation (in partial_order)
    89 abbreviation (in partial_order)
    97   greatest where "greatest == order_syntax.greatest carrier le"
    90   greatest where "greatest == order_syntax.greatest L le"
    98 abbreviation (in partial_order)
    91 abbreviation (in partial_order)
    99   sup ("\<Squnion>") (* FIXME precedence *) "sup == order_syntax.sup carrier le"
    92   sup ("\<Squnion>_" [90] 90) "sup == order_syntax.sup L le"
   100 abbreviation (in partial_order)
    93 abbreviation (in partial_order)
   101   inf ("\<Sqinter>") (* FIXME precedence *) "inf == order_syntax.inf carrier le"
    94   inf ("\<Sqinter>_" [90] 90) "inf == order_syntax.inf L le"
   102 abbreviation (in partial_order)
    95 abbreviation (in partial_order)
   103   join (infixl "\<squnion>" 65) "join == order_syntax.join carrier le"
    96   join (infixl "\<squnion>" 65) "join == order_syntax.join L le"
   104 abbreviation (in partial_order)
    97 abbreviation (in partial_order)
   105   meet (infixl "\<sqinter>" 70) "meet == order_syntax.meet carrier le"
    98   meet (infixl "\<sqinter>" 70) "meet == order_syntax.meet L le"
   106 
    99 
   107 
   100 
   108 subsubsection {* Upper *}
   101 subsubsection {* Upper *}
   109 
   102 
   110 lemma (in order_syntax) Upper_closed [intro, simp]:
   103 lemma (in order_syntax) Upper_closed [intro, simp]:
   111   "Upper A \<subseteq> carrier"
   104   "Upper A \<subseteq> L"
   112   by (unfold Upper_def) clarify
   105   by (unfold Upper_def) clarify
   113 
   106 
   114 lemma (in order_syntax) UpperD [dest]:
   107 lemma (in order_syntax) UpperD [dest]:
   115   fixes L (structure)
   108   "[| u \<in> Upper A; x \<in> A; A \<subseteq> L |] ==> x \<sqsubseteq> u"
   116   shows "[| u \<in> Upper A; x \<in> A; A \<subseteq> carrier |] ==> x \<sqsubseteq> u"
       
   117   by (unfold Upper_def) blast
   109   by (unfold Upper_def) blast
   118 
   110 
   119 lemma (in order_syntax) Upper_memI:
   111 lemma (in order_syntax) Upper_memI:
   120   fixes L (structure)
   112   "[| !! y. y \<in> A ==> y \<sqsubseteq> x; x \<in> L |] ==> x \<in> Upper A"
   121   shows "[| !! y. y \<in> A ==> y \<sqsubseteq> x; x \<in> carrier |] ==> x \<in> Upper A"
       
   122   by (unfold Upper_def) blast
   113   by (unfold Upper_def) blast
   123 
   114 
   124 lemma (in order_syntax) Upper_antimono:
   115 lemma (in order_syntax) Upper_antimono:
   125   "A \<subseteq> B ==> Upper B \<subseteq> Upper A"
   116   "A \<subseteq> B ==> Upper B \<subseteq> Upper A"
   126   by (unfold Upper_def) blast
   117   by (unfold Upper_def) blast
   127 
   118 
   128 
   119 
   129 subsubsection {* Lower *}
   120 subsubsection {* Lower *}
   130 
   121 
   131 lemma (in order_syntax) Lower_closed [intro, simp]:
   122 lemma (in order_syntax) Lower_closed [intro, simp]:
   132   "Lower A \<subseteq> carrier"
   123   "Lower A \<subseteq> L"
   133   by (unfold Lower_def) clarify
   124   by (unfold Lower_def) clarify
   134 
   125 
   135 lemma (in order_syntax) LowerD [dest]:
   126 lemma (in order_syntax) LowerD [dest]:
   136   "[| l \<in> Lower A; x \<in> A; A \<subseteq> carrier |] ==> l \<sqsubseteq> x"
   127   "[| l \<in> Lower A; x \<in> A; A \<subseteq> L |] ==> l \<sqsubseteq> x"
   137   by (unfold Lower_def) blast
   128   by (unfold Lower_def) blast
   138 
   129 
   139 lemma (in order_syntax) Lower_memI:
   130 lemma (in order_syntax) Lower_memI:
   140   "[| !! y. y \<in> A ==> x \<sqsubseteq> y; x \<in> carrier |] ==> x \<in> Lower A"
   131   "[| !! y. y \<in> A ==> x \<sqsubseteq> y; x \<in> L |] ==> x \<in> Lower A"
   141   by (unfold Lower_def) blast
   132   by (unfold Lower_def) blast
   142 
   133 
   143 lemma (in order_syntax) Lower_antimono:
   134 lemma (in order_syntax) Lower_antimono:
   144   "A \<subseteq> B ==> Lower B \<subseteq> Lower A"
   135   "A \<subseteq> B ==> Lower B \<subseteq> Lower A"
   145   by (unfold Lower_def) blast
   136   by (unfold Lower_def) blast
   146 
   137 
   147 
   138 
   148 subsubsection {* least *}
   139 subsubsection {* least *}
   149 
   140 
   150 lemma (in order_syntax) least_carrier [intro, simp]:
   141 lemma (in order_syntax) least_closed [intro, simp]:
   151   "least l A ==> l \<in> carrier"
   142   "least l A ==> l \<in> L"
   152   by (unfold least_def) fast
   143   by (unfold least_def) fast
   153 
   144 
   154 lemma (in order_syntax) least_mem:
   145 lemma (in order_syntax) least_mem:
   155   "least l A ==> l \<in> A"
   146   "least l A ==> l \<in> A"
   156   by (unfold least_def) fast
   147   by (unfold least_def) fast
   164   by (unfold least_def) fast
   155   by (unfold least_def) fast
   165 
   156 
   166 lemma (in order_syntax) least_UpperI:
   157 lemma (in order_syntax) least_UpperI:
   167   assumes above: "!! x. x \<in> A ==> x \<sqsubseteq> s"
   158   assumes above: "!! x. x \<in> A ==> x \<sqsubseteq> s"
   168     and below: "!! y. y \<in> Upper A ==> s \<sqsubseteq> y"
   159     and below: "!! y. y \<in> Upper A ==> s \<sqsubseteq> y"
   169     and L: "A \<subseteq> carrier"  "s \<in> carrier"
   160     and L: "A \<subseteq> L"  "s \<in> L"
   170   shows "least s (Upper A)"
   161   shows "least s (Upper A)"
   171 proof -
   162 proof -
   172   have "Upper A \<subseteq> carrier" by simp
   163   have "Upper A \<subseteq> L" by simp
   173   moreover from above L have "s \<in> Upper A" by (simp add: Upper_def)
   164   moreover from above L have "s \<in> Upper A" by (simp add: Upper_def)
   174   moreover from below have "ALL x : Upper A. s \<sqsubseteq> x" by fast
   165   moreover from below have "ALL x : Upper A. s \<sqsubseteq> x" by fast
   175   ultimately show ?thesis by (simp add: least_def)
   166   ultimately show ?thesis by (simp add: least_def)
   176 qed
   167 qed
   177 
   168 
   178 
   169 
   179 subsubsection {* greatest *}
   170 subsubsection {* greatest *}
   180 
   171 
   181 lemma (in order_syntax) greatest_carrier [intro, simp]:
   172 lemma (in order_syntax) greatest_closed [intro, simp]:
   182   "greatest l A ==> l \<in> carrier"
   173   "greatest l A ==> l \<in> L"
   183   by (unfold greatest_def) fast
   174   by (unfold greatest_def) fast
   184 
   175 
   185 lemma (in order_syntax) greatest_mem:
   176 lemma (in order_syntax) greatest_mem:
   186   "greatest l A ==> l \<in> A"
   177   "greatest l A ==> l \<in> A"
   187   by (unfold greatest_def) fast
   178   by (unfold greatest_def) fast
   195   by (unfold greatest_def) fast
   186   by (unfold greatest_def) fast
   196 
   187 
   197 lemma (in order_syntax) greatest_LowerI:
   188 lemma (in order_syntax) greatest_LowerI:
   198   assumes below: "!! x. x \<in> A ==> i \<sqsubseteq> x"
   189   assumes below: "!! x. x \<in> A ==> i \<sqsubseteq> x"
   199     and above: "!! y. y \<in> Lower A ==> y \<sqsubseteq> i"
   190     and above: "!! y. y \<in> Lower A ==> y \<sqsubseteq> i"
   200     and L: "A \<subseteq> carrier"  "i \<in> carrier"
   191     and L: "A \<subseteq> L"  "i \<in> L"
   201   shows "greatest i (Lower A)"
   192   shows "greatest i (Lower A)"
   202 proof -
   193 proof -
   203   have "Lower A \<subseteq> carrier" by simp
   194   have "Lower A \<subseteq> L" by simp
   204   moreover from below L have "i \<in> Lower A" by (simp add: Lower_def)
   195   moreover from below L have "i \<in> Lower A" by (simp add: Lower_def)
   205   moreover from above have "ALL x : Lower A. x \<sqsubseteq> i" by fast
   196   moreover from above have "ALL x : Lower A. x \<sqsubseteq> i" by fast
   206   ultimately show ?thesis by (simp add: greatest_def)
   197   ultimately show ?thesis by (simp add: greatest_def)
   207 qed
   198 qed
   208 
   199 
   209 
   200 
   210 subsection {* Lattices *}
   201 subsection {* Lattices *}
   211 
   202 
   212 locale lattice = partial_order +
   203 locale lattice = partial_order +
   213   assumes sup_of_two_exists:
   204   assumes sup_of_two_exists:
   214     "[| x \<in> carrier; y \<in> carrier |] ==> EX s. order_syntax.least carrier le s (order_syntax.Upper carrier le {x, y})"
   205     "[| x \<in> L; y \<in> L |] ==> EX s. order_syntax.least L le s (order_syntax.Upper L le {x, y})"
   215     and inf_of_two_exists:
   206     and inf_of_two_exists:
   216     "[| x \<in> carrier; y \<in> carrier |] ==> EX s. order_syntax.greatest carrier le s (order_syntax.Lower carrier le {x, y})"
   207     "[| x \<in> L; y \<in> L |] ==> EX s. order_syntax.greatest L le s (order_syntax.Lower L le {x, y})"
   217 
   208 
   218 abbreviation (in lattice)
   209 abbreviation (in lattice)
   219   less (infixl "\<sqsubset>" 50) "less == order_syntax.less le"
   210   less (infixl "\<sqsubset>" 50) "less == order_syntax.less le"
   220 abbreviation (in lattice)
   211 abbreviation (in lattice)
   221   Upper where "Upper == order_syntax.Upper carrier le"
   212   Upper where "Upper == order_syntax.Upper L le"
   222 abbreviation (in lattice)
   213 abbreviation (in lattice)
   223   Lower where "Lower == order_syntax.Lower carrier le"
   214   Lower where "Lower == order_syntax.Lower L le"
   224 abbreviation (in lattice)
   215 abbreviation (in lattice)
   225   least where "least == order_syntax.least carrier le"
   216   least where "least == order_syntax.least L le"
   226 abbreviation (in lattice)
   217 abbreviation (in lattice)
   227   greatest where "greatest == order_syntax.greatest carrier le"
   218   greatest where "greatest == order_syntax.greatest L le"
   228 abbreviation (in lattice)
   219 abbreviation (in lattice)
   229   sup ("\<Squnion>") (* FIXME precedence *) "sup == order_syntax.sup carrier le"
   220   sup ("\<Squnion>_" [90] 90) "sup == order_syntax.sup L le"
   230 abbreviation (in lattice)
   221 abbreviation (in lattice)
   231   inf ("\<Sqinter>") (* FIXME precedence *) "inf == order_syntax.inf carrier le"
   222   inf ("\<Sqinter>_" [90] 90) "inf == order_syntax.inf L le"
   232 abbreviation (in lattice)
   223 abbreviation (in lattice)
   233   join (infixl "\<squnion>" 65) "join == order_syntax.join carrier le"
   224   join (infixl "\<squnion>" 65) "join == order_syntax.join L le"
   234 abbreviation (in lattice)
   225 abbreviation (in lattice)
   235   meet (infixl "\<sqinter>" 70) "meet == order_syntax.meet carrier le"
   226   meet (infixl "\<sqinter>" 70) "meet == order_syntax.meet L le"
   236 
   227 
   237 lemma (in order_syntax) least_Upper_above:
   228 lemma (in order_syntax) least_Upper_above:
   238   "[| least s (Upper A); x \<in> A; A \<subseteq> carrier |] ==> x \<sqsubseteq> s"
   229   "[| least s (Upper A); x \<in> A; A \<subseteq> L |] ==> x \<sqsubseteq> s"
   239   by (unfold least_def) blast
   230   by (unfold least_def) blast
   240 
   231 
   241 lemma (in order_syntax) greatest_Lower_above:
   232 lemma (in order_syntax) greatest_Lower_above:
   242   "[| greatest i (Lower A); x \<in> A; A \<subseteq> carrier |] ==> i \<sqsubseteq> x"
   233   "[| greatest i (Lower A); x \<in> A; A \<subseteq> L |] ==> i \<sqsubseteq> x"
   243   by (unfold greatest_def) blast
   234   by (unfold greatest_def) blast
   244 
   235 
   245 
   236 
   246 subsubsection {* Supremum *}
   237 subsubsection {* Supremum *}
   247 
   238 
   248 lemma (in lattice) joinI:
   239 lemma (in lattice) joinI:
   249   "[| !!l. least l (Upper {x, y}) ==> P l; x \<in> carrier; y \<in> carrier |]
   240   "[| !!l. least l (Upper {x, y}) ==> P l; x \<in> L; y \<in> L |]
   250   ==> P (x \<squnion> y)"
   241   ==> P (x \<squnion> y)"
   251 proof (unfold join_def sup_def)
   242 proof (unfold join_def sup_def)
   252   assume L: "x \<in> carrier"  "y \<in> carrier"
   243   assume L: "x \<in> L"  "y \<in> L"
   253     and P: "!!l. least l (Upper {x, y}) ==> P l"
   244     and P: "!!l. least l (Upper {x, y}) ==> P l"
   254   with sup_of_two_exists obtain s where "least s (Upper {x, y})" by fast
   245   with sup_of_two_exists obtain s where "least s (Upper {x, y})" by fast
   255   with L show "P (THE l. least l (Upper {x, y}))"
   246   with L show "P (THE l. least l (Upper {x, y}))"
   256     by (fast intro: theI2 least_unique P)
   247     by (fast intro: theI2 least_unique P)
   257 qed
   248 qed
   258 
   249 
   259 lemma (in lattice) join_closed [simp]:
   250 lemma (in lattice) join_closed [simp]:
   260   "[| x \<in> carrier; y \<in> carrier |] ==> x \<squnion> y \<in> carrier"
   251   "[| x \<in> L; y \<in> L |] ==> x \<squnion> y \<in> L"
   261   by (rule joinI) (rule least_carrier)
   252   by (rule joinI) (rule least_closed)
   262 
   253 
   263 lemma (in partial_order) sup_of_singletonI:     (* only reflexivity needed ? *)
   254 lemma (in partial_order) sup_of_singletonI:     (* only reflexivity needed ? *)
   264   "x \<in> carrier ==> least x (Upper {x})"
   255   "x \<in> L ==> least x (Upper {x})"
   265   by (rule least_UpperI) fast+
   256   by (rule least_UpperI) fast+
   266 
   257 
   267 lemma (in partial_order) sup_of_singleton [simp]:
   258 lemma (in partial_order) sup_of_singleton [simp]:
   268   "x \<in> carrier ==> \<Squnion>{x} = x"
   259   "x \<in> L ==> \<Squnion>{x} = x"
   269   by (unfold sup_def) (blast intro: least_unique least_UpperI sup_of_singletonI)
   260   by (unfold sup_def) (blast intro: least_unique least_UpperI sup_of_singletonI)
   270 
   261 
   271 
   262 
   272 text {* Condition on @{text A}: supremum exists. *}
   263 text {* Condition on @{text A}: supremum exists. *}
   273 
   264 
   274 lemma (in lattice) sup_insertI:
   265 lemma (in lattice) sup_insertI:
   275   "[| !!s. least s (Upper (insert x A)) ==> P s;
   266   "[| !!s. least s (Upper (insert x A)) ==> P s;
   276   least a (Upper A); x \<in> carrier; A \<subseteq> carrier |]
   267   least a (Upper A); x \<in> L; A \<subseteq> L |]
   277   ==> P (\<Squnion>(insert x A))"
   268   ==> P (\<Squnion>(insert x A))"
   278 proof (unfold sup_def)
   269 proof (unfold sup_def)
   279   assume L: "x \<in> carrier"  "A \<subseteq> carrier"
   270   assume L: "x \<in> L"  "A \<subseteq> L"
   280     and P: "!!l. least l (Upper (insert x A)) ==> P l"
   271     and P: "!!l. least l (Upper (insert x A)) ==> P l"
   281     and least_a: "least a (Upper A)"
   272     and least_a: "least a (Upper A)"
   282   from least_a have La: "a \<in> carrier" by simp
   273   from least_a have La: "a \<in> L" by simp
   283   from L sup_of_two_exists least_a
   274   from L sup_of_two_exists least_a
   284   obtain s where least_s: "least s (Upper {a, x})" by blast
   275   obtain s where least_s: "least s (Upper {a, x})" by blast
   285   show "P (THE l. least l (Upper (insert x A)))"
   276   show "P (THE l. least l (Upper (insert x A)))"
   286   proof (rule theI2)
   277   proof (rule theI2)
   287     show "least s (Upper (insert x A))"
   278     show "least s (Upper (insert x A))"
   316 	  assume "z \<in> {x}"  (* FIXME "z = x"; declare specific elim rule for "insert x {}" (!?) *)
   307 	  assume "z \<in> {x}"  (* FIXME "z = x"; declare specific elim rule for "insert x {}" (!?) *)
   317           with y L show ?thesis by blast
   308           with y L show ?thesis by blast
   318 	qed
   309 	qed
   319       qed (rule Upper_closed [THEN subsetD])
   310       qed (rule Upper_closed [THEN subsetD])
   320     next
   311     next
   321       from L show "insert x A \<subseteq> carrier" by simp
   312       from L show "insert x A \<subseteq> L" by simp
   322       from least_s show "s \<in> carrier" by simp
   313       from least_s show "s \<in> L" by simp
   323     qed
   314     qed
   324   next
   315   next
   325     fix l
   316     fix l
   326     assume least_l: "least l (Upper (insert x A))"
   317     assume least_l: "least l (Upper (insert x A))"
   327     show "l = s"
   318     show "l = s"
   358             assume "z \<in> {x}"
   349             assume "z \<in> {x}"
   359             with y L show ?thesis by blast
   350             with y L show ?thesis by blast
   360           qed
   351           qed
   361         qed (rule Upper_closed [THEN subsetD])
   352         qed (rule Upper_closed [THEN subsetD])
   362       next
   353       next
   363         from L show "insert x A \<subseteq> carrier" by simp
   354         from L show "insert x A \<subseteq> L" by simp
   364         from least_s show "s \<in> carrier" by simp
   355         from least_s show "s \<in> L" by simp
   365       qed
   356       qed
   366     qed
   357     qed
   367   qed
   358   qed
   368 qed
   359 qed
   369 
   360 
   370 lemma (in lattice) finite_sup_least:
   361 lemma (in lattice) finite_sup_least:
   371   "[| finite A; A \<subseteq> carrier; A ~= {} |] ==> least (\<Squnion>A) (Upper A)"
   362   "[| finite A; A \<subseteq> L; A ~= {} |] ==> least (\<Squnion>A) (Upper A)"
   372 proof (induct set: Finites)
   363 proof (induct set: Finites)
   373   case empty
   364   case empty
   374   then show ?case by simp
   365   then show ?case by simp
   375 next
   366 next
   376   case (insert x A)
   367   case (insert x A)
   386   qed
   377   qed
   387 qed
   378 qed
   388 
   379 
   389 lemma (in lattice) finite_sup_insertI:
   380 lemma (in lattice) finite_sup_insertI:
   390   assumes P: "!!l. least l (Upper (insert x A)) ==> P l"
   381   assumes P: "!!l. least l (Upper (insert x A)) ==> P l"
   391     and xA: "finite A"  "x \<in> carrier"  "A \<subseteq> carrier"
   382     and xA: "finite A"  "x \<in> L"  "A \<subseteq> L"
   392   shows "P (\<Squnion> (insert x A))"
   383   shows "P (\<Squnion> (insert x A))"
   393 proof (cases "A = {}")
   384 proof (cases "A = {}")
   394   case True with P and xA show ?thesis
   385   case True with P and xA show ?thesis
   395     by (simp add: sup_of_singletonI)
   386     by (simp add: sup_of_singletonI)
   396 next
   387 next
   397   case False with P and xA show ?thesis
   388   case False with P and xA show ?thesis
   398     by (simp add: sup_insertI finite_sup_least)
   389     by (simp add: sup_insertI finite_sup_least)
   399 qed
   390 qed
   400 
   391 
   401 lemma (in lattice) finite_sup_closed:
   392 lemma (in lattice) finite_sup_closed:
   402   "[| finite A; A \<subseteq> carrier; A ~= {} |] ==> \<Squnion>A \<in> carrier"
   393   "[| finite A; A \<subseteq> L; A ~= {} |] ==> \<Squnion>A \<in> L"
   403 proof (induct set: Finites)
   394 proof (induct set: Finites)
   404   case empty then show ?case by simp
   395   case empty then show ?case by simp
   405 next
   396 next
   406   case insert then show ?case
   397   case insert then show ?case
   407     by - (rule finite_sup_insertI, simp_all)
   398     by - (rule finite_sup_insertI, simp_all)
   408 qed
   399 qed
   409 
   400 
   410 lemma (in lattice) join_left:
   401 lemma (in lattice) join_left:
   411   "[| x \<in> carrier; y \<in> carrier |] ==> x \<sqsubseteq> x \<squnion> y"
   402   "[| x \<in> L; y \<in> L |] ==> x \<sqsubseteq> x \<squnion> y"
   412   by (rule joinI [folded join_def]) (blast dest: least_mem)
   403   by (rule joinI [folded join_def]) (blast dest: least_mem)
   413 
   404 
   414 lemma (in lattice) join_right:
   405 lemma (in lattice) join_right:
   415   "[| x \<in> carrier; y \<in> carrier |] ==> y \<sqsubseteq> x \<squnion> y"
   406   "[| x \<in> L; y \<in> L |] ==> y \<sqsubseteq> x \<squnion> y"
   416   by (rule joinI [folded join_def]) (blast dest: least_mem)
   407   by (rule joinI [folded join_def]) (blast dest: least_mem)
   417 
   408 
   418 lemma (in lattice) sup_of_two_least:
   409 lemma (in lattice) sup_of_two_least:
   419   "[| x \<in> carrier; y \<in> carrier |] ==> least (\<Squnion>{x, y}) (Upper {x, y})"
   410   "[| x \<in> L; y \<in> L |] ==> least (\<Squnion>{x, y}) (Upper {x, y})"
   420 proof (unfold sup_def)
   411 proof (unfold sup_def)
   421   assume L: "x \<in> carrier"  "y \<in> carrier"
   412   assume L: "x \<in> L"  "y \<in> L"
   422   with sup_of_two_exists obtain s where "least s (Upper {x, y})" by fast
   413   with sup_of_two_exists obtain s where "least s (Upper {x, y})" by fast
   423   with L show "least (THE xa. least xa (Upper {x, y})) (Upper {x, y})"
   414   with L show "least (THE xa. least xa (Upper {x, y})) (Upper {x, y})"
   424   by (fast intro: theI2 least_unique)  (* blast fails *)
   415   by (fast intro: theI2 least_unique)  (* blast fails *)
   425 qed
   416 qed
   426 
   417 
   427 lemma (in lattice) join_le:
   418 lemma (in lattice) join_le:
   428   assumes sub: "x \<sqsubseteq> z"  "y \<sqsubseteq> z"
   419   assumes sub: "x \<sqsubseteq> z"  "y \<sqsubseteq> z"
   429     and L: "x \<in> carrier"  "y \<in> carrier"  "z \<in> carrier"
   420     and L: "x \<in> L"  "y \<in> L"  "z \<in> L"
   430   shows "x \<squnion> y \<sqsubseteq> z"
   421   shows "x \<squnion> y \<sqsubseteq> z"
   431 proof (rule joinI)
   422 proof (rule joinI)
   432   fix s
   423   fix s
   433   assume "least s (Upper {x, y})"
   424   assume "least s (Upper {x, y})"
   434   with sub L show "s \<sqsubseteq> z" by (fast elim: least_le intro: Upper_memI)
   425   with sub L show "s \<sqsubseteq> z" by (fast elim: least_le intro: Upper_memI)
   435 qed
   426 qed
   436 
   427 
   437 lemma (in lattice) join_assoc_lemma:
   428 lemma (in lattice) join_assoc_lemma:
   438   assumes L: "x \<in> carrier"  "y \<in> carrier"  "z \<in> carrier"
   429   assumes L: "x \<in> L"  "y \<in> L"  "z \<in> L"
   439   shows "x \<squnion> (y \<squnion> z) = \<Squnion>{x, y, z}"
   430   shows "x \<squnion> (y \<squnion> z) = \<Squnion>{x, y, z}"
   440 proof (rule finite_sup_insertI)
   431 proof (rule finite_sup_insertI)
   441   -- {* The textbook argument in Jacobson I, p 457 *}
   432   -- {* The textbook argument in Jacobson I, p 457 *}
   442   fix s
   433   fix s
   443   assume sup: "least s (Upper {x, y, z})"
   434   assume sup: "least s (Upper {x, y, z})"
   447       by (fastsimp intro!: join_le elim: least_Upper_above)
   438       by (fastsimp intro!: join_le elim: least_Upper_above)
   448   next
   439   next
   449     from sup L show "s \<sqsubseteq> x \<squnion> (y \<squnion> z)"
   440     from sup L show "s \<sqsubseteq> x \<squnion> (y \<squnion> z)"
   450     by (erule_tac least_le)
   441     by (erule_tac least_le)
   451       (blast intro!: Upper_memI intro: trans join_left join_right join_closed)
   442       (blast intro!: Upper_memI intro: trans join_left join_right join_closed)
   452   qed (simp_all add: L least_carrier [OF sup])
   443   qed (simp_all add: L least_closed [OF sup])
   453 qed (simp_all add: L)
   444 qed (simp_all add: L)
   454 
   445 
   455 lemma (in order_syntax) join_comm:
   446 lemma (in order_syntax) join_comm:
   456   "x \<squnion> y = y \<squnion> x"
   447   "x \<squnion> y = y \<squnion> x"
   457   by (unfold join_def) (simp add: insert_commute)
   448   by (unfold join_def) (simp add: insert_commute)
   458 
   449 
   459 lemma (in lattice) join_assoc:
   450 lemma (in lattice) join_assoc:
   460   assumes L: "x \<in> carrier"  "y \<in> carrier"  "z \<in> carrier"
   451   assumes L: "x \<in> L"  "y \<in> L"  "z \<in> L"
   461   shows "(x \<squnion> y) \<squnion> z = x \<squnion> (y \<squnion> z)"
   452   shows "(x \<squnion> y) \<squnion> z = x \<squnion> (y \<squnion> z)"
   462 proof -
   453 proof -
   463   have "(x \<squnion> y) \<squnion> z = z \<squnion> (x \<squnion> y)" by (simp only: join_comm)
   454   have "(x \<squnion> y) \<squnion> z = z \<squnion> (x \<squnion> y)" by (simp only: join_comm)
   464   also from L have "... = \<Squnion>{z, x, y}" by (simp add: join_assoc_lemma)
   455   also from L have "... = \<Squnion>{z, x, y}" by (simp add: join_assoc_lemma)
   465   also from L have "... = \<Squnion>{x, y, z}" by (simp add: insert_commute)
   456   also from L have "... = \<Squnion>{x, y, z}" by (simp add: insert_commute)
   469 
   460 
   470 
   461 
   471 subsubsection {* Infimum *}
   462 subsubsection {* Infimum *}
   472 
   463 
   473 lemma (in lattice) meetI:
   464 lemma (in lattice) meetI:
   474   "[| !!i. greatest i (Lower {x, y}) ==> P i;
   465   "[| !!i. greatest i (Lower {x, y}) ==> P i; x \<in> L; y \<in> L |]
   475   x \<in> carrier; y \<in> carrier |]
       
   476   ==> P (x \<sqinter> y)"
   466   ==> P (x \<sqinter> y)"
   477 proof (unfold meet_def inf_def)
   467 proof (unfold meet_def inf_def)
   478   assume L: "x \<in> carrier"  "y \<in> carrier"
   468   assume L: "x \<in> L"  "y \<in> L"
   479     and P: "!!g. greatest g (Lower {x, y}) ==> P g"
   469     and P: "!!g. greatest g (Lower {x, y}) ==> P g"
   480   with inf_of_two_exists obtain i where "greatest i (Lower {x, y})" by fast
   470   with inf_of_two_exists obtain i where "greatest i (Lower {x, y})" by fast
   481   with L show "P (THE g. greatest g (Lower {x, y}))"
   471   with L show "P (THE g. greatest g (Lower {x, y}))"
   482   by (fast intro: theI2 greatest_unique P)
   472   by (fast intro: theI2 greatest_unique P)
   483 qed
   473 qed
   484 
   474 
   485 lemma (in lattice) meet_closed [simp]:
   475 lemma (in lattice) meet_closed [simp]:
   486   "[| x \<in> carrier; y \<in> carrier |] ==> x \<sqinter> y \<in> carrier"
   476   "[| x \<in> L; y \<in> L |] ==> x \<sqinter> y \<in> L"
   487   by (rule meetI) (rule greatest_carrier)
   477   by (rule meetI) (rule greatest_closed)
   488 
   478 
   489 lemma (in partial_order) inf_of_singletonI:      (* only reflexivity needed ? *)
   479 lemma (in partial_order) inf_of_singletonI:      (* only reflexivity needed ? *)
   490   "x \<in> carrier ==> greatest x (Lower {x})"
   480   "x \<in> L ==> greatest x (Lower {x})"
   491   by (rule greatest_LowerI) fast+
   481   by (rule greatest_LowerI) fast+
   492 
   482 
   493 lemma (in partial_order) inf_of_singleton [simp]:
   483 lemma (in partial_order) inf_of_singleton [simp]:
   494   "x \<in> carrier ==> \<Sqinter> {x} = x"
   484   "x \<in> L ==> \<Sqinter> {x} = x"
   495   by (unfold inf_def) (blast intro: greatest_unique greatest_LowerI inf_of_singletonI)
   485   by (unfold inf_def) (blast intro: greatest_unique greatest_LowerI inf_of_singletonI)
   496 
   486 
   497 text {* Condition on A: infimum exists. *}
   487 text {* Condition on A: infimum exists. *}
   498 
   488 
   499 lemma (in lattice) inf_insertI:
   489 lemma (in lattice) inf_insertI:
   500   "[| !!i. greatest i (Lower (insert x A)) ==> P i;
   490   "[| !!i. greatest i (Lower (insert x A)) ==> P i;
   501   greatest a (Lower A); x \<in> carrier; A \<subseteq> carrier |]
   491   greatest a (Lower A); x \<in> L; A \<subseteq> L |]
   502   ==> P (\<Sqinter>(insert x A))"
   492   ==> P (\<Sqinter>(insert x A))"
   503 proof (unfold inf_def)
   493 proof (unfold inf_def)
   504   assume L: "x \<in> carrier"  "A \<subseteq> carrier"
   494   assume L: "x \<in> L"  "A \<subseteq> L"
   505     and P: "!!g. greatest g (Lower (insert x A)) ==> P g"
   495     and P: "!!g. greatest g (Lower (insert x A)) ==> P g"
   506     and greatest_a: "greatest a (Lower A)"
   496     and greatest_a: "greatest a (Lower A)"
   507   from greatest_a have La: "a \<in> carrier" by simp
   497   from greatest_a have La: "a \<in> L" by simp
   508   from L inf_of_two_exists greatest_a
   498   from L inf_of_two_exists greatest_a
   509   obtain i where greatest_i: "greatest i (Lower {a, x})" by blast
   499   obtain i where greatest_i: "greatest i (Lower {a, x})" by blast
   510   show "P (THE g. greatest g (Lower (insert x A)))"
   500   show "P (THE g. greatest g (Lower (insert x A)))"
   511   proof (rule theI2)
   501   proof (rule theI2)
   512     show "greatest i (Lower (insert x A))"
   502     show "greatest i (Lower (insert x A))"
   541           assume "z \<in> {x}"
   531           assume "z \<in> {x}"
   542           with y L show ?thesis by blast
   532           with y L show ?thesis by blast
   543 	qed
   533 	qed
   544       qed (rule Lower_closed [THEN subsetD])
   534       qed (rule Lower_closed [THEN subsetD])
   545     next
   535     next
   546       from L show "insert x A \<subseteq> carrier" by simp
   536       from L show "insert x A \<subseteq> L" by simp
   547       from greatest_i show "i \<in> carrier" by simp
   537       from greatest_i show "i \<in> L" by simp
   548     qed
   538     qed
   549   next
   539   next
   550     fix g
   540     fix g
   551     assume greatest_g: "greatest g (Lower (insert x A))"
   541     assume greatest_g: "greatest g (Lower (insert x A))"
   552     show "g = i"
   542     show "g = i"
   583             assume "z \<in> {x}"
   573             assume "z \<in> {x}"
   584             with y L show ?thesis by blast
   574             with y L show ?thesis by blast
   585 	  qed
   575 	  qed
   586         qed (rule Lower_closed [THEN subsetD])
   576         qed (rule Lower_closed [THEN subsetD])
   587       next
   577       next
   588         from L show "insert x A \<subseteq> carrier" by simp
   578         from L show "insert x A \<subseteq> L" by simp
   589         from greatest_i show "i \<in> carrier" by simp
   579         from greatest_i show "i \<in> L" by simp
   590       qed
   580       qed
   591     qed
   581     qed
   592   qed
   582   qed
   593 qed
   583 qed
   594 
   584 
   595 lemma (in lattice) finite_inf_greatest:
   585 lemma (in lattice) finite_inf_greatest:
   596   "[| finite A; A \<subseteq> carrier; A ~= {} |] ==> greatest (\<Sqinter>A) (Lower A)"
   586   "[| finite A; A \<subseteq> L; A ~= {} |] ==> greatest (\<Sqinter>A) (Lower A)"
   597 proof (induct set: Finites)
   587 proof (induct set: Finites)
   598   case empty then show ?case by simp
   588   case empty then show ?case by simp
   599 next
   589 next
   600   case (insert x A)
   590   case (insert x A)
   601   show ?case
   591   show ?case
   611   qed
   601   qed
   612 qed
   602 qed
   613 
   603 
   614 lemma (in lattice) finite_inf_insertI:
   604 lemma (in lattice) finite_inf_insertI:
   615   assumes P: "!!i. greatest i (Lower (insert x A)) ==> P i"
   605   assumes P: "!!i. greatest i (Lower (insert x A)) ==> P i"
   616     and xA: "finite A"  "x \<in> carrier"  "A \<subseteq> carrier"
   606     and xA: "finite A"  "x \<in> L"  "A \<subseteq> L"
   617   shows "P (\<Sqinter> (insert x A))"
   607   shows "P (\<Sqinter> (insert x A))"
   618 proof (cases "A = {}")
   608 proof (cases "A = {}")
   619   case True with P and xA show ?thesis
   609   case True with P and xA show ?thesis
   620     by (simp add: inf_of_singletonI)
   610     by (simp add: inf_of_singletonI)
   621 next
   611 next
   622   case False with P and xA show ?thesis
   612   case False with P and xA show ?thesis
   623     by (simp add: inf_insertI finite_inf_greatest)
   613     by (simp add: inf_insertI finite_inf_greatest)
   624 qed
   614 qed
   625 
   615 
   626 lemma (in lattice) finite_inf_closed:
   616 lemma (in lattice) finite_inf_closed:
   627   "[| finite A; A \<subseteq> carrier; A ~= {} |] ==> \<Sqinter>A \<in> carrier"
   617   "[| finite A; A \<subseteq> L; A ~= {} |] ==> \<Sqinter>A \<in> L"
   628 proof (induct set: Finites)
   618 proof (induct set: Finites)
   629   case empty then show ?case by simp
   619   case empty then show ?case by simp
   630 next
   620 next
   631   case insert then show ?case
   621   case insert then show ?case
   632     by (rule_tac finite_inf_insertI) (simp_all)
   622     by (rule_tac finite_inf_insertI) (simp_all)
   633 qed
   623 qed
   634 
   624 
   635 lemma (in lattice) meet_left:
   625 lemma (in lattice) meet_left:
   636   "[| x \<in> carrier; y \<in> carrier |] ==> x \<sqinter> y \<sqsubseteq> x"
   626   "[| x \<in> L; y \<in> L |] ==> x \<sqinter> y \<sqsubseteq> x"
   637   by (rule meetI [folded meet_def]) (blast dest: greatest_mem)
   627   by (rule meetI [folded meet_def]) (blast dest: greatest_mem)
   638 
   628 
   639 lemma (in lattice) meet_right:
   629 lemma (in lattice) meet_right:
   640   "[| x \<in> carrier; y \<in> carrier |] ==> x \<sqinter> y \<sqsubseteq> y"
   630   "[| x \<in> L; y \<in> L |] ==> x \<sqinter> y \<sqsubseteq> y"
   641   by (rule meetI [folded meet_def]) (blast dest: greatest_mem)
   631   by (rule meetI [folded meet_def]) (blast dest: greatest_mem)
   642 
   632 
   643 lemma (in lattice) inf_of_two_greatest:
   633 lemma (in lattice) inf_of_two_greatest:
   644   "[| x \<in> carrier; y \<in> carrier |] ==>
   634   "[| x \<in> L; y \<in> L |] ==> greatest (\<Sqinter> {x, y}) (Lower {x, y})"
   645   greatest (\<Sqinter> {x, y}) (Lower {x, y})"
       
   646 proof (unfold inf_def)
   635 proof (unfold inf_def)
   647   assume L: "x \<in> carrier"  "y \<in> carrier"
   636   assume L: "x \<in> L"  "y \<in> L"
   648   with inf_of_two_exists obtain s where "greatest s (Lower {x, y})" by fast
   637   with inf_of_two_exists obtain s where "greatest s (Lower {x, y})" by fast
   649   with L
   638   with L
   650   show "greatest (THE xa. greatest xa (Lower {x, y})) (Lower {x, y})"
   639   show "greatest (THE xa. greatest xa (Lower {x, y})) (Lower {x, y})"
   651   by (fast intro: theI2 greatest_unique)  (* blast fails *)
   640   by (fast intro: theI2 greatest_unique)  (* blast fails *)
   652 qed
   641 qed
   653 
   642 
   654 lemma (in lattice) meet_le:
   643 lemma (in lattice) meet_le:
   655   assumes sub: "z \<sqsubseteq> x"  "z \<sqsubseteq> y"
   644   assumes sub: "z \<sqsubseteq> x"  "z \<sqsubseteq> y"
   656     and L: "x \<in> carrier"  "y \<in> carrier"  "z \<in> carrier"
   645     and L: "x \<in> L"  "y \<in> L"  "z \<in> L"
   657   shows "z \<sqsubseteq> x \<sqinter> y"
   646   shows "z \<sqsubseteq> x \<sqinter> y"
   658 proof (rule meetI)
   647 proof (rule meetI)
   659   fix i
   648   fix i
   660   assume "greatest i (Lower {x, y})"
   649   assume "greatest i (Lower {x, y})"
   661   with sub L show "z \<sqsubseteq> i" by (fast elim: greatest_le intro: Lower_memI)
   650   with sub L show "z \<sqsubseteq> i" by (fast elim: greatest_le intro: Lower_memI)
   662 qed
   651 qed
   663 
   652 
   664 lemma (in lattice) meet_assoc_lemma:
   653 lemma (in lattice) meet_assoc_lemma:
   665   assumes L: "x \<in> carrier"  "y \<in> carrier"  "z \<in> carrier"
   654   assumes L: "x \<in> L"  "y \<in> L"  "z \<in> L"
   666   shows "x \<sqinter> (y \<sqinter> z) = \<Sqinter>{x, y, z}"
   655   shows "x \<sqinter> (y \<sqinter> z) = \<Sqinter>{x, y, z}"
   667 proof (rule finite_inf_insertI)
   656 proof (rule finite_inf_insertI)
   668   txt {* The textbook argument in Jacobson I, p 457 *}
   657   txt {* The textbook argument in Jacobson I, p 457 *}
   669   fix i
   658   fix i
   670   assume inf: "greatest i (Lower {x, y, z})"
   659   assume inf: "greatest i (Lower {x, y, z})"
   674       by (fastsimp intro!: meet_le elim: greatest_Lower_above)
   663       by (fastsimp intro!: meet_le elim: greatest_Lower_above)
   675   next
   664   next
   676     from inf L show "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> i"
   665     from inf L show "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> i"
   677     by (erule_tac greatest_le)
   666     by (erule_tac greatest_le)
   678       (blast intro!: Lower_memI intro: trans meet_left meet_right meet_closed)
   667       (blast intro!: Lower_memI intro: trans meet_left meet_right meet_closed)
   679   qed (simp_all add: L greatest_carrier [OF inf])
   668   qed (simp_all add: L greatest_closed [OF inf])
   680 qed (simp_all add: L)
   669 qed (simp_all add: L)
   681 
   670 
   682 lemma (in order_syntax) meet_comm:
   671 lemma (in order_syntax) meet_comm:
   683   "x \<sqinter> y = y \<sqinter> x"
   672   "x \<sqinter> y = y \<sqinter> x"
   684   by (unfold meet_def) (simp add: insert_commute)
   673   by (unfold meet_def) (simp add: insert_commute)
   685 
   674 
   686 lemma (in lattice) meet_assoc:
   675 lemma (in lattice) meet_assoc:
   687   assumes L: "x \<in> carrier"  "y \<in> carrier"  "z \<in> carrier"
   676   assumes L: "x \<in> L"  "y \<in> L"  "z \<in> L"
   688   shows "(x \<sqinter> y) \<sqinter> z = x \<sqinter> (y \<sqinter> z)"
   677   shows "(x \<sqinter> y) \<sqinter> z = x \<sqinter> (y \<sqinter> z)"
   689 proof -
   678 proof -
   690   have "(x \<sqinter> y) \<sqinter> z = z \<sqinter> (x \<sqinter> y)" by (simp only: meet_comm)
   679   have "(x \<sqinter> y) \<sqinter> z = z \<sqinter> (x \<sqinter> y)" by (simp only: meet_comm)
   691   also from L have "... = \<Sqinter> {z, x, y}" by (simp add: meet_assoc_lemma)
   680   also from L have "... = \<Sqinter> {z, x, y}" by (simp add: meet_assoc_lemma)
   692   also from L have "... = \<Sqinter> {x, y, z}" by (simp add: insert_commute)
   681   also from L have "... = \<Sqinter> {x, y, z}" by (simp add: insert_commute)
   696 
   685 
   697 
   686 
   698 subsection {* Total Orders *}
   687 subsection {* Total Orders *}
   699 
   688 
   700 locale total_order = lattice +
   689 locale total_order = lattice +
   701   assumes total: "[| x \<in> carrier; y \<in> carrier |] ==> x \<sqsubseteq> y | y \<sqsubseteq> x"
   690   assumes total: "[| x \<in> L; y \<in> L |] ==> x \<sqsubseteq> y | y \<sqsubseteq> x"
   702 
   691 
   703 abbreviation (in total_order)
   692 abbreviation (in total_order)
   704   less (infixl "\<sqsubset>" 50) "less == order_syntax.less le"
   693   less (infixl "\<sqsubset>" 50) "less == order_syntax.less le"
   705 abbreviation (in total_order)
   694 abbreviation (in total_order)
   706   Upper where "Upper == order_syntax.Upper carrier le"
   695   Upper where "Upper == order_syntax.Upper L le"
   707 abbreviation (in total_order)
   696 abbreviation (in total_order)
   708   Lower where "Lower == order_syntax.Lower carrier le"
   697   Lower where "Lower == order_syntax.Lower L le"
   709 abbreviation (in total_order)
   698 abbreviation (in total_order)
   710   least where "least == order_syntax.least carrier le"
   699   least where "least == order_syntax.least L le"
   711 abbreviation (in total_order)
   700 abbreviation (in total_order)
   712   greatest where "greatest == order_syntax.greatest carrier le"
   701   greatest where "greatest == order_syntax.greatest L le"
   713 abbreviation (in total_order)
   702 abbreviation (in total_order)
   714   sup ("\<Squnion>") (* FIXME precedence *) "sup == order_syntax.sup carrier le"
   703   sup ("\<Squnion>_" [90] 90) "sup == order_syntax.sup L le"
   715 abbreviation (in total_order)
   704 abbreviation (in total_order)
   716   inf ("\<Sqinter>") (* FIXME precedence *) "inf == order_syntax.inf carrier le"
   705   inf ("\<Sqinter>_" [90] 90) "inf == order_syntax.inf L le"
   717 abbreviation (in total_order)
   706 abbreviation (in total_order)
   718   join (infixl "\<squnion>" 65) "join == order_syntax.join carrier le"
   707   join (infixl "\<squnion>" 65) "join == order_syntax.join L le"
   719 abbreviation (in total_order)
   708 abbreviation (in total_order)
   720   meet (infixl "\<sqinter>" 70) "meet == order_syntax.meet carrier le"
   709   meet (infixl "\<sqinter>" 70) "meet == order_syntax.meet L le"
   721 
   710 
   722 text {* Introduction rule: the usual definition of total order *}
   711 text {* Introduction rule: the usual definition of total order *}
   723 
   712 
   724 lemma (in partial_order) total_orderI:
   713 lemma (in partial_order) total_orderI:
   725   assumes total: "!!x y. [| x \<in> carrier; y \<in> carrier |] ==> x \<sqsubseteq> y | y \<sqsubseteq> x"
   714   assumes total: "!!x y. [| x \<in> L; y \<in> L |] ==> x \<sqsubseteq> y | y \<sqsubseteq> x"
   726   shows "total_order carrier le"
   715   shows "total_order L le"
   727 proof intro_locales
   716 proof intro_locales
   728   show "lattice_axioms carrier le"
   717   show "lattice_axioms L le"
   729   proof (rule lattice_axioms.intro)
   718   proof (rule lattice_axioms.intro)
   730     fix x y
   719     fix x y
   731     assume L: "x \<in> carrier"  "y \<in> carrier"
   720     assume L: "x \<in> L"  "y \<in> L"
   732     show "EX s. least s (Upper {x, y})"
   721     show "EX s. least s (Upper {x, y})"
   733     proof -
   722     proof -
   734       note total L
   723       note total L
   735       moreover
   724       moreover
   736       {
   725       {
   746       }
   735       }
   747       ultimately show ?thesis by blast
   736       ultimately show ?thesis by blast
   748     qed
   737     qed
   749   next
   738   next
   750     fix x y
   739     fix x y
   751     assume L: "x \<in> carrier"  "y \<in> carrier"
   740     assume L: "x \<in> L"  "y \<in> L"
   752     show "EX i. greatest i (Lower {x, y})"
   741     show "EX i. greatest i (Lower {x, y})"
   753     proof -
   742     proof -
   754       note total L
   743       note total L
   755       moreover
   744       moreover
   756       {
   745       {
   772 
   761 
   773 subsection {* Complete lattices *}
   762 subsection {* Complete lattices *}
   774 
   763 
   775 locale complete_lattice = lattice +
   764 locale complete_lattice = lattice +
   776   assumes sup_exists:
   765   assumes sup_exists:
   777     "[| A \<subseteq> carrier |] ==> EX s. order_syntax.least carrier le s (order_syntax.Upper carrier le A)"
   766     "[| A \<subseteq> L |] ==> EX s. order_syntax.least L le s (order_syntax.Upper L le A)"
   778     and inf_exists:
   767     and inf_exists:
   779     "[| A \<subseteq> carrier |] ==> EX i. order_syntax.greatest carrier le i (order_syntax.Lower carrier le A)"
   768     "[| A \<subseteq> L |] ==> EX i. order_syntax.greatest L le i (order_syntax.Lower L le A)"
   780 
   769 
   781 abbreviation (in complete_lattice)
   770 abbreviation (in complete_lattice)
   782   less (infixl "\<sqsubset>" 50) "less == order_syntax.less le"
   771   less (infixl "\<sqsubset>" 50) "less == order_syntax.less le"
   783 abbreviation (in complete_lattice)
   772 abbreviation (in complete_lattice)
   784   Upper where "Upper == order_syntax.Upper carrier le"
   773   Upper where "Upper == order_syntax.Upper L le"
   785 abbreviation (in complete_lattice)
   774 abbreviation (in complete_lattice)
   786   Lower where "Lower == order_syntax.Lower carrier le"
   775   Lower where "Lower == order_syntax.Lower L le"
   787 abbreviation (in complete_lattice)
   776 abbreviation (in complete_lattice)
   788   least where "least == order_syntax.least carrier le"
   777   least where "least == order_syntax.least L le"
   789 abbreviation (in complete_lattice)
   778 abbreviation (in complete_lattice)
   790   greatest where "greatest == order_syntax.greatest carrier le"
   779   greatest where "greatest == order_syntax.greatest L le"
   791 abbreviation (in complete_lattice)
   780 abbreviation (in complete_lattice)
   792   sup ("\<Squnion>") (* FIXME precedence *) "sup == order_syntax.sup carrier le"
   781   sup ("\<Squnion>_" [90] 90) "sup == order_syntax.sup L le"
   793 abbreviation (in complete_lattice)
   782 abbreviation (in complete_lattice)
   794   inf ("\<Sqinter>") (* FIXME precedence *) "inf == order_syntax.inf carrier le"
   783   inf ("\<Sqinter>_" [90] 90) "inf == order_syntax.inf L le"
   795 abbreviation (in complete_lattice)
   784 abbreviation (in complete_lattice)
   796   join (infixl "\<squnion>" 65) "join == order_syntax.join carrier le"
   785   join (infixl "\<squnion>" 65) "join == order_syntax.join L le"
   797 abbreviation (in complete_lattice)
   786 abbreviation (in complete_lattice)
   798   meet (infixl "\<sqinter>" 70) "meet == order_syntax.meet carrier le"
   787   meet (infixl "\<sqinter>" 70) "meet == order_syntax.meet L le"
   799 
   788 
   800 text {* Introduction rule: the usual definition of complete lattice *}
   789 text {* Introduction rule: the usual definition of complete lattice *}
   801 
   790 
   802 lemma (in partial_order) complete_latticeI:
   791 lemma (in partial_order) complete_latticeI:
   803   assumes sup_exists:
   792   assumes sup_exists:
   804     "!!A. [| A \<subseteq> carrier |] ==> EX s. least s (Upper A)"
   793     "!!A. [| A \<subseteq> L |] ==> EX s. least s (Upper A)"
   805     and inf_exists:
   794     and inf_exists:
   806     "!!A. [| A \<subseteq> carrier |] ==> EX i. greatest i (Lower A)"
   795     "!!A. [| A \<subseteq> L |] ==> EX i. greatest i (Lower A)"
   807   shows "complete_lattice carrier le"
   796   shows "complete_lattice L le"
   808 proof intro_locales
   797 proof intro_locales
   809   show "lattice_axioms carrier le"
   798   show "lattice_axioms L le"
   810     by (rule lattice_axioms.intro) (blast intro: sup_exists inf_exists)+
   799     by (rule lattice_axioms.intro) (blast intro: sup_exists inf_exists)+
   811 qed (assumption | rule complete_lattice_axioms.intro)+
   800 qed (assumption | rule complete_lattice_axioms.intro)+
   812 
   801 
   813 definition (in order_syntax)
   802 definition (in order_syntax)
   814   top ("\<top>")
   803   top ("\<top>")
   815   "\<top> == sup carrier"
   804   "\<top> == sup L"
   816 
   805 
   817 definition (in order_syntax)
   806 definition (in order_syntax)
   818   bottom ("\<bottom>")
   807   bottom ("\<bottom>")
   819   "\<bottom> == inf carrier"
   808   "\<bottom> == inf L"
   820 
   809 
   821 abbreviation (in partial_order)
   810 abbreviation (in partial_order)
   822   top ("\<top>") "top == order_syntax.top carrier le"
   811   top ("\<top>") "top == order_syntax.top L le"
   823 abbreviation (in partial_order)
   812 abbreviation (in partial_order)
   824   bottom ("\<bottom>") "bottom == order_syntax.bottom carrier le"
   813   bottom ("\<bottom>") "bottom == order_syntax.bottom L le"
   825 abbreviation (in lattice)
   814 abbreviation (in lattice)
   826   top ("\<top>") "top == order_syntax.top carrier le"
   815   top ("\<top>") "top == order_syntax.top L le"
   827 abbreviation (in lattice)
   816 abbreviation (in lattice)
   828   bottom ("\<bottom>") "bottom == order_syntax.bottom carrier le"
   817   bottom ("\<bottom>") "bottom == order_syntax.bottom L le"
   829 abbreviation (in total_order)
   818 abbreviation (in total_order)
   830   top ("\<top>") "top == order_syntax.top carrier le"
   819   top ("\<top>") "top == order_syntax.top L le"
   831 abbreviation (in total_order)
   820 abbreviation (in total_order)
   832   bottom ("\<bottom>") "bottom == order_syntax.bottom carrier le"
   821   bottom ("\<bottom>") "bottom == order_syntax.bottom L le"
   833 abbreviation (in complete_lattice)
   822 abbreviation (in complete_lattice)
   834   top ("\<top>") "top == order_syntax.top carrier le"
   823   top ("\<top>") "top == order_syntax.top L le"
   835 abbreviation (in complete_lattice)
   824 abbreviation (in complete_lattice)
   836   bottom ("\<bottom>") "bottom == order_syntax.bottom carrier le"
   825   bottom ("\<bottom>") "bottom == order_syntax.bottom L le"
   837 
   826 
   838 
   827 
   839 lemma (in complete_lattice) supI:
   828 lemma (in complete_lattice) supI:
   840   "[| !!l. least l (Upper A) ==> P l; A \<subseteq> carrier |]
   829   "[| !!l. least l (Upper A) ==> P l; A \<subseteq> L |]
   841   ==> P (\<Squnion>A)"
   830   ==> P (\<Squnion>A)"
   842 proof (unfold sup_def)
   831 proof (unfold sup_def)
   843   assume L: "A \<subseteq> carrier"
   832   assume L: "A \<subseteq> L"
   844     and P: "!!l. least l (Upper A) ==> P l"
   833     and P: "!!l. least l (Upper A) ==> P l"
   845   with sup_exists obtain s where "least s (Upper A)" by blast
   834   with sup_exists obtain s where "least s (Upper A)" by blast
   846   with L show "P (THE l. least l (Upper A))"
   835   with L show "P (THE l. least l (Upper A))"
   847   by (fast intro: theI2 least_unique P)
   836   by (fast intro: theI2 least_unique P)
   848 qed
   837 qed
   849 
   838 
   850 lemma (in complete_lattice) sup_closed [simp]:
   839 lemma (in complete_lattice) sup_closed [simp]:
   851   "A \<subseteq> carrier ==> \<Squnion>A \<in> carrier"
   840   "A \<subseteq> L ==> \<Squnion>A \<in> L"
   852   by (rule supI) simp_all
   841   by (rule supI) simp_all
   853 
   842 
   854 lemma (in complete_lattice) top_closed [simp, intro]:
   843 lemma (in complete_lattice) top_closed [simp, intro]:
   855   "\<top> \<in> carrier"
   844   "\<top> \<in> L"
   856   by (unfold top_def) simp
   845   by (unfold top_def) simp
   857 
   846 
   858 lemma (in complete_lattice) infI:
   847 lemma (in complete_lattice) infI:
   859   "[| !!i. greatest i (Lower A) ==> P i; A \<subseteq> carrier |]
   848   "[| !!i. greatest i (Lower A) ==> P i; A \<subseteq> L |]
   860   ==> P (\<Sqinter>A)"
   849   ==> P (\<Sqinter>A)"
   861 proof (unfold inf_def)
   850 proof (unfold inf_def)
   862   assume L: "A \<subseteq> carrier"
   851   assume L: "A \<subseteq> L"
   863     and P: "!!l. greatest l (Lower A) ==> P l"
   852     and P: "!!l. greatest l (Lower A) ==> P l"
   864   with inf_exists obtain s where "greatest s (Lower A)" by blast
   853   with inf_exists obtain s where "greatest s (Lower A)" by blast
   865   with L show "P (THE l. greatest l (Lower A))"
   854   with L show "P (THE l. greatest l (Lower A))"
   866   by (fast intro: theI2 greatest_unique P)
   855   by (fast intro: theI2 greatest_unique P)
   867 qed
   856 qed
   868 
   857 
   869 lemma (in complete_lattice) inf_closed [simp]:
   858 lemma (in complete_lattice) inf_closed [simp]:
   870   "A \<subseteq> carrier ==> \<Sqinter>A \<in> carrier"
   859   "A \<subseteq> L ==> \<Sqinter>A \<in> L"
   871   by (rule infI) simp_all
   860   by (rule infI) simp_all
   872 
   861 
   873 lemma (in complete_lattice) bottom_closed [simp, intro]:
   862 lemma (in complete_lattice) bottom_closed [simp, intro]:
   874   "\<bottom> \<in> carrier"
   863   "\<bottom> \<in> L"
   875   by (unfold bottom_def) simp
   864   by (unfold bottom_def) simp
   876 
   865 
   877 text {* Jacobson: Theorem 8.1 *}
   866 text {* Jacobson: Theorem 8.1 *}
   878 
   867 
   879 lemma (in order_syntax) Lower_empty [simp]:
   868 lemma (in order_syntax) Lower_empty [simp]:
   880   "Lower {} = carrier"
   869   "Lower {} = L"
   881   by (unfold Lower_def) simp
   870   by (unfold Lower_def) simp
   882 
   871 
   883 lemma (in order_syntax) Upper_empty [simp]:
   872 lemma (in order_syntax) Upper_empty [simp]:
   884   "Upper {} = carrier"
   873   "Upper {} = L"
   885   by (unfold Upper_def) simp
   874   by (unfold Upper_def) simp
   886 
   875 
   887 theorem (in partial_order) complete_lattice_criterion1:
   876 theorem (in partial_order) complete_lattice_criterion1:
   888   assumes top_exists: "EX g. greatest g (carrier)"
   877   assumes top_exists: "EX g. greatest g L"
   889     and inf_exists:
   878     and inf_exists:
   890       "!!A. [| A \<subseteq> carrier; A ~= {} |] ==> EX i. greatest i (Lower A)"
   879       "!!A. [| A \<subseteq> L; A ~= {} |] ==> EX i. greatest i (Lower A)"
   891   shows "complete_lattice carrier le"
   880   shows "complete_lattice L le"
   892 proof (rule complete_latticeI)
   881 proof (rule complete_latticeI)
   893   from top_exists obtain top where top: "greatest top (carrier)" ..
   882   from top_exists obtain top where top: "greatest top L" ..
   894   fix A
   883   fix A
   895   assume L: "A \<subseteq> carrier"
   884   assume L: "A \<subseteq> L"
   896   let ?B = "Upper A"
   885   let ?B = "Upper A"
   897   from L top have "top \<in> ?B" by (fast intro!: Upper_memI intro: greatest_le)
   886   from L top have "top \<in> ?B" by (fast intro!: Upper_memI intro: greatest_le)
   898   then have B_non_empty: "?B ~= {}" by fast
   887   then have B_non_empty: "?B ~= {}" by fast
   899   have B_L: "?B \<subseteq> carrier" by simp
   888   have B_L: "?B \<subseteq> L" by simp
   900   from inf_exists [OF B_L B_non_empty]
   889   from inf_exists [OF B_L B_non_empty]
   901   obtain b where b_inf_B: "greatest b (Lower ?B)" ..
   890   obtain b where b_inf_B: "greatest b (Lower ?B)" ..
   902   have "least b (Upper A)"
   891   have "least b (Upper A)"
   903 apply (rule least_UpperI)
   892 apply (rule least_UpperI)
   904    apply (rule greatest_le [where A = "Lower ?B"])
   893    apply (rule greatest_le [where A = "Lower ?B"])
   909     apply (rule L)
   898     apply (rule L)
   910    apply (fast intro: L [THEN subsetD])
   899    apply (fast intro: L [THEN subsetD])
   911   apply (erule greatest_Lower_above [OF b_inf_B])
   900   apply (erule greatest_Lower_above [OF b_inf_B])
   912   apply simp
   901   apply simp
   913  apply (rule L)
   902  apply (rule L)
   914 apply (rule greatest_carrier [OF b_inf_B]) (* rename rule: _closed *)
   903 apply (rule greatest_closed [OF b_inf_B]) (* rename rule: _closed *)
   915 done
   904 done
   916   then show "EX s. least s (Upper A)" ..
   905   then show "EX s. least s (Upper A)" ..
   917 next
   906 next
   918   fix A
   907   fix A
   919   assume L: "A \<subseteq> carrier"
   908   assume L: "A \<subseteq> L"
   920   show "EX i. greatest i (Lower A)"
   909   show "EX i. greatest i (Lower A)"
   921   proof (cases "A = {}")
   910   proof (cases "A = {}")
   922     case True then show ?thesis
   911     case True then show ?thesis
   923       by (simp add: top_exists)
   912       by (simp add: top_exists)
   924   next
   913   next
   934 
   923 
   935 subsubsection {* Powerset of a Set is a Complete Lattice *}
   924 subsubsection {* Powerset of a Set is a Complete Lattice *}
   936 
   925 
   937 theorem powerset_is_complete_lattice:
   926 theorem powerset_is_complete_lattice:
   938   "complete_lattice (Pow A) (op \<subseteq>)"
   927   "complete_lattice (Pow A) (op \<subseteq>)"
   939   (is "complete_lattice ?car ?le")
   928   (is "complete_lattice ?L ?le")
   940 proof (rule partial_order.complete_latticeI)
   929 proof (rule partial_order.complete_latticeI)
   941   show "partial_order ?car ?le"
   930   show "partial_order ?L ?le"
   942     by (rule partial_order.intro) auto
   931     by (rule partial_order.intro) auto
   943 next
   932 next
   944   fix B
   933   fix B
   945   assume "B \<subseteq> ?car"
   934   assume "B \<subseteq> ?L"
   946   then have "order_syntax.least ?car ?le (\<Union> B) (order_syntax.Upper ?car ?le B)"
   935   then have "order_syntax.least ?L ?le (\<Union> B) (order_syntax.Upper ?L ?le B)"
   947     by (fastsimp intro!: order_syntax.least_UpperI simp: order_syntax.Upper_def)
   936     by (fastsimp intro!: order_syntax.least_UpperI simp: order_syntax.Upper_def)
   948   then show "EX s. order_syntax.least ?car ?le s (order_syntax.Upper ?car ?le B)" ..
   937   then show "EX s. order_syntax.least ?L ?le s (order_syntax.Upper ?L ?le B)" ..
   949 next
   938 next
   950   fix B
   939   fix B
   951   assume "B \<subseteq> ?car"
   940   assume "B \<subseteq> ?L"
   952   then have "order_syntax.greatest ?car ?le (\<Inter> B \<inter> A) (order_syntax.Lower ?car ?le B)"
   941   then have "order_syntax.greatest ?L ?le (\<Inter> B \<inter> A) (order_syntax.Lower ?L ?le B)"
   953     txt {* @{term "\<Inter> B"} is not the infimum of @{term B}:
   942     txt {* @{term "\<Inter> B"} is not the infimum of @{term B}:
   954       @{term "\<Inter> {} = UNIV"} which is in general bigger than @{term "A"}! *}
   943       @{term "\<Inter> {} = UNIV"} which is in general bigger than @{term "A"}! *}
   955     by (fastsimp intro!: order_syntax.greatest_LowerI simp: order_syntax.Lower_def)
   944     by (fastsimp intro!: order_syntax.greatest_LowerI simp: order_syntax.Lower_def)
   956   then show "EX i. order_syntax.greatest ?car ?le i (order_syntax.Lower ?car ?le B)" ..
   945   then show "EX i. order_syntax.greatest ?L ?le i (order_syntax.Lower ?L ?le B)" ..
   957 qed
   946 qed
   958 
   947 
   959 text {* An other example, that of the lattice of subgroups of a group,
   948 text {* An other example, that of the lattice of subgroups of a group,
   960   can be found in Group theory (Section~\ref{sec:subgroup-lattice}). *}
   949   can be found in Group theory (Section~\ref{sec:subgroup-lattice}). *}
   961 
   950