src/HOL/Lattice/Lattice.thy
author bulwahn
Fri, 11 Mar 2011 15:21:13 +0100
changeset 41919 e180c2a9873b
parent 37678 0040bafffdef
child 58879 143c85e3cdb5
permissions -rw-r--r--
correcting dependencies after renaming
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/Lattice.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
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
     5
header {* Lattices *}
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: 12818
diff changeset
     7
theory Lattice imports Bounds begin
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
     8
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
     9
subsection {* Lattice operations *}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    10
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    11
text {*
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    12
  A \emph{lattice} is a partial order with infimum and supremum of any
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    13
  two elements (thus any \emph{finite} number of elements have bounds
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    14
  as well).
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    15
*}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    16
35317
d57da4abb47d dropped axclass; dropped Id
haftmann
parents: 25469
diff changeset
    17
class lattice =
d57da4abb47d dropped axclass; dropped Id
haftmann
parents: 25469
diff changeset
    18
  assumes ex_inf: "\<exists>inf. is_inf x y inf"
d57da4abb47d dropped axclass; dropped Id
haftmann
parents: 25469
diff changeset
    19
  assumes ex_sup: "\<exists>sup. is_sup x y sup"
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
text {*
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    22
  The @{text \<sqinter>} (meet) and @{text \<squnion>} (join) operations select such
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    23
  infimum and supremum elements.
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    24
*}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    25
19736
wenzelm
parents: 16417
diff changeset
    26
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    27
  meet :: "'a::lattice \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl "&&" 70) where
19736
wenzelm
parents: 16417
diff changeset
    28
  "x && y = (THE inf. is_inf x y inf)"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    29
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    30
  join :: "'a::lattice \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl "||" 65) where
19736
wenzelm
parents: 16417
diff changeset
    31
  "x || y = (THE sup. is_sup x y sup)"
wenzelm
parents: 16417
diff changeset
    32
21210
c17fd2df4e9e renamed 'const_syntax' to 'notation';
wenzelm
parents: 20523
diff changeset
    33
notation (xsymbols)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    34
  meet  (infixl "\<sqinter>" 70) and
