src/HOL/LOrder.thy
author haftmann
Wed, 15 Nov 2006 17:05:40 +0100
changeset 21381 79e065f2be95
parent 21380 c4f79922bc81
child 21733 131dd2a27137
permissions -rw-r--r--
reworking of min/max lemmas
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
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
     4
*)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
     5
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
     6
header "Lattice Orders"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
     7
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15010
diff changeset
     8
theory LOrder
21249
d594c58e24ed renamed Lattice_Locales to Lattices
haftmann
parents: 21213
diff changeset
     9
imports Lattices
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15010
diff changeset
    10
begin
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    11
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
    12
text {* The theory of lattices developed here is taken from
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
    13
\cite{Birkhoff79}.  *}
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    14
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    15
constdefs
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    16
  is_meet :: "(('a::order) \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    17
  "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
    18
  is_join :: "(('a::order) \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    19
  "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
    20
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    21
lemma is_meet_unique: 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    22
  assumes "is_meet u" "is_meet v" shows "u = v"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    23
proof -
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    24
  {
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    25
    fix a b :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    26
    assume a: "is_meet a"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    27
    assume b: "is_meet b"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    28
    {
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    29
      fix x y 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    30
      let ?za = "a x y"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    31
      let ?zb = "b x y"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    32
      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
    33
      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
    34
    }
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    35
  }
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    36
  note f_le = this
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    37
  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
    38
qed
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
lemma is_join_unique: 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    41
  assumes "is_join u" "is_join v" shows "u = v"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    42
proof -
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    43
  {
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    44
    fix a b :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    45
    assume a: "is_join a"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    46
    assume b: "is_join b"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    47
    {
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    48
      fix x y 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    49
      let ?za = "a x y"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    50
      let ?zb = "b x y"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    51
      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
    52
      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
    53
    }
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    54
  }
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    55
  note f_le = this
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    56
  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
    57
qed
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
axclass join_semilorder < order
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    60
  join_exists: "? j. is_join j"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    61
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    62
axclass meet_semilorder < order
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    63
  meet_exists: "? m. is_meet m"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    64
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    65
axclass lorder < join_semilorder, meet_semilorder
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
constdefs
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    68
  meet :: "('a::meet_semilorder) \<Rightarrow> 'a \<Rightarrow> 'a"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    69
  "meet == THE m. is_meet m"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    70
  join :: "('a::join_semilorder) \<Rightarrow> 'a \<Rightarrow> 'a"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    71
  "join ==  THE j. is_join j"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    72
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    73
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
    74
proof -
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    75
  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
    76
  with is_meet_unique[of _ k] show ?thesis
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    77
    by (simp add: meet_def theI[of is_meet])    
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    78
qed
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    79
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    80
lemma meet_unique: "(is_meet m) = (m = meet)" 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    81
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
    82
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    83
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
    84
proof -
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    85
  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
    86
  with is_join_unique[of _ k] show ?thesis
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    87
    by (simp add: join_def theI[of is_join])    
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    88
qed
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    89
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    90
lemma join_unique: "(is_join j) = (j = join)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
    91
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
    92
21380
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
    93
interpretation lattice:
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
    94
  lattice ["op \<le> \<Colon> 'a\<Colon>lorder \<Rightarrow> 'a \<Rightarrow> bool" "op <" meet join]
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
    95
proof unfold_locales
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
    96
  fix x y z :: "'a\<Colon>lorder"
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
    97
  from is_meet_meet have "is_meet meet" by blast
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
    98
  note meet = this is_meet_def
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
    99
  from meet show "meet x y \<le> x" by blast
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
   100
  from meet show "meet x y \<le> y" by blast
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
   101
  from meet show "x \<le> y \<Longrightarrow> x \<le> z \<Longrightarrow> x \<le> meet y z" by blast
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
   102
  from is_join_join have "is_join join" by blast
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
   103
  note join = this is_join_def
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
   104
  from join show "x \<le> join x y" by blast
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
   105
  from join show "y \<le> join x y" by blast
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
   106
  from join show "y \<le> x \<Longrightarrow> z \<le> x \<Longrightarrow> join y z \<le> x" by blast
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
   107
