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