src/HOL/LOrder.thy
author wenzelm
Sat, 30 Dec 2006 16:08:06 +0100
changeset 21966 edab0ecfbd7c
parent 21734 283461c15fa7
child 22422 ee19cdb07528
permissions -rw-r--r--
removed misleading OuterLex.eq_token;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
     1
(*  Title:   HOL/LOrder.thy
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
     2
    ID:      $Id$
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
     3
    Author:  Steven Obua, TU Muenchen
21734
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
     4
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
     5
FIXME integrate properly with lattice locales
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
     6
- make it a class?
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
     7
- get rid of the implicit there-is-a-meet/join but require eplicit ops.
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
     8
Also rename meet/join to inf/sup. 
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
     9
*)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    10
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
    11
header "Lattice Orders"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    12
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15010
diff changeset
    13
theory LOrder
21249
d594c58e24ed renamed Lattice_Locales to Lattices
haftmann
parents: 21213
diff changeset
    14
imports Lattices
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15010
diff changeset
    15
begin
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    16
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
    17
text {* The theory of lattices developed here is taken from
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
    18
\cite{Birkhoff79}.  *}
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    19
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    20
constdefs
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    21
  is_meet :: "(('a::order) \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    22
  "is_meet m == ! a b x. m a b \<le> a \<and> m a b \<le> b \<and> (x \<le> a \<and> x \<le> b \<longrightarrow> x \<le> m a b)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    23
  is_join :: "(('a::order) \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    24
  "is_join j == ! a b x. a \<le> j a b \<and> b \<le> j a b \<and> (a \<le> x \<and> b \<le> x \<longrightarrow> j a b \<le> x)"  
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    25
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    26
lemma is_meet_unique: 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    27
  assumes "is_meet u" "is_meet v" shows "u = v"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    28
proof -
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    29
  {
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    30
    fix a b :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    31
    assume a: "is_meet a"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    32
    assume b: "is_meet b"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    33
    {
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    34
      fix x y 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    35
      let ?za = "a x y"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    36
      let ?zb = "b x y"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    37
      from a have za_le: "?za <= x & ?za <= y" by (auto simp add: is_meet_def)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    38
      with b have "?za <= ?zb" by (auto simp add: is_meet_def)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    39
    }
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    40
  }
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    41
  note f_le = this
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    42
  show "u = v" by ((rule ext)+, simp_all add: order_antisym prems f_le) 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    43
qed
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    44
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    45
lemma is_join_unique: 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    46
  assumes "is_join u" "is_join v" shows "u = v"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    47
proof -
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    48
  {
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    49
    fix a b :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    50
    assume a: "is_join a"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    51
    assume b: "is_join b"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    52
    {
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    53
      fix x y 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    54
      let ?za = "a x y"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    55
      let ?zb = "b x y"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    56
      from a have za_le: "x <= ?za & y <= ?za" by (auto simp add: is_join_def)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    57
      with b have "?zb <= ?za" by (auto simp add: is_join_def)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    58
    }
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    59
  }
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    60
  note f_le = this
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    61
  show "u = v" by ((rule ext)+, simp_all add: order_antisym prems f_le) 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    62
