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