qed
c4f79922bc81 added interpretation
haftmann
parents: 21312
diff changeset
   108
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   109
lemma meet_left_le: "meet a b \<le> (a::'a::meet_semilorder)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   110
by (insert is_meet_meet, auto simp add: is_meet_def)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   111
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   112
lemma meet_right_le: "meet a b \<le> (b::'a::meet_semilorder)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   113
by (insert is_meet_meet, auto simp add: is_meet_def)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   114
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   115
(* intro! breaks a proof in Hyperreal/SEQ and NumberTheory/IntPrimes *)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   116
lemma le_meetI:
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   117
 "x \<le> a \<Longrightarrow> x \<le> b \<Longrightarrow> x \<le> meet a (b::'a::meet_semilorder)"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   118
by (insert is_meet_meet, auto simp add: is_meet_def)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   119
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   120
lemma join_left_le: "a \<le> join a (b::'a::join_semilorder)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   121
by (insert is_join_join, auto simp add: is_join_def)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   122
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   123
lemma join_right_le: "b \<le> join a (b::'a::join_semilorder)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   124
by (insert is_join_join, auto simp add: is_join_def)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   125
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   126
lemma join_leI:
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   127
 "a \<le> x \<Longrightarrow> b \<le> x \<Longrightarrow> join a b \<le> (x::'a::join_semilorder)"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   128
by (insert is_join_join, auto simp add: is_join_def)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   129
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   130
lemmas meet_join_le[simp] = meet_left_le meet_right_le join_left_le join_right_le
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   131
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   132
lemma le_meet[simp]: "(x <= meet y z) = (x <= y & x <= z)" (is "?L = ?R")
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   133
proof
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   134
  assume ?L
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   135
  moreover have "meet y z \<le> y" "meet y z <= z" by(simp_all)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   136
  ultimately show ?R by(blast intro:order_trans)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   137
next
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   138
  assume ?R thus ?L by (blast intro!:le_meetI)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   139
qed
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   140
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   141
lemma join_le[simp]: "(join x y <= z) = (x <= z & y <= z)" (is "?L = ?R")
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   142
proof
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   143
  assume ?L
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   144
  moreover have "x \<le> join x y" "y \<le> join x y" by(simp_all)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   145
  ultimately show ?R by(blast intro:order_trans)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   146
next
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   147
  assume ?R thus ?L by (blast intro:join_leI)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   148
qed
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   149
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   150
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   151
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
   152
by (auto simp add: is_meet_def min_def)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   153
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   154
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
   155
by (auto simp add: is_join_def max_def)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   156
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   157
instance linorder \<subseteq> meet_semilorder
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   158
proof
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   159
  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
   160
qed
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   161
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   162
instance linorder \<subseteq> join_semilorder
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   163
proof
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   164
  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
   165
qed
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   166
    
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   167
instance linorder \<subseteq> lorder ..
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   168
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   169
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
   170
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
   171
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   172
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
   173
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
   174
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   175
lemma meet_idempotent[simp]: "meet x x = x"
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   176
by (rule order_antisym, simp_all add: le_meetI)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   177
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   178
lemma join_idempotent[simp]: "join x x = x"
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   179
by (rule order_antisym, simp_all add: join_leI)
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
lemma meet_comm: "meet x y = meet y x" 
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   182
by (rule order_antisym, (simp add: le_meetI)+)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   183
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   184
lemma join_comm: "join x y = join y x"
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   185
by (rule order_antisym, (simp add: join_leI)+)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   186
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   187
lemma meet_leI1: "x \<le> z \<Longrightarrow> meet x y \<le> z"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   188
apply(subgoal_tac "meet x y <= x")
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   189
 apply(blast intro:order_trans)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   190
apply simp
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   191
done
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   192
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   193
lemma meet_leI2: "y \<le> z \<Longrightarrow> meet x y \<le> z"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   194
apply(subgoal_tac "meet x y <= y")
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   195
 apply(blast intro:order_trans)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   196
apply simp
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   197
done
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   198
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   199
lemma le_joinI1: "x \<le> y \<Longrightarrow> x \<le> join y z"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   200
apply(subgoal_tac "y <= join y z")
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   201
 apply(blast intro:order_trans)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   202
apply simp
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   203
done
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   204
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   205
lemma le_joinI2: "x \<le> z \<Longrightarrow> x \<le> join y z"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   206
apply(subgoal_tac "z <= join y z")
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   207
 apply(blast intro:order_trans)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   208