19736
wenzelm
parents: 16417
diff changeset
    35
  join  (infixl "\<squnion>" 65)
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    36
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    37
text {*
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    38
  Due to unique existence of bounds, the lattice operations may be
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    39
  exhibited as follows.
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    40
*}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    41
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    42
lemma meet_equality [elim?]: "is_inf x y inf \<Longrightarrow> x \<sqinter> y = inf"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    43
proof (unfold meet_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    44
  assume "is_inf x y inf"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
    45
  then show "(THE inf. is_inf x y inf) = inf"
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
    46
    by (rule the_equality) (rule is_inf_uniq [OF _ `is_inf x y inf`])
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    47
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    48
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    49
lemma meetI [intro?]:
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    50
    "inf \<sqsubseteq> x \<Longrightarrow> inf \<sqsubseteq> y \<Longrightarrow> (\<And>z. z \<sqsubseteq> x \<Longrightarrow> z \<sqsubseteq> y \<Longrightarrow> z \<sqsubseteq> inf) \<Longrightarrow> x \<sqinter> y = inf"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    51
  by (rule meet_equality, rule is_infI) blast+
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    52
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    53
lemma join_equality [elim?]: "is_sup x y sup \<Longrightarrow> x \<squnion> y = sup"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    54
proof (unfold join_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    55
  assume "is_sup x y sup"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
    56
  then show "(THE sup. is_sup x y sup) = sup"
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
    57
    by (rule the_equality) (rule is_sup_uniq [OF _ `is_sup x y sup`])
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    58
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    59
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    60
lemma joinI [intro?]: "x \<sqsubseteq> sup \<Longrightarrow> y \<sqsubseteq> sup \<Longrightarrow>
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    61
    (\<And>z. x \<sqsubseteq> z \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> sup \<sqsubseteq> z) \<Longrightarrow> x \<squnion> y = sup"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    62
  by (rule join_equality, rule is_supI) blast+
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    63
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
text {*
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    66
  \medskip The @{text \<sqinter>} and @{text \<squnion>} operations indeed determine
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    67
  bounds on a lattice structure.
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    68
*}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    69
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    70
lemma is_inf_meet [intro?]: "is_inf x y (x \<sqinter> y)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    71
proof (unfold meet_def)
11441
54b236801671 replaced SOME by THE;
wenzelm
parents: 11099
diff changeset
    72
  from ex_inf obtain inf where "is_inf x y inf" ..
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
    73
  then show "is_inf x y (THE inf. is_inf x y inf)"
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
    74
    by (rule theI) (rule is_inf_uniq [OF _ `is_inf x y inf`])
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    75
qed
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 meet_greatest [intro?]: "z \<sqsubseteq> x \<Longrightarrow> z \<sqsubseteq> y \<Longrightarrow> z \<sqsubseteq> x \<sqinter> y"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    78
  by (rule is_inf_greatest) (rule is_inf_meet)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    79
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    80
lemma meet_lower1 [intro?]: "x \<sqinter> y \<sqsubseteq> x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    81
  by (rule is_inf_lower) (rule is_inf_meet)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    82
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    83
lemma meet_lower2 [intro?]: "x \<sqinter> y \<sqsubseteq> y"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    84
  by (rule is_inf_lower) (rule is_inf_meet)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    85
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 is_sup_join [intro?]: "is_sup x y (x \<squnion> y)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    88
proof (unfold join_def)
11441
54b236801671 replaced SOME by THE;
wenzelm
parents: 11099
diff changeset
    89
  from ex_sup obtain sup where "is_sup x y sup" ..
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
    90
  then show "is_sup x y (THE sup. is_sup x y sup)"
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
    91
    by (rule theI) (rule is_sup_uniq [OF _ `is_sup x y sup`])
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    92
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    93
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    94
lemma join_least [intro?]: "x \<sqsubseteq> z \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<squnion> y \<sqsubseteq> z"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    95
  by (rule is_sup_least) (rule is_sup_join)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    96
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    97
lemma join_upper1 [intro?]: "x \<sqsubseteq> x \<squnion> y"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    98
  by (rule is_sup_upper) (rule is_sup_join)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
    99
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   100
lemma join_upper2 [intro?]: "y \<sqsubseteq> x \<squnion> y"
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
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   104
subsection {* Duality *}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   105
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   106
text {*
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   107
  The class of lattices is closed under formation of dual structures.
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   108
  This means that for any theorem of lattice theory, the dualized
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   109
  statement holds as well; this important fact simplifies many proofs
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   110
  of lattice theory.
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   111
*}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   112
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   113
instance dual :: (lattice) lattice
10309
a7f961fb62c6 intro_classes by default;
wenzelm
parents: 10183
diff changeset
   114
proof
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   115
  fix x' y' :: "'a::lattice dual"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   116
  show "\<exists>inf'. is_inf x' y' inf'"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   117
  proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   118
    have "\<exists>sup. is_sup (undual x') (undual y') sup" by (rule ex_sup)
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   119
    then have "\<exists>sup. is_inf (dual (undual x')) (dual (undual y')) (dual sup)"
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   120
      by (simp only: dual_inf)
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   121
    then show ?thesis by (simp add: dual_ex [symmetric])
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   122
  qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   123
  show "\<exists>sup'. is_sup x' y' sup'"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   124
  proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   125
    have "\<exists>inf. is_inf (undual x') (undual y') inf" by (rule ex_inf)
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   126
    then have "\<exists>inf. is_sup (dual (undual x')) (dual (undual y')) (dual inf)"
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   127
      by (simp only: dual_sup)
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   128
    then show ?thesis by (simp add: dual_ex [symmetric])
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   129
  qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   130
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   131
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   132
text {*
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   133
  Apparently, the @{text \<sqinter>} and @{text \<squnion>} operations are dual to each
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   134
  other.
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   135
*}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   136
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   137
theorem dual_meet [intro?]: "dual (x \<sqinter> y) = dual x \<squnion> dual y"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   138
proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   139
  from is_inf_meet have "is_sup (dual x) (dual y) (dual (x \<sqinter> y))" ..
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   140
  then have "dual x \<squnion> dual y = dual (x \<sqinter> y)" ..
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   141
  then show ?thesis ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   142
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   143
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   144
theorem dual_join [intro?]: "dual (x \<squnion> y) = dual x \<sqinter> dual y"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   145
proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   146
  from is_sup_join have "is_inf (dual x) (dual y) (dual (x \<squnion> y))" ..
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   147
  then have "dual x \<sqinter> dual y = dual (x \<squnion> y)" ..
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   148
  then show ?thesis ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   149
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   150
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   151
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   152
subsection {* Algebraic properties \label{sec:lattice-algebra} *}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   153
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   154
text {*
12818
e7b4c0731d57 fixed typos
kleing
parents: 12338
diff changeset
   155
  The @{text \<sqinter>} and @{text \<squnion>} operations have the following
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   156
  characteristic algebraic properties: associative (A), commutative
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   157
  (C), and absorptive (AB).
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   158
*}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   159
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   160
theorem meet_assoc: "(x \<sqinter> y) \<sqinter> z = x \<sqinter> (y \<sqinter> z)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   161
proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   162
  show "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> x \<sqinter> y"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   163
  proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   164
    show "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> x" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   165
    show "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> y"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   166
    proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   167
      have "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> y \<sqinter> z" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   168
      also have "\<dots> \<sqsubseteq> y" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   169
      finally show ?thesis .
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   170
    qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   171
  qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   172
  show "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> z"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   173
  proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   174
    have "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> y \<sqinter> z" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   175
    also have "\<dots> \<sqsubseteq> z" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   176
    finally show ?thesis .
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   177
  qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   178
  fix w assume "w \<sqsubseteq> x \<sqinter> y" and "w \<sqsubseteq> z"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   179
  show "w \<sqsubseteq> x \<sqinter> (y \<sqinter> z)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   180
  proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   181
    show "w \<sqsubseteq> x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   182
    proof -
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   183
      have "w \<sqsubseteq> x \<sqinter> y" by fact
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   184
      also have "\<dots> \<sqsubseteq> x" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   185
      finally show ?thesis .
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   186
    qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   187
    show "w \<sqsubseteq> y \<sqinter> z"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   188
    proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   189
      show "w \<sqsubseteq> y"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   190
      proof -
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   191
        have "w \<sqsubseteq> x \<sqinter> y" by fact
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   192
        also have "\<dots> \<sqsubseteq> y" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   193
        finally show ?thesis .
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   194
      qed
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   195
      show "w \<sqsubseteq> z" by fact
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   196
    qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   197
  qed
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
theorem join_assoc: "(x \<squnion> y) \<squnion> z = x \<squnion> (y \<squnion> z)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   201
proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   202
  have "dual ((x \<squnion> y) \<squnion> z) = (dual x \<sqinter> dual y) \<sqinter> dual z"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   203
    by (simp only: dual_join)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   204
  also have "\<dots> = dual x \<sqinter> (dual y \<sqinter> dual z)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   205
    by (rule meet_assoc)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   206
  also have "\<dots> = dual (x \<squnion> (y \<squnion> z))"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   207
    by (simp only: dual_join)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   208
  finally show ?thesis ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   209
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   210
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   211
theorem meet_commute: "x \<sqinter> y = y \<sqinter> x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   212
proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   213
  show "y \<sqinter> x \<sqsubseteq> x" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   214
  show "y \<sqinter> x \<sqsubseteq> y" ..
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   215
  fix z assume "z \<sqsubseteq> y" and "z \<sqsubseteq> x"
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   216
  then show "z \<sqsubseteq> y \<sqinter> x" ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   217
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   218
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   219
theorem join_commute: "x \<squnion> y = y \<squnion> x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   220
proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   221
  have "dual (x \<squnion> y) = dual x \<sqinter> dual y" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   222
  also have "\<dots> = dual y \<sqinter> dual x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   223
    by (rule meet_commute)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   224
  also have "\<dots> = dual (y \<squnion> x)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   225
    by (simp only: dual_join)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   226
  finally show ?thesis ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   227
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   228
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   229
theorem meet_join_absorb: "x \<sqinter> (x \<squnion> y) = x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   230
proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   231
  show "x \<sqsubseteq> x" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   232
  show "x \<sqsubseteq> x \<squnion> y" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   233
  fix z assume "z \<sqsubseteq> x" and "z \<sqsubseteq> x \<squnion> y"
23393
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23373
diff changeset
   234
  show "z \<sqsubseteq> x" by fact
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   235
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   236
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   237
theorem join_meet_absorb: "x \<squnion> (x \<sqinter> y) = x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   238
proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   239
  have "dual x \<sqinter> (dual x \<squnion> dual y) = dual x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   240
    by (rule meet_join_absorb)
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   241
  then have "dual (x \<squnion> (x \<sqinter> y)) = dual x"
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   242
    by (simp only: dual_meet dual_join)
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   243
  then show ?thesis ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   244
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   245
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   246
text {*
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   247
  \medskip Some further algebraic properties hold as well.  The
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   248
  property idempotent (I) is a basic algebraic consequence of (AB).
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   249
*}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   250
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   251
theorem meet_idem: "x \<sqinter> x = x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   252
proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   253
  have "x \<sqinter> (x \<squnion> (x \<sqinter> x)) = x" by (rule meet_join_absorb)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   254
  also have "x \<squnion> (x \<sqinter> x) = x" by (rule join_meet_absorb)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   255
  finally show ?thesis .
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   256
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   257
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   258
theorem join_idem: "x \<squnion> x = x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   259
proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   260
  have "dual x \<sqinter> dual x = dual x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   261
    by (rule meet_idem)
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   262
  then have "dual (x \<squnion> x) = dual x"
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   263
    by (simp only: dual_join)
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   264
  then show ?thesis ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   265
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   266
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   267
text {*
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   268
  Meet and join are trivial for related elements.
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   269
*}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   270
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   271
theorem meet_related [elim?]: "x \<sqsubseteq> y \<Longrightarrow> x \<sqinter> y = x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   272
proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   273
  assume "x \<sqsubseteq> y"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   274
  show "x \<sqsubseteq> x" ..
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   275
  show "x \<sqsubseteq> y" by fact
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   276
  fix z assume "z \<sqsubseteq> x" and "z \<sqsubseteq> y"
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   277
  show "z \<sqsubseteq> x" by fact
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   278
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   279
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   280
theorem join_related [elim?]: "x \<sqsubseteq> y \<Longrightarrow> x \<squnion> y = y"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   281
proof -
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   282
  assume "x \<sqsubseteq> y" then have "dual y \<sqsubseteq> dual x" ..
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   283
  then have "dual y \<sqinter> dual x = dual y" by (rule meet_related)
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   284
  also have "dual y \<sqinter> dual x = dual (y \<squnion> x)" by (simp only: dual_join)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   285
  also have "y \<squnion> x = x \<squnion> y" by (rule join_commute)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   286
  finally show ?thesis ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   287
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   288
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   289
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   290
subsection {* Order versus algebraic structure *}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   291
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   292
text {*
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   293
  The @{text \<sqinter>} and @{text \<squnion>} operations are connected with the
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   294
  underlying @{text \<sqsubseteq>} relation in a canonical manner.
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   295
*}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   296
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   297
theorem meet_connection: "(x \<sqsubseteq> y) = (x \<sqinter> y = x)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   298
proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   299
  assume "x \<sqsubseteq> y"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   300
  then have "is_inf x y x" ..
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   301
  then show "x \<sqinter> y = x" ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   302
next
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   303
  have "x \<sqinter> y \<sqsubseteq> y" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   304
  also assume "x \<sqinter> y = x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   305
  finally show "x \<sqsubseteq> y" .
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   306
qed
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
theorem join_connection: "(x \<sqsubseteq> y) = (x \<squnion> y = 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
  assume "x \<sqsubseteq> y"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   311
  then have "is_sup x y y" ..
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   312
  then show "x \<squnion> y = y" ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   313
next
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   314
  have "x \<sqsubseteq> x \<squnion> y" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   315
  also assume "x \<squnion> y = y"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   316
  finally show "x \<sqsubseteq> y" .
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   317
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   318
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   319
text {*
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   320
  \medskip The most fundamental result of the meta-theory of lattices
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   321
  is as follows (we do not prove it here).
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   322
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   323
  Given a structure with binary operations @{text \<sqinter>} and @{text \<squnion>}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   324
  such that (A), (C), and (AB) hold (cf.\
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   325
  \S\ref{sec:lattice-algebra}).  This structure represents a lattice,
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   326
  if the relation @{term "x \<sqsubseteq> y"} is defined as @{term "x \<sqinter> y = x"}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   327
  (alternatively as @{term "x \<squnion> y = y"}).  Furthermore, infimum and
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   328
  supremum with respect to this ordering coincide with the original
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   329
  @{text \<sqinter>} and @{text \<squnion>} operations.
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   330
*}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   331
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
subsection {* Example instances *}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   334
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   335
subsubsection {* Linear orders *}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   336
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   337
text {*
12818
e7b4c0731d57 fixed typos
kleing
parents: 12338
diff changeset
   338
  Linear orders with @{term minimum} and @{term maximum} operations
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   339
  are a (degenerate) example of lattice structures.
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
19736
wenzelm
parents: 16417
diff changeset
   342
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
   343
  minimum :: "'a::linear_order \<Rightarrow> 'a \<Rightarrow> 'a" where
19736
wenzelm
parents: 16417
diff changeset
   344
  "minimum x y = (if x \<sqsubseteq> y then x else y)"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
   345
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
   346
  maximum :: "'a::linear_order \<Rightarrow> 'a \<Rightarrow> 'a" where
19736
wenzelm
parents: 16417
diff changeset
   347
  "maximum x y = (if x \<sqsubseteq> y then y else x)"
10157
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
lemma is_inf_minimum: "is_inf x y (minimum x y)"
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
  let ?min = "minimum x y"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   352
  from leq_linear show "?min \<sqsubseteq> x" by (auto simp add: minimum_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   353
  from leq_linear show "?min \<sqsubseteq> y" by (auto simp add: minimum_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   354
  fix z assume "z \<sqsubseteq> x" and "z \<sqsubseteq> y"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   355
  with leq_linear show "z \<sqsubseteq> ?min" by (auto simp add: minimum_def)
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
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   358
lemma is_sup_maximum: "is_sup x y (maximum x y)"      (* FIXME dualize!? *)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   359
proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   360
  let ?max = "maximum x y"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   361
  from leq_linear show "x \<sqsubseteq> ?max" by (auto simp add: maximum_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   362
  from leq_linear show "y \<sqsubseteq> ?max" by (auto simp add: maximum_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   363
  fix z assume "x \<sqsubseteq> z" and "y \<sqsubseteq> z"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   364
  with leq_linear show "?max \<sqsubseteq> z" by (auto simp add: maximum_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   365
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   366
11099
b301d1f72552 \<subseteq>;
wenzelm
parents: 10309
diff changeset
   367
instance linear_order \<subseteq> lattice
10309
a7f961fb62c6 intro_classes by default;
wenzelm
parents: 10183
diff changeset
   368
proof
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   369
  fix x y :: "'a::linear_order"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   370
  from is_inf_minimum show "\<exists>inf. is_inf x y inf" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   371
  from is_sup_maximum show "\<exists>sup. is_sup x y sup" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   372
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   373
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   374
text {*
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   375
  The lattice operations on linear orders indeed coincide with @{term
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   376
  minimum} and @{term maximum}.
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   377
*}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   378
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   379
theorem meet_mimimum: "x \<sqinter> y = minimum x y"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   380
  by (rule meet_equality) (rule is_inf_minimum)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   381
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   382
theorem meet_maximum: "x \<squnion> y = maximum x y"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   383
  by (rule join_equality) (rule is_sup_maximum)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   384
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   385
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   386
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   387
subsubsection {* Binary products *}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   388
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   389
text {*
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   390
  The class of lattices is closed under direct binary products (cf.\
10158
wenzelm
parents: 10157
diff changeset
   391
  \S\ref{sec:prod-order}).
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   392
*}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   393
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   394
lemma is_inf_prod: "is_inf p q (fst p \<sqinter> fst q, snd p \<sqinter> snd q)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   395
proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   396
  show "(fst p \<sqinter> fst q, snd p \<sqinter> snd q) \<sqsubseteq> p"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   397
  proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   398
    have "fst p \<sqinter> fst q \<sqsubseteq> fst p" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   399
    moreover have "snd p \<sqinter> snd q \<sqsubseteq> snd p" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   400
    ultimately show ?thesis by (simp add: leq_prod_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   401
  qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   402
  show "(fst p \<sqinter> fst q, snd p \<sqinter> snd q) \<sqsubseteq> q"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   403
  proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   404
    have "fst p \<sqinter> fst q \<sqsubseteq> fst q" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   405
    moreover have "snd p \<sqinter> snd q \<sqsubseteq> snd q" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   406
    ultimately show ?thesis by (simp add: leq_prod_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   407
  qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   408
  fix r assume rp: "r \<sqsubseteq> p" and rq: "r \<sqsubseteq> q"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   409
  show "r \<sqsubseteq> (fst p \<sqinter> fst q, snd p \<sqinter> snd q)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   410
  proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   411
    have "fst r \<sqsubseteq> fst p \<sqinter> fst q"
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
      from rp show "fst r \<sqsubseteq> fst p" by (simp add: leq_prod_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   414
      from rq show "fst r \<sqsubseteq> fst q" by (simp add: leq_prod_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   415
    qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   416
    moreover have "snd r \<sqsubseteq> snd p \<sqinter> snd q"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   417
    proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   418
      from rp show "snd r \<sqsubseteq> snd p" by (simp add: leq_prod_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   419
      from rq show "snd r \<sqsubseteq> snd q" by (simp add: leq_prod_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   420
    qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   421
    ultimately show ?thesis by (simp add: leq_prod_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   422
  qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   423
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   424
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   425
lemma is_sup_prod: "is_sup p q (fst p \<squnion> fst q, snd p \<squnion> snd q)"  (* FIXME dualize!? *)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   426
proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   427
  show "p \<sqsubseteq> (fst p \<squnion> fst q, snd p \<squnion> snd q)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   428
  proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   429
    have "fst p \<sqsubseteq> fst p \<squnion> fst q" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   430
    moreover have "snd p \<sqsubseteq> snd p \<squnion> snd q" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   431
    ultimately show ?thesis by (simp add: leq_prod_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   432
  qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   433
  show "q \<sqsubseteq> (fst p \<squnion> fst q, snd p \<squnion> snd q)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   434
  proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   435
    have "fst q \<sqsubseteq> fst p \<squnion> fst q" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   436
    moreover have "snd q \<sqsubseteq> snd p \<squnion> snd q" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   437
    ultimately show ?thesis by (simp add: leq_prod_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   438
  qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   439
  fix r assume "pr": "p \<sqsubseteq> r" and qr: "q \<sqsubseteq> r"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   440
  show "(fst p \<squnion> fst q, snd p \<squnion> snd q) \<sqsubseteq> r"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   441
  proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   442
    have "fst p \<squnion> fst q \<sqsubseteq> fst r"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   443
    proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   444
      from "pr" show "fst p \<sqsubseteq> fst r" by (simp add: leq_prod_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   445
      from qr show "fst q \<sqsubseteq> fst r" by (simp add: leq_prod_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   446
    qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   447
    moreover have "snd p \<squnion> snd q \<sqsubseteq> snd r"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   448
    proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   449
      from "pr" show "snd p \<sqsubseteq> snd r" by (simp add: leq_prod_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   450
      from qr show "snd q \<sqsubseteq> snd r" by (simp add: leq_prod_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   451
    qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   452
    ultimately show ?thesis by (simp add: leq_prod_def)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   453
  qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   454
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   455
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 35317
diff changeset
   456
instance prod :: (lattice, lattice) lattice
10309
a7f961fb62c6 intro_classes by default;
wenzelm
parents: 10183
diff changeset
   457
proof
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   458
  fix p q :: "'a::lattice \<times> 'b::lattice"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   459
  from is_inf_prod show "\<exists>inf. is_inf p q inf" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   460
  from is_sup_prod show "\<exists>sup. is_sup p q sup" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   461
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   462
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   463
text {*
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   464
  The lattice operations on a binary product structure indeed coincide
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   465
  with the products of the original ones.
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   466
*}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   467
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   468
theorem meet_prod: "p \<sqinter> q = (fst p \<sqinter> fst q, snd p \<sqinter> snd q)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   469
  by (rule meet_equality) (rule is_inf_prod)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   470
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   471
theorem join_prod: "p \<squnion> q = (fst p \<squnion> fst q, snd p \<squnion> snd q)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   472
  by (rule join_equality) (rule is_sup_prod)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   473
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   474
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   475
subsubsection {* General products *}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   476
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   477
text {*
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   478
  The class of lattices is closed under general products (function
10158
wenzelm
parents: 10157
diff changeset
   479
  spaces) as well (cf.\ \S\ref{sec:fun-order}).
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   480
*}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   481
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   482
lemma is_inf_fun: "is_inf f g (\<lambda>x. f x \<sqinter> g x)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   483
proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   484
  show "(\<lambda>x. f x \<sqinter> g x) \<sqsubseteq> f"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   485
  proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   486
    fix x show "f x \<sqinter> g x \<sqsubseteq> f x" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   487
  qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   488
  show "(\<lambda>x. f x \<sqinter> g x) \<sqsubseteq> g"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   489
  proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   490
    fix x show "f x \<sqinter> g x \<sqsubseteq> g x" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   491
  qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   492
  fix h assume hf: "h \<sqsubseteq> f" and hg: "h \<sqsubseteq> g"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   493
  show "h \<sqsubseteq> (\<lambda>x. f x \<sqinter> g x)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   494
  proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   495
    fix x
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   496
    show "h x \<sqsubseteq> f x \<sqinter> g x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   497
    proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   498
      from hf show "h x \<sqsubseteq> f x" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   499
      from hg show "h x \<sqsubseteq> g x" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   500
    qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   501
  qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   502
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   503
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   504
lemma is_sup_fun: "is_sup f g (\<lambda>x. f x \<squnion> g x)"   (* FIXME dualize!? *)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   505
proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   506
  show "f \<sqsubseteq> (\<lambda>x. f x \<squnion> g x)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   507
  proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   508
    fix x show "f x \<sqsubseteq> f x \<squnion> g x" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   509
  qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   510
  show "g \<sqsubseteq> (\<lambda>x. f x \<squnion> g x)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   511
  proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   512
    fix x show "g x \<sqsubseteq> f x \<squnion> g x" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   513
  qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   514
  fix h assume fh: "f \<sqsubseteq> h" and gh: "g \<sqsubseteq> h"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   515
  show "(\<lambda>x. f x \<squnion> g x) \<sqsubseteq> h"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   516
  proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   517
    fix x
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   518
    show "f x \<squnion> g x \<sqsubseteq> h x"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   519
    proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   520
      from fh show "f x \<sqsubseteq> h x" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   521
      from gh show "g x \<sqsubseteq> h x" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   522
    qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   523
  qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   524
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   525
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19736
diff changeset
   526
instance "fun" :: (type, lattice) lattice
10309
a7f961fb62c6 intro_classes by default;
wenzelm
parents: 10183
diff changeset
   527
proof
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   528
  fix f g :: "'a \<Rightarrow> 'b::lattice"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   529
  show "\<exists>inf. is_inf f g inf" by rule (rule is_inf_fun) (* FIXME @{text "from \<dots> show \<dots> .."} does not work!? unification incompleteness!? *)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   530
  show "\<exists>sup. is_sup f g sup" by rule (rule is_sup_fun)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   531
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   532
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   533
text {*
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   534
  The lattice operations on a general product structure (function
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   535
  space) indeed emerge by point-wise lifting of the original ones.
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   536
*}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   537
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   538
theorem meet_fun: "f \<sqinter> g = (\<lambda>x. f x \<sqinter> g x)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   539
  by (rule meet_equality) (rule is_inf_fun)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   540
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   541
theorem join_fun: "f \<squnion> g = (\<lambda>x. f x \<squnion> g x)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   542
  by (rule join_equality) (rule is_sup_fun)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   543
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   544
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   545
subsection {* Monotonicity and semi-morphisms *}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   546
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   547
text {*
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   548
  The lattice operations are monotone in both argument positions.  In
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   549
  fact, monotonicity of the second position is trivial due to
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   550
  commutativity.
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   551
*}
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   552
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   553
theorem meet_mono: "x \<sqsubseteq> z \<Longrightarrow> y \<sqsubseteq> w \<Longrightarrow> x \<sqinter> y \<sqsubseteq> z \<sqinter> w"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   554
proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   555
  {
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   556
    fix a b c :: "'a::lattice"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   557
    assume "a \<sqsubseteq> c" have "a \<sqinter> b \<sqsubseteq> c \<sqinter> b"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   558
    proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   559
      have "a \<sqinter> b \<sqsubseteq> a" ..
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   560
      also have "\<dots> \<sqsubseteq> c" by fact
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   561
      finally show "a \<sqinter> b \<sqsubseteq> c" .
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   562
      show "a \<sqinter> b \<sqsubseteq> b" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   563
    qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   564
  } note this [elim?]
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   565
  assume "x \<sqsubseteq> z" then have "x \<sqinter> y \<sqsubseteq> z \<sqinter> y" ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   566
  also have "\<dots> = y \<sqinter> z" by (rule meet_commute)
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   567
  also assume "y \<sqsubseteq> w" then have "y \<sqinter> z \<sqsubseteq> w \<sqinter> z" ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   568
  also have "\<dots> = z \<sqinter> w" by (rule meet_commute)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   569
  finally show ?thesis .
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   570
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   571
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   572
theorem join_mono: "x \<sqsubseteq> z \<Longrightarrow> y \<sqsubseteq> w \<Longrightarrow> x \<squnion> y \<sqsubseteq> z \<squnion> w"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   573
proof -
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   574
  assume "x \<sqsubseteq> z" then have "dual z \<sqsubseteq> dual x" ..
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   575
  moreover assume "y \<sqsubseteq> w" then have "dual w \<sqsubseteq> dual y" ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   576
  ultimately have "dual z \<sqinter> dual w \<sqsubseteq> dual x \<sqinter> dual y"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   577
    by (rule meet_mono)
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   578
  then have "dual (z \<squnion> w) \<sqsubseteq> dual (x \<squnion> y)"
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   579
    by (simp only: dual_join)
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   580
  then show ?thesis ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   581
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   582
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   583
text {*
25469
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   584
  \medskip A semi-morphisms is a function @{text f} that preserves the
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   585
  lattice operations in the following manner: @{term "f (x \<sqinter> y) \<sqsubseteq> f x
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   586
  \<sqinter> f y"} and @{term "f x \<squnion> f y \<sqsubseteq> f (x \<squnion> y)"}, respectively.  Any of
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   587
  these properties is equivalent with monotonicity.
25469
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   588
*}
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   589
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   590
theorem meet_semimorph:
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   591
  "(\<And>x y. f (x \<sqinter> y) \<sqsubseteq> f x \<sqinter> f y) \<equiv> (\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y)"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   592
proof
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   593
  assume morph: "\<And>x y. f (x \<sqinter> y) \<sqsubseteq> f x \<sqinter> f y"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   594
  fix x y :: "'a::lattice"
25469
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   595
  assume "x \<sqsubseteq> y"
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   596
  then have "x \<sqinter> y = x" ..
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   597
  then have "x = x \<sqinter> y" ..
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   598
  also have "f \<dots> \<sqsubseteq> f x \<sqinter> f y" by (rule morph)
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   599
  also have "\<dots> \<sqsubseteq> f y" ..
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   600
  finally show "f x \<sqsubseteq> f y" .
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   601
next
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   602
  assume mono: "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   603
  show "\<And>x y. f (x \<sqinter> y) \<sqsubseteq> f x \<sqinter> f y"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   604
  proof -
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   605
    fix x y
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   606
    show "f (x \<sqinter> y) \<sqsubseteq> f x \<sqinter> f y"
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   607
    proof
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   608
      have "x \<sqinter> y \<sqsubseteq> x" .. then show "f (x \<sqinter> y) \<sqsubseteq> f x" by (rule mono)
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   609
      have "x \<sqinter> y \<sqsubseteq> y" .. then show "f (x \<sqinter> y) \<sqsubseteq> f y" by (rule mono)
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   610
    qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   611
  qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   612
qed
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   613
25469
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   614
lemma join_semimorph:
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   615
  "(\<And>x y. f x \<squnion> f y \<sqsubseteq> f (x \<squnion> y)) \<equiv> (\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y)"
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   616
proof
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   617
  assume morph: "\<And>x y. f x \<squnion> f y \<sqsubseteq> f (x \<squnion> y)"
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   618
  fix x y :: "'a::lattice"
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   619
  assume "x \<sqsubseteq> y" then have "x \<squnion> y = y" ..
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   620
  have "f x \<sqsubseteq> f x \<squnion> f y" ..
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   621
  also have "\<dots> \<sqsubseteq> f (x \<squnion> y)" by (rule morph)
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   622
  also from `x \<sqsubseteq> y` have "x \<squnion> y = y" ..
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   623
  finally show "f x \<sqsubseteq> f y" .
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   624
next
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   625
  assume mono: "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y"
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   626
  show "\<And>x y. f x \<squnion> f y \<sqsubseteq> f (x \<squnion> y)"
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   627
  proof -
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   628
    fix x y
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   629
    show "f x \<squnion> f y \<sqsubseteq> f (x \<squnion> y)"
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   630
    proof
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   631
      have "x \<sqsubseteq> x \<squnion> y" .. then show "f x \<sqsubseteq> f (x \<squnion> y)" by (rule mono)
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   632
      have "y \<sqsubseteq> x \<squnion> y" .. then show "f y \<sqsubseteq> f (x \<squnion> y)" by (rule mono)
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   633
    qed
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   634
  qed
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   635
qed
f81b3be9dfdd some more lemmas due to Peter Lammich;
wenzelm
parents: 23393
diff changeset
   636
10157
6d3987f3aad9 * HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff changeset
   637
end