src/HOL/Lattice/CompleteLattice.thy
author wenzelm
Wed, 30 Dec 2015 18:25:39 +0100
changeset 61986 2461779da2b8
parent 61983 8fb53badad99
child 76987 4c275405faae
permissions -rw-r--r--
isabelle update_cartouches -c -t;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/Lattice/CompleteLattice.thy
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
     2
    Author:     Markus Wenzel, TU Muenchen
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
     3
*)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
     4
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
     5
section \<open>Complete lattices\<close>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
     6
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 12399
diff changeset
     7
theory CompleteLattice imports Lattice begin
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
     8
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
     9
subsection \<open>Complete lattice operations\<close>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    10
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
    11
text \<open>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    12
  A \emph{complete lattice} is a partial order with general
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    13
  (infinitary) infimum of any set of elements.  General supremum
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    14
  exists as well, as a consequence of the connection of infinitary
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    15
  bounds (see \S\ref{sec:connect-bounds}).
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
    16
\<close>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    17
35317
d57da4abb47d dropped axclass; dropped Id
haftmann
parents: 25474
diff changeset
    18
class complete_lattice =
d57da4abb47d dropped axclass; dropped Id
haftmann
parents: 25474
diff changeset
    19
  assumes ex_Inf: "\<exists>inf. is_Inf A inf"
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    20
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    21
theorem ex_Sup: "\<exists>sup::'a::complete_lattice. is_Sup A sup"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    22
proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    23
  from ex_Inf obtain sup where "is_Inf {b. \<forall>a\<in>A. a \<sqsubseteq> b} sup" by blast
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
    24
  then have "is_Sup A sup" by (rule Inf_Sup)
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
    25
  then show ?thesis ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    26
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    27
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
    28
text \<open>
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
    29
  The general \<open>\<Sqinter>\<close> (meet) and \<open>\<Squnion>\<close> (join) operations select
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    30
  such infimum and supremum elements.
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
    31
\<close>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    32
19736
wenzelm
parents: 16417
diff changeset
    33
definition
61983
8fb53badad99 more symbols;
wenzelm
parents: 58879
diff changeset
    34
  Meet :: "'a::complete_lattice set \<Rightarrow> 'a"  ("\<Sqinter>_" [90] 90) where
8fb53badad99 more symbols;
wenzelm
parents: 58879
diff changeset
    35
  "\<Sqinter>A = (THE inf. is_Inf A inf)"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    36
definition
61983
8fb53badad99 more symbols;
wenzelm
parents: 58879
diff changeset
    37
  Join :: "'a::complete_lattice set \<Rightarrow> 'a"  ("\<Squnion>_" [90] 90) where
8fb53badad99 more symbols;
wenzelm
parents: 58879
diff changeset
    38
  "\<Squnion>A = (THE sup. is_Sup A sup)"
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    39
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
    40
text \<open>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    41
  Due to unique existence of bounds, the complete lattice operations
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    42
  may be exhibited as follows.
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
    43