apply simp
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   209
done
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   210
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   211
lemma meet_assoc: "meet (meet x y) z = meet x (meet y z)"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   212
apply(rule order_antisym)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   213
apply (simp add:meet_leI1 meet_leI2)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   214
apply (simp add:meet_leI1 meet_leI2)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   215
done
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   216
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   217
lemma join_assoc: "join (join x y) z = join x (join y z)"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   218
apply(rule order_antisym)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   219
apply (simp add:le_joinI1 le_joinI2)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   220
apply (simp add:le_joinI1 le_joinI2)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   221
done
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   222
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   223
lemma meet_left_comm: "meet a (meet b c) = meet b (meet a c)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   224
by (simp add: meet_assoc[symmetric, of a b c], simp add: meet_comm[of a b], simp add: meet_assoc)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   225
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   226
lemma meet_left_idempotent: "meet y (meet y x) = meet y x"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   227
by (simp add: meet_assoc meet_comm meet_left_comm)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   228
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   229
lemma join_left_comm: "join a (join b c) = join b (join a c)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   230
by (simp add: join_assoc[symmetric, of a b c], simp add: join_comm[of a b], simp add: join_assoc)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   231
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   232
lemma join_left_idempotent: "join y (join y x) = join y x"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   233
by (simp add: join_assoc join_comm join_left_comm)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   234
    
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   235
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
   236
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   237
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
   238
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   239
lemma le_def_meet: "(x <= y) = (meet x y = x)"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   240
apply rule
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   241
apply(simp add: order_antisym)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   242
apply(subgoal_tac "meet x y <= y")
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   243
apply(simp)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   244
apply(simp (no_asm))
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   245
done
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   246
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   247
lemma le_def_join: "(x <= y) = (join x y = y)"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   248
apply rule
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   249
apply(simp add: order_antisym)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   250
apply(subgoal_tac "x <= join x y")
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   251
apply(simp)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   252
apply(simp (no_asm))
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   253
done
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   254
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   255
lemma join_absorp2: "a \<le> b \<Longrightarrow> join a b = b" 
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   256
by (simp add: le_def_join)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   257
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   258
lemma join_absorp1: "b \<le> a \<Longrightarrow> join a b = a"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   259
by (simp add: le_def_join join_aci)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   260
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   261
lemma meet_absorp1: "a \<le> b \<Longrightarrow> meet a b = a"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   262
by (simp add: le_def_meet)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   263
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   264
lemma meet_absorp2: "b \<le> a \<Longrightarrow> meet a b = b"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   265
by (simp add: le_def_meet meet_aci)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   266
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   267
lemma meet_join_absorp: "meet x (join x y) = x"
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   268
by(simp add:meet_absorp1)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   269
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   270
lemma join_meet_absorp: "join x (meet x y) = x"
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   271
by(simp add:join_absorp1)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   272
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   273
lemma meet_mono: "y \<le> z \<Longrightarrow> meet x y \<le> meet x z"
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   274
by(simp add:meet_leI2)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   275
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   276
lemma join_mono: "y \<le> z \<Longrightarrow> join x y \<le> join x z"
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   277
by(simp add:le_joinI2)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   278
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   279
lemma distrib_join_le: "join x (meet y z) \<le> meet (join x y) (join x z)" (is "_ <= ?r")
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   280
proof -
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   281
  have a: "x <= ?r" by (simp_all add:le_meetI)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   282
  have b: "meet y z <= ?r" by (simp add:le_joinI2)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   283
  from a b show ?thesis by (simp add: join_leI)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   284
qed
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   285
  
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   286
lemma distrib_meet_le: "join (meet x y) (meet x z) \<le> meet x (join y z)" (is "?l <= _")
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   287
proof -
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   288
  have a: "?l <= x" by (simp_all add: join_leI)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   289
  have b: "?l <= join y z" by (simp add:meet_leI2)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   290
  from a b show ?thesis by (simp add: le_meetI)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   291
qed
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   292
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   293
lemma meet_join_eq_imp_le: "a = c \<or> a = d \<or> b = c \<or> b = d \<Longrightarrow> meet a b \<le> join c d"
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   294
by (auto simp:meet_leI2 meet_leI1)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   295
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   296
lemma modular_le: "x \<le> z \<Longrightarrow> join x (meet y z) \<le> meet (join x y) z" (is "_ \<Longrightarrow> ?t <= _")
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   297
proof -
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   298
  assume a: "x <= z"
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   299
  have b: "?t <= join x y" by (simp_all add: join_leI meet_join_eq_imp_le )
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   300
  have c: "?t <= z" by (simp_all add: a join_leI)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21249
diff changeset
   301
  from b c show ?thesis by (simp add: le_meetI)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   302
qed
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents:
diff changeset
   303
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15010
diff changeset
   304
end