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