qed
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    63
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    64
axclass join_semilorder < order
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    65
  join_exists: "? j. is_join j"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    66
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    67
axclass meet_semilorder < order
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    68
  meet_exists: "? m. is_meet m"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    69
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    70
axclass lorder < join_semilorder, meet_semilorder
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    71
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    72
constdefs
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    73
  meet :: "('a::meet_semilorder) \<Rightarrow> 'a \<Rightarrow> 'a"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    74
  "meet == THE m. is_meet m"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    75
  join :: "('a::join_semilorder) \<Rightarrow> 'a \<Rightarrow> 'a"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    76
  "join ==  THE j. is_join j"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    77
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    78
lemma is_meet_meet: "is_meet (meet::'a \<Rightarrow> 'a \<Rightarrow> ('a::meet_semilorder))"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    79
proof -
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    80
  from meet_exists obtain k::"'a \<Rightarrow> 'a \<Rightarrow> 'a" where "is_meet k" ..
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    81
  with is_meet_unique[of _ k] show ?thesis
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    82
    by (simp add: meet_def theI[of is_meet])    
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    83
qed
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    84
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    85
lemma meet_unique: "(is_meet m) = (m = meet)" 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    86
by (insert is_meet_meet, auto simp add: is_meet_unique)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    87
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    88
lemma is_join_join: "is_join (join::'a \<Rightarrow> 'a \<Rightarrow> ('a::join_semilorder))"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    89
proof -
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    90
  from join_exists obtain k::"'a \<Rightarrow> 'a \<Rightarrow> 'a" where "is_join k" ..
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    91
  with is_join_unique[of _ k] show ?thesis
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    92
    by (simp add: join_def theI[of is_join])    
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    93
qed
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    94
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    95
lemma join_unique: "(is_join j) = (j = join)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    96
by (insert is_join_join, auto simp add: is_join_unique)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    97
21733
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
    98
interpretation meet_semilat:
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
    99
  lower_semilattice ["op \<le> \<Colon> 'a\<Colon>meet_semilorder \<Rightarrow> 'a \<Rightarrow> bool" "op <" meet]
21380
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
   100
proof unfold_locales
21733
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   101
  fix x y z :: "'a\<Colon>meet_semilorder"
21380
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
   102
  from is_meet_meet have "is_meet meet" by blast
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
   103
  note meet = this is_meet_def
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
   104
  from meet show "meet x y \<le> x" by blast
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
   105
  from meet show "meet x y \<le> y" by blast
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
   106
  from meet show "x \<le> y \<Longrightarrow> x \<le> z \<Longrightarrow> x \<le> meet y z" by blast
21733
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   107
qed
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   108
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   109
interpretation join_semilat:
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   110
  upper_semilattice ["op \<le> \<Colon> 'a\<Colon>join_semilorder \<Rightarrow> 'a \<Rightarrow> bool" "op <" join]
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   111
proof unfold_locales
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   112
  fix x y z :: "'a\<Colon>join_semilorder"
21380
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
   113
  from is_join_join have "is_join join" by blast
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
   114
  note join = this is_join_def
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
   115
  from join show "x \<le> join x y" by blast
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
   116
  from join show "y \<le> join x y" by blast
21733
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   117
  from join show "x \<le> z \<Longrightarrow> y \<le> z \<Longrightarrow> join x y \<le> z" by blast
21380
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
   118
qed
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
   119
21733
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   120
declare
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   121
 join_semilat.antisym_intro[rule del] meet_semilat.antisym_intro[rule del]
21734
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
   122
 join_semilat.le_supE[rule del] meet_semilat.le_infE[rule del]
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   123
21733
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   124
interpretation meet_join_lat:
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   125
  lattice ["op \<le> \<Colon> 'a\<Colon>lorder \<Rightarrow> 'a \<Rightarrow> bool" "op <" meet join]
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   126
by unfold_locales
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   127
21733
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   128
lemmas meet_left_le = meet_semilat.inf_le1
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   129
lemmas meet_right_le = meet_semilat.inf_le2
21734
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
   130
lemmas le_meetI[rule del] = meet_semilat.le_infI
21733
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   131
(* intro! breaks a proof in Hyperreal/SEQ and NumberTheory/IntPrimes *)
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   132
lemmas join_left_le = join_semilat.sup_ge1
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   133
lemmas join_right_le = join_semilat.sup_ge2
21734
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
   134
lemmas join_leI[rule del] = join_semilat.le_supI
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   135
21733
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   136
lemmas meet_join_le = meet_left_le meet_right_le join_left_le join_right_le
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   137
21734
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
   138
lemmas le_meet = meet_semilat.le_inf_iff
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   139
21734
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
   140