\<close>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    44
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    45
lemma Meet_equality [elim?]: "is_Inf A inf \<Longrightarrow> \<Sqinter>A = inf"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    46
proof (unfold Meet_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    47
  assume "is_Inf A inf"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
    48
  then show "(THE inf. is_Inf A inf) = inf"
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
    49
    by (rule the_equality) (rule is_Inf_uniq [OF _ \<open>is_Inf A inf\<close>])
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    50
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    51
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    52
lemma MeetI [intro?]:
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    53
  "(\<And>a. a \<in> A \<Longrightarrow> inf \<sqsubseteq> a) \<Longrightarrow>
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    54
    (\<And>b. \<forall>a \<in> A. b \<sqsubseteq> a \<Longrightarrow> b \<sqsubseteq> inf) \<Longrightarrow>
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    55
    \<Sqinter>A = inf"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    56
  by (rule Meet_equality, rule is_InfI) blast+
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    57
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    58
lemma Join_equality [elim?]: "is_Sup A sup \<Longrightarrow> \<Squnion>A = sup"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    59
proof (unfold Join_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    60
  assume "is_Sup A sup"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
    61
  then show "(THE sup. is_Sup A sup) = sup"
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
    62
    by (rule the_equality) (rule is_Sup_uniq [OF _ \<open>is_Sup A sup\<close>])
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    63
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    64
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    65
lemma JoinI [intro?]:
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    66
  "(\<And>a. a \<in> A \<Longrightarrow> a \<sqsubseteq> sup) \<Longrightarrow>
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    67
    (\<And>b. \<forall>a \<in> A. a \<sqsubseteq> b \<Longrightarrow> sup \<sqsubseteq> b) \<Longrightarrow>
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    68
    \<Squnion>A = sup"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    69
  by (rule Join_equality, rule is_SupI) blast+
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    70
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    71
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
    72
text \<open>
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
    73
  \medskip The \<open>\<Sqinter>\<close> and \<open>\<Squnion>\<close> operations indeed determine
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    74
  bounds on a complete lattice structure.
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
    75
\<close>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    76
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    77
lemma is_Inf_Meet [intro?]: "is_Inf A (\<Sqinter>A)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    78
proof (unfold Meet_def)
11441
54b236801671 replaced SOME by THE;
wenzelm
parents: 11099
diff changeset
    79
  from ex_Inf obtain inf where "is_Inf A inf" ..
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
    80
  then show "is_Inf A (THE inf. is_Inf A inf)"
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
    81
    by (rule theI) (rule is_Inf_uniq [OF _ \<open>is_Inf A inf\<close>])
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    82
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    83
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    84
lemma Meet_greatest [intro?]: "(\<And>a. a \<in> A \<Longrightarrow> x \<sqsubseteq> a) \<Longrightarrow> x \<sqsubseteq> \<Sqinter>A"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    85
  by (rule is_Inf_greatest, rule is_Inf_Meet) blast
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    86
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    87
lemma Meet_lower [intro?]: "a \<in> A \<Longrightarrow> \<Sqinter>A \<sqsubseteq> a"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    88
  by (rule is_Inf_lower) (rule is_Inf_Meet)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    89
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    90
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    91
lemma is_Sup_Join [intro?]: "is_Sup A (\<Squnion>A)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    92
proof (unfold Join_def)
11441
54b236801671 replaced SOME by THE;
wenzelm
parents: 11099
diff changeset
    93
  from ex_Sup obtain sup where "is_Sup A sup" ..
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
    94
  then show "is_Sup A (THE sup. is_Sup A sup)"
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
    95
    by (rule theI) (rule is_Sup_uniq [OF _ \<open>is_Sup A sup\<close>])
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    96
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    97
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    98
lemma Join_least [intro?]: "(\<And>a. a \<in> A \<Longrightarrow> a \<sqsubseteq> x) \<Longrightarrow> \<Squnion>A \<sqsubseteq> x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    99
  by (rule is_Sup_least, rule is_Sup_Join) blast
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   100
lemma Join_lower [intro?]: "a \<in> A \<Longrightarrow> a \<sqsubseteq> \<Squnion>A"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   101
  by (rule is_Sup_upper) (rule is_Sup_Join)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   102
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   103
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   104
subsection \<open>The Knaster-Tarski Theorem\<close>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   105
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   106
text \<open>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   107
  The Knaster-Tarski Theorem (in its simplest formulation) states that
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   108
  any monotone function on a complete lattice has a least fixed-point
58622
aa99568f56de more antiquotations;
wenzelm
parents: 56154
diff changeset
   109
  (see @{cite \<open>pages 93--94\<close> "Davey-Priestley:1990"} for example).  This
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   110
  is a consequence of the basic boundary properties of the complete
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   111
  lattice operations.
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   112
\<close>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   113
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   114
theorem Knaster_Tarski:
25469
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   115
  assumes mono: "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y"
25474
c41b433b0f65 Knaster_Tarski: turned into Isar statement, tuned proofs;
wenzelm
parents: 25469
diff changeset
   116
  obtains a :: "'a::complete_lattice" where
c41b433b0f65 Knaster_Tarski: turned into Isar statement, tuned proofs;
wenzelm
parents: 25469
diff changeset
   117
    "f a = a" and "\<And>a'. f a' = a' \<Longrightarrow> a \<sqsubseteq> a'"
c41b433b0f65 Knaster_Tarski: turned into Isar statement, tuned proofs;
wenzelm
parents: 25469
diff changeset
   118
proof
25469
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   119
  let ?H = "{u. f u \<sqsubseteq> u}"
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   120
  let ?a = "\<Sqinter>?H"
25474
c41b433b0f65 Knaster_Tarski: turned into Isar statement, tuned proofs;
wenzelm
parents: 25469
diff changeset
   121
  show "f ?a = ?a"
25469
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   122
  proof -
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   123
    have ge: "f ?a \<sqsubseteq> ?a"
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   124
    proof
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   125
      fix x assume x: "x \<in> ?H"
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   126
      then have "?a \<sqsubseteq> x" ..
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   127
      then have "f ?a \<sqsubseteq> f x" by (rule mono)
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   128
      also from x have "... \<sqsubseteq> x" ..
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   129
      finally show "f ?a \<sqsubseteq> x" .
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   130
    qed
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   131
    also have "?a \<sqsubseteq> f ?a"
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   132
    proof
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   133
      from ge have "f (f ?a) \<sqsubseteq> f ?a" by (rule mono)
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   134
      then show "f ?a \<in> ?H" ..
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   135
    qed
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   136
    finally show ?thesis .
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   137
  qed
25474
c41b433b0f65 Knaster_Tarski: turned into Isar statement, tuned proofs;
wenzelm
parents: 25469
diff changeset
   138
c41b433b0f65 Knaster_Tarski: turned into Isar statement, tuned proofs;
wenzelm
parents: 25469
diff changeset
   139
  fix a'
c41b433b0f65 Knaster_Tarski: turned into Isar statement, tuned proofs;
wenzelm
parents: 25469
diff changeset
   140
  assume "f a' = a'"
c41b433b0f65 Knaster_Tarski: turned into Isar statement, tuned proofs;
wenzelm
parents: 25469
diff changeset
   141
  then have "f a' \<sqsubseteq> a'" by (simp only: leq_refl)
c41b433b0f65 Knaster_Tarski: turned into Isar statement, tuned proofs;
wenzelm
parents: 25469
diff changeset
   142
  then have "a' \<in> ?H" ..
c41b433b0f65 Knaster_Tarski: turned into Isar statement, tuned proofs;
wenzelm
parents: 25469
diff changeset
   143
  then show "?a \<sqsubseteq> a'" ..
25469
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   144
qed
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   145
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   146
theorem Knaster_Tarski_dual:
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   147
  assumes mono: "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y"
25474
c41b433b0f65 Knaster_Tarski: turned into Isar statement, tuned proofs;
wenzelm
parents: 25469
diff changeset
   148
  obtains a :: "'a::complete_lattice" where
c41b433b0f65 Knaster_Tarski: turned into Isar statement, tuned proofs;
wenzelm
parents: 25469
diff changeset
   149
    "f a = a" and "\<And>a'. f a' = a' \<Longrightarrow> a' \<sqsubseteq> a"
c41b433b0f65 Knaster_Tarski: turned into Isar statement, tuned proofs;
wenzelm
parents: 25469
diff changeset
   150
proof
25469
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   151
  let ?H = "{u. u \<sqsubseteq> f u}"
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   152
  let ?a = "\<Squnion>?H"
25474
c41b433b0f65 Knaster_Tarski: turned into Isar statement, tuned proofs;
wenzelm
parents: 25469
diff changeset
   153
  show "f ?a = ?a"
25469
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   154
  proof -
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   155
    have le: "?a \<sqsubseteq> f ?a"
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   156
    proof
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   157
      fix x assume x: "x \<in> ?H"
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   158
      then have "x \<sqsubseteq> f x" ..
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   159
      also from x have "x \<sqsubseteq> ?a" ..
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   160
      then have "f x \<sqsubseteq> f ?a" by (rule mono)
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   161
      finally show "x \<sqsubseteq> f ?a" .
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   162
    qed
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   163
    have "f ?a \<sqsubseteq> ?a"
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   164
    proof
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   165
      from le have "f ?a \<sqsubseteq> f (f ?a)" by (rule mono)
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   166
      then show "f ?a \<in> ?H" ..
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   167
    qed
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   168
    from this and le show ?thesis by (rule leq_antisym)
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   169
  qed
25474
c41b433b0f65 Knaster_Tarski: turned into Isar statement, tuned proofs;
wenzelm
parents: 25469
diff changeset
   170
c41b433b0f65 Knaster_Tarski: turned into Isar statement, tuned proofs;
wenzelm
parents: 25469
diff changeset
   171
  fix a'
c41b433b0f65 Knaster_Tarski: turned into Isar statement, tuned proofs;
wenzelm
parents: 25469
diff changeset
   172
  assume "f a' = a'"
c41b433b0f65 Knaster_Tarski: turned into Isar statement, tuned proofs;
wenzelm
parents: 25469
diff changeset
   173
  then have "a' \<sqsubseteq> f a'" by (simp only: leq_refl)
c41b433b0f65 Knaster_Tarski: turned into Isar statement, tuned proofs;
wenzelm
parents: 25469
diff changeset
   174
  then have "a' \<in> ?H" ..
c41b433b0f65 Knaster_Tarski: turned into Isar statement, tuned proofs;
wenzelm
parents: 25469
diff changeset
   175
  then show "a' \<sqsubseteq> ?a" ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   176
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   177
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   178
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   179
subsection \<open>Bottom and top elements\<close>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   180
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   181
text \<open>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   182
  With general bounds available, complete lattices also have least and
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   183
  greatest elements.
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   184
\<close>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   185
19736
wenzelm
parents: 16417
diff changeset
   186
definition
25469
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   187
  bottom :: "'a::complete_lattice"  ("\<bottom>") where
19736
wenzelm
parents: 16417
diff changeset
   188
  "\<bottom> = \<Sqinter>UNIV"
25469
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   189
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
   190
definition
25469
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23373
diff changeset
   191
  top :: "'a::complete_lattice"  ("\<top>") where
19736
wenzelm
parents: 16417
diff changeset
   192
  "\<top> = \<Squnion>UNIV"
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   193
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   194
lemma bottom_least [intro?]: "\<bottom> \<sqsubseteq> x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   195
proof (unfold bottom_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   196
  have "x \<in> UNIV" ..
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   197
  then show "\<Sqinter>UNIV \<sqsubseteq> x" ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   198
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   199
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   200
lemma bottomI [intro?]: "(\<And>a. x \<sqsubseteq> a) \<Longrightarrow> \<bottom> = x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   201
proof (unfold bottom_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   202
  assume "\<And>a. x \<sqsubseteq> a"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   203
  show "\<Sqinter>UNIV = x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   204
  proof
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   205
    fix a show "x \<sqsubseteq> a" by fact
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   206
  next
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   207
    fix b :: "'a::complete_lattice"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   208
    assume b: "\<forall>a \<in> UNIV. b \<sqsubseteq> a"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   209
    have "x \<in> UNIV" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   210
    with b show "b \<sqsubseteq> x" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   211
  qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   212
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   213
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   214
lemma top_greatest [intro?]: "x \<sqsubseteq> \<top>"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   215
proof (unfold top_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   216
  have "x \<in> UNIV" ..
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   217
  then show "x \<sqsubseteq> \<Squnion>UNIV" ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   218
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   219
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   220
lemma topI [intro?]: "(\<And>a. a \<sqsubseteq> x) \<Longrightarrow> \<top> = x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   221
proof (unfold top_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   222
  assume "\<And>a. a \<sqsubseteq> x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   223
  show "\<Squnion>UNIV = x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   224
  proof
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   225
    fix a show "a \<sqsubseteq> x" by fact
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   226
  next
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   227
    fix b :: "'a::complete_lattice"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   228
    assume b: "\<forall>a \<in> UNIV. a \<sqsubseteq> b"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   229
    have "x \<in> UNIV" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   230
    with b show "x \<sqsubseteq> b" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   231
  qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   232
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   233
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   234
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   235
subsection \<open>Duality\<close>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   236
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   237
text \<open>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   238
  The class of complete lattices is closed under formation of dual
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   239
  structures.
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   240
\<close>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   241
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   242
instance dual :: (complete_lattice) complete_lattice
10309
a7f961fb62c6 intro_classes by default;
wenzelm
parents: 10183
diff changeset
   243
proof
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   244
  fix A' :: "'a::complete_lattice dual set"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   245
  show "\<exists>inf'. is_Inf A' inf'"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   246
  proof -
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10309
diff changeset
   247
    have "\<exists>sup. is_Sup (undual ` A') sup" by (rule ex_Sup)
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   248
    then have "\<exists>sup. is_Inf (dual ` undual ` A') (dual sup)" by (simp only: dual_Inf)
56154
f0a927235162 more complete set of lemmas wrt. image and composition
haftmann
parents: 35317
diff changeset
   249
    then show ?thesis by (simp add: dual_ex [symmetric] image_comp)
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   250
  qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   251
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   252
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   253
text \<open>
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   254
  Apparently, the \<open>\<Sqinter>\<close> and \<open>\<Squnion>\<close> operations are dual to each
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   255
  other.
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   256
\<close>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   257
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10309
diff changeset
   258
theorem dual_Meet [intro?]: "dual (\<Sqinter>A) = \<Squnion>(dual ` A)"
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   259
proof -
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10309
diff changeset
   260
  from is_Inf_Meet have "is_Sup (dual ` A) (dual (\<Sqinter>A))" ..
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   261
  then have "\<Squnion>(dual ` A) = dual (\<Sqinter>A)" ..
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   262
  then show ?thesis ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   263
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   264
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10309
diff changeset
   265
theorem dual_Join [intro?]: "dual (\<Squnion>A) = \<Sqinter>(dual ` A)"
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   266
proof -
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10309
diff changeset
   267
  from is_Sup_Join have "is_Inf (dual ` A) (dual (\<Squnion>A))" ..
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   268
  then have "\<Sqinter>(dual ` A) = dual (\<Squnion>A)" ..
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   269
  then show ?thesis ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   270
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   271
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   272
text \<open>
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   273
  Likewise are \<open>\<bottom>\<close> and \<open>\<top>\<close> duals of each other.
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   274
\<close>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   275
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   276
theorem dual_bottom [intro?]: "dual \<bottom> = \<top>"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   277
proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   278
  have "\<top> = dual \<bottom>"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   279
  proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   280
    fix a' have "\<bottom> \<sqsubseteq> undual a'" ..
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   281
    then have "dual (undual a') \<sqsubseteq> dual \<bottom>" ..
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   282
    then show "a' \<sqsubseteq> dual \<bottom>" by simp
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   283
  qed
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   284
  then show ?thesis ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   285
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   286
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   287
theorem dual_top [intro?]: "dual \<top> = \<bottom>"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   288
proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   289
  have "\<bottom> = dual \<top>"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   290
  proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   291
    fix a' have "undual a' \<sqsubseteq> \<top>" ..
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   292
    then have "dual \<top> \<sqsubseteq> dual (undual a')" ..
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   293
    then show "dual \<top> \<sqsubseteq> a'" by simp
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   294
  qed
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   295
  then show ?thesis ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   296
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   297
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   298
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   299
subsection \<open>Complete lattices are lattices\<close>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   300
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   301
text \<open>
10176
2e38e3c94990 tuned text;
wenzelm
parents: 10175
diff changeset
   302
  Complete lattices (with general bounds available) are indeed plain
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   303
  lattices as well.  This holds due to the connection of general
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   304
  versus binary bounds that has been formally established in
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   305
  \S\ref{sec:gen-bin-bounds}.
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   306
\<close>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   307
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   308
lemma is_inf_binary: "is_inf x y (\<Sqinter>{x, y})"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   309
proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   310
  have "is_Inf {x, y} (\<Sqinter>{x, y})" ..
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   311
  then show ?thesis by (simp only: is_Inf_binary)
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   312
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   313
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   314
lemma is_sup_binary: "is_sup x y (\<Squnion>{x, y})"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   315
proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   316
  have "is_Sup {x, y} (\<Squnion>{x, y})" ..
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   317
  then show ?thesis by (simp only: is_Sup_binary)
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   318
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   319
11099
b301d1f72552 \<subseteq>;
wenzelm
parents: 10834
diff changeset
   320
instance complete_lattice \<subseteq> lattice
10309
a7f961fb62c6 intro_classes by default;
wenzelm
parents: 10183
diff changeset
   321
proof
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   322
  fix x y :: "'a::complete_lattice"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   323
  from is_inf_binary show "\<exists>inf. is_inf x y inf" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   324
  from is_sup_binary show "\<exists>sup. is_sup x y sup" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   325
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   326
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   327
theorem meet_binary: "x \<sqinter> y = \<Sqinter>{x, y}"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   328
  by (rule meet_equality) (rule is_inf_binary)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   329
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   330
theorem join_binary: "x \<squnion> y = \<Squnion>{x, y}"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   331
  by (rule join_equality) (rule is_sup_binary)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   332
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   333
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   334
subsection \<open>Complete lattices and set-theory operations\<close>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   335
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   336
text \<open>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   337
  The complete lattice operations are (anti) monotone wrt.\ set
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   338
  inclusion.
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   339
\<close>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   340
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   341
theorem Meet_subset_antimono: "A \<subseteq> B \<Longrightarrow> \<Sqinter>B \<sqsubseteq> \<Sqinter>A"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   342
proof (rule Meet_greatest)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   343
  fix a assume "a \<in> A"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   344
  also assume "A \<subseteq> B"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   345
  finally have "a \<in> B" .
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   346
  then show "\<Sqinter>B \<sqsubseteq> a" ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   347
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   348
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   349
theorem Join_subset_mono: "A \<subseteq> B \<Longrightarrow> \<Squnion>A \<sqsubseteq> \<Squnion>B"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   350
proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   351
  assume "A \<subseteq> B"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   352
  then have "dual ` A \<subseteq> dual ` B" by blast
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   353
  then have "\<Sqinter>(dual ` B) \<sqsubseteq> \<Sqinter>(dual ` A)" by (rule Meet_subset_antimono)
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   354
  then have "dual (\<Squnion>B) \<sqsubseteq> dual (\<Squnion>A)" by (simp only: dual_Join)
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   355
  then show ?thesis by (simp only: dual_leq)
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   356
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   357
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   358
text \<open>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   359
  Bounds over unions of sets may be obtained separately.
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   360
\<close>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   361
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   362
theorem Meet_Un: "\<Sqinter>(A \<union> B) = \<Sqinter>A \<sqinter> \<Sqinter>B"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   363
proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   364
  fix a assume "a \<in> A \<union> B"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   365
  then show "\<Sqinter>A \<sqinter> \<Sqinter>B \<sqsubseteq> a"
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   366
  proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   367
    assume a: "a \<in> A"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   368
    have "\<Sqinter>A \<sqinter> \<Sqinter>B \<sqsubseteq> \<Sqinter>A" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   369
    also from a have "\<dots> \<sqsubseteq> a" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   370
    finally show ?thesis .
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   371
  next
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   372
    assume a: "a \<in> B"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   373
    have "\<Sqinter>A \<sqinter> \<Sqinter>B \<sqsubseteq> \<Sqinter>B" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   374
    also from a have "\<dots> \<sqsubseteq> a" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   375
    finally show ?thesis .
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   376
  qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   377
next
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   378
  fix b assume b: "\<forall>a \<in> A \<union> B. b \<sqsubseteq> a"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   379
  show "b \<sqsubseteq> \<Sqinter>A \<sqinter> \<Sqinter>B"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   380
  proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   381
    show "b \<sqsubseteq> \<Sqinter>A"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   382
    proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   383
      fix a assume "a \<in> A"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   384
      then have "a \<in>  A \<union> B" ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   385
      with b show "b \<sqsubseteq> a" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   386
    qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   387
    show "b \<sqsubseteq> \<Sqinter>B"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   388
    proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   389
      fix a assume "a \<in> B"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   390
      then have "a \<in>  A \<union> B" ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   391
      with b show "b \<sqsubseteq> a" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   392
    qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   393
  qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   394
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   395
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   396
theorem Join_Un: "\<Squnion>(A \<union> B) = \<Squnion>A \<squnion> \<Squnion>B"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   397
proof -
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10309
diff changeset
   398
  have "dual (\<Squnion>(A \<union> B)) = \<Sqinter>(dual ` A \<union> dual ` B)"
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   399
    by (simp only: dual_Join image_Un)
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10309
diff changeset
   400
  also have "\<dots> = \<Sqinter>(dual ` A) \<sqinter> \<Sqinter>(dual ` B)"
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   401
    by (rule Meet_Un)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   402
  also have "\<dots> = dual (\<Squnion>A \<squnion> \<Squnion>B)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   403
    by (simp only: dual_join dual_Join)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   404
  finally show ?thesis ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   405
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   406
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   407
text \<open>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   408
  Bounds over singleton sets are trivial.
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   409
\<close>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   410
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   411
theorem Meet_singleton: "\<Sqinter>{x} = x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   412
proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   413
  fix a assume "a \<in> {x}"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   414
  then have "a = x" by simp
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   415
  then show "x \<sqsubseteq> a" by (simp only: leq_refl)
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   416
next
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   417
  fix b assume "\<forall>a \<in> {x}. b \<sqsubseteq> a"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   418
  then show "b \<sqsubseteq> x" by simp
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   419
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   420
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   421
theorem Join_singleton: "\<Squnion>{x} = x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   422
proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   423
  have "dual (\<Squnion>{x}) = \<Sqinter>{dual x}" by (simp add: dual_Join)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   424
  also have "\<dots> = dual x" by (rule Meet_singleton)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   425
  finally show ?thesis ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   426
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   427
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   428
text \<open>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   429
  Bounds over the empty and universal set correspond to each other.
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 61983
diff changeset
   430
\<close>
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   431
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   432
theorem Meet_empty: "\<Sqinter>{} = \<Squnion>UNIV"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   433
proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   434
  fix a :: "'a::complete_lattice"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   435
  assume "a \<in> {}"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   436
  then have False by simp
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   437
  then show "\<Squnion>UNIV \<sqsubseteq> a" ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   438
next
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   439
  fix b :: "'a::complete_lattice"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   440
  have "b \<in> UNIV" ..
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   441
  then show "b \<sqsubseteq> \<Squnion>UNIV" ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   442
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   443
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   444
theorem Join_empty: "\<Squnion>{} = \<Sqinter>UNIV"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   445
proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   446
  have "dual (\<Squnion>{}) = \<Sqinter>{}" by (simp add: dual_Join)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   447
  also have "\<dots> = \<Squnion>UNIV" by (rule Meet_empty)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   448
  also have "\<dots> = dual (\<Sqinter>UNIV)" by (simp add: dual_Meet)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   449
  finally show ?thesis ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   450
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   451
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   452
end