lemmas le_join = join_semilat.ge_sup_conv
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   141
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   142
lemma is_meet_min: "is_meet (min::'a \<Rightarrow> 'a \<Rightarrow> ('a::linorder))"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   143
by (auto simp add: is_meet_def min_def)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   144
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   145
lemma is_join_max: "is_join (max::'a \<Rightarrow> 'a \<Rightarrow> ('a::linorder))"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   146
by (auto simp add: is_join_def max_def)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   147
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   148
instance linorder \<subseteq> meet_semilorder
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   149
proof
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   150
  from is_meet_min show "? (m::'a\<Rightarrow>'a\<Rightarrow>('a::linorder)). is_meet m" by auto
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   151
qed
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   152
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   153
instance linorder \<subseteq> join_semilorder
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   154
proof
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   155
  from is_join_max show "? (j::'a\<Rightarrow>'a\<Rightarrow>('a::linorder)). is_join j" by auto 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   156
qed
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   157
    
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   158
instance linorder \<subseteq> lorder ..
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   159
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   160
lemma meet_min: "meet = (min :: 'a\<Rightarrow>'a\<Rightarrow>('a::linorder))" 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   161
by (simp add: is_meet_meet is_meet_min is_meet_unique)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   162
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   163
lemma join_max: "join = (max :: 'a\<Rightarrow>'a\<Rightarrow>('a::linorder))"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   164
by (simp add: is_join_join is_join_max is_join_unique)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   165
21733
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   166
lemmas meet_idempotent = meet_semilat.inf_idem
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   167
lemmas join_idempotent = join_semilat.sup_idem
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   168
lemmas meet_comm = meet_semilat.inf_commute
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   169
lemmas join_comm = join_semilat.sup_commute
21734
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
   170
lemmas meet_leI1[rule del] = meet_semilat.le_infI1
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
   171
lemmas meet_leI2[rule del] = meet_semilat.le_infI2
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
   172
lemmas le_joinI1[rule del] = join_semilat.le_supI1
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
   173
lemmas le_joinI2[rule del] = join_semilat.le_supI2
21733
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   174
lemmas meet_assoc = meet_semilat.inf_assoc
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   175
lemmas join_assoc = join_semilat.sup_assoc
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   176
lemmas meet_left_comm = meet_semilat.inf_left_commute
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   177
lemmas meet_left_idempotent = meet_semilat.inf_left_idem
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   178
lemmas join_left_comm = join_semilat.sup_left_commute
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   179
lemmas join_left_idempotent= join_semilat.sup_left_idem
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   180
    
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   181
lemmas meet_aci = meet_assoc meet_comm meet_left_comm meet_left_idempotent
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   182
lemmas join_aci = join_assoc join_comm join_left_comm join_left_idempotent
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   183
21734
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
   184
lemmas le_def_meet = meet_semilat.le_iff_inf
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
   185
lemmas le_def_join = join_semilat.le_iff_sup
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   186
21733
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   187
lemmas join_absorp2 = join_semilat.sup_absorb2
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   188
lemmas join_absorp1 = join_semilat.sup_absorb1
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   189
21733
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   190
lemmas meet_absorp1 = meet_semilat.inf_absorb1
131dd2a27137 Modified lattice locale
nipkow
parents: 21380
diff changeset
   191
lemmas meet_absorp2 = meet_semilat.inf_absorb2
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   192
21734
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
   193
interpretation meet_join_lat:
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
   194
  lattice ["op \<le> \<Colon> 'a\<Colon>{meet_semilorder,join_semilorder} \<Rightarrow> 'a \<Rightarrow> bool" "op <" meet join]
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
   195
by unfold_locales
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   196
21734
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
   197
lemmas meet_join_absorp = meet_join_lat.inf_sup_absorb
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
   198
lemmas join_meet_absorp = meet_join_lat.sup_inf_absorb
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   199
21734
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
   200
lemmas distrib_join_le = meet_join_lat.distrib_sup_le
283461c15fa7 renaming
nipkow
parents: 21733
diff changeset
   201
lemmas distrib_meet_le = meet_join_lat.distrib_inf_le
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   202
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15010
diff changeset
   203
end