src/HOL/Algebra/Lattice.thy
author ballarin
Wed, 30 Jul 2008 19:03:33 +0200
changeset 27713 95b36bfe7fc4
parent 27700 ef4b26efa8b6
child 27714 27b4d7c01f8b
permissions -rw-r--r--
New locales for orders and lattices where the equivalence relation is not restricted to equality.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
     1
(*
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
     2
  Title:     HOL/Algebra/GLattice.thy
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
     3
  Id:        $Id$
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
     4
  Author:    Clemens Ballarin, started 7 November 2003
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
     5
  Copyright: Clemens Ballarin
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
     6
*)
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
     7
27700
ef4b26efa8b6 Renamed theorems;
ballarin
parents: 26805
diff changeset
     8
theory Lattice imports Congruence begin
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
     9
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19984
diff changeset
    10
section {* Orders and Lattices *}
14751
0d7850e27fed Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents: 14706
diff changeset
    11
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
    12
subsection {* Partial Orders *}
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
    13
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    14
record 'a gorder = "'a eq_object" +
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
    15
  le :: "['a, 'a] => bool" (infixl "\<sqsubseteq>\<index>" 50)
21041
60e418260b4d Order and lattice structures no longer based on records.
ballarin
parents: 20318
diff changeset
    16
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    17
locale weak_partial_order = equivalence L +
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    18
  assumes le_refl [intro, simp]:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    19
      "x \<in> carrier L ==> x \<sqsubseteq> x"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    20
    and weak_le_anti_sym [intro]:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    21
      "[| x \<sqsubseteq> y; y \<sqsubseteq> x; x \<in> carrier L; y \<in> carrier L |] ==> x .= y"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    22
    and le_trans [trans]:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    23
      "[| x \<sqsubseteq> y; y \<sqsubseteq> z; x \<in> carrier L; y \<in> carrier L; z \<in> carrier L |] ==> x \<sqsubseteq> z"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    24
    and le_cong:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    25
      "\<lbrakk> x .= y; z .= w; x \<in> carrier L; y \<in> carrier L; z \<in> carrier L; w \<in> carrier L \<rbrakk> \<Longrightarrow> x \<sqsubseteq> z \<longleftrightarrow> y \<sqsubseteq> w"
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
    26
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
    27
constdefs (structure L)
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
    28
  lless :: "[_, 'a, 'a] => bool" (infixl "\<sqsubset>\<index>" 50)
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    29
  "x \<sqsubset> y == x \<sqsubseteq> y & x .\<noteq> y"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    30
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    31
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    32
subsubsection {* The order relation *}
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    33
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    34
context weak_partial_order begin
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    35
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    36
lemma le_cong_l [intro, trans]:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    37
  "\<lbrakk> x .= y; y \<sqsubseteq> z; x \<in> carrier L; y \<in> carrier L; z \<in> carrier L \<rbrakk> \<Longrightarrow> x \<sqsubseteq> z"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    38
  by (auto intro: le_cong [THEN iffD2])
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    39
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    40
lemma le_cong_r [intro, trans]:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    41
  "\<lbrakk> x \<sqsubseteq> y; y .= z; x \<in> carrier L; y \<in> carrier L; z \<in> carrier L \<rbrakk> \<Longrightarrow> x \<sqsubseteq> z"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    42
  by (auto intro: le_cong [THEN iffD1])
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    43
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    44
lemma gen_refl [intro, simp]: "\<lbrakk> x .= y; x \<in> carrier L; y \<in> carrier L \<rbrakk> \<Longrightarrow> x \<sqsubseteq> y"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    45
  by (simp add: le_cong_l)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    46
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    47
end
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    48
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    49
lemma weak_llessI:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    50
  fixes R (structure)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    51
  assumes "x \<sqsubseteq> y" and "~(x .= y)"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    52
  shows "x \<sqsubset> y"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    53
  using assms unfolding lless_def by simp
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    54
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    55
lemma lless_imp_le:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    56
  fixes R (structure)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    57
  assumes "x \<sqsubset> y"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    58
  shows "x \<sqsubseteq> y"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    59
  using assms unfolding lless_def by simp
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    60
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    61
lemma weak_lless_imp_not_eq:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    62
  fixes R (structure)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    63
  assumes "x \<sqsubset> y"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    64
  shows "\<not> (x .= y)"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    65
  using assms unfolding lless_def by simp
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
    66
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    67
lemma weak_llessE:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    68
  fixes R (structure)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    69
  assumes p: "x \<sqsubset> y" and e: "\<lbrakk>x \<sqsubseteq> y; \<not> (x .= y)\<rbrakk> \<Longrightarrow> P"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    70
  shows "P"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    71
  using p by (blast dest: lless_imp_le weak_lless_imp_not_eq e)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    72
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    73
lemma (in weak_partial_order) lless_cong_l [trans]:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    74
  assumes xx': "x .= x'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    75
    and xy: "x' \<sqsubset> y"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    76
    and carr: "x \<in> carrier L" "x' \<in> carrier L" "y \<in> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    77
  shows "x \<sqsubset> y"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    78
  using assms unfolding lless_def by (auto intro: trans sym)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    79
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    80
lemma (in weak_partial_order) lless_cong_r [trans]:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    81
  assumes xy: "x \<sqsubset> y"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    82
    and  yy': "y .= y'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    83
    and carr: "x \<in> carrier L" "y \<in> carrier L" "y' \<in> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    84
  shows "x \<sqsubset> y'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    85
  using assms unfolding lless_def by (auto intro: trans sym)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    86
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    87
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    88
lemma (in weak_partial_order) lless_antisym:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    89
  assumes "a \<in> carrier L" "b \<in> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    90
    and "a \<sqsubset> b" "b \<sqsubset> a"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    91
  shows "P"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    92
  using assms
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    93
  by (elim weak_llessE) auto
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    94
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    95
lemma (in weak_partial_order) lless_trans [trans]:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    96
  assumes "a \<sqsubset> b" "b \<sqsubset> c"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    97
    and carr[simp]: "a \<in> carrier L" "b \<in> carrier L" "c \<in> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    98
  shows "a \<sqsubset> c"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
    99
  using assms unfolding lless_def by (blast dest: le_trans intro: sym)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   100
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   101
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   102
subsubsection {* Upper and lower bounds of a set *}
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   103
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   104
constdefs (structure L)
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   105
  Upper :: "[_, 'a set] => 'a set"
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   106
  "Upper L A == {u. (ALL x. x \<in> A \<inter> carrier L --> x \<sqsubseteq> u)} \<inter> carrier L"
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   107
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   108
  Lower :: "[_, 'a set] => 'a set"
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   109
  "Lower L A == {l. (ALL x. x \<in> A \<inter> carrier L --> l \<sqsubseteq> x)} \<inter> carrier L"
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   110
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   111
lemma Upper_closed [intro!, simp]:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   112
  "Upper L A \<subseteq> carrier L"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   113
  by (unfold Upper_def) clarify
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   114
27700
ef4b26efa8b6 Renamed theorems;
ballarin
parents: 26805
diff changeset
   115
lemma Upper_memD [dest]:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   116
  fixes L (structure)
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   117
  shows "[| u \<in> Upper L A; x \<in> A; A \<subseteq> carrier L |] ==> x \<sqsubseteq> u \<and> u \<in> carrier L"
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   118
  by (unfold Upper_def) blast
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   119
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   120
lemma (in weak_partial_order) Upper_elemD [dest]:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   121
  "[| u .\<in> Upper L A; u \<in> carrier L; x \<in> A; A \<subseteq> carrier L |] ==> x \<sqsubseteq> u"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   122
  unfolding Upper_def elem_def
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   123
  by (blast dest: sym)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   124
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   125
lemma Upper_memI:
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   126
  fixes L (structure)
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   127
  shows "[| !! y. y \<in> A ==> y \<sqsubseteq> x; x \<in> carrier L |] ==> x \<in> Upper L A"
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   128
  by (unfold Upper_def) blast
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   129
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   130
lemma (in weak_partial_order) Upper_elemI:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   131
  "[| !! y. y \<in> A ==> y \<sqsubseteq> x; x \<in> carrier L |] ==> x .\<in> Upper L A"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   132
  unfolding Upper_def by blast
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   133
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   134
lemma Upper_antimono:
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   135
  "A \<subseteq> B ==> Upper L B \<subseteq> Upper L A"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   136
  by (unfold Upper_def) blast
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   137
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   138
lemma (in weak_partial_order) Upper_is_closed [simp]:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   139
  "A \<subseteq> carrier L ==> is_closed (Upper L A)"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   140
  by (rule is_closedI) (blast intro: Upper_memI)+
14651
02b8f3bcf7fe improved notation;
wenzelm
parents: 14577
diff changeset
   141
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   142
lemma (in weak_partial_order) Upper_mem_cong:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   143
  assumes a'carr: "a' \<in> carrier L" and Acarr: "A \<subseteq> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   144
    and aa': "a .= a'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   145
    and aelem: "a \<in> Upper L A"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   146
  shows "a' \<in> Upper L A"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   147
proof (rule Upper_memI[OF _ a'carr])
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   148
  fix y
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   149
  assume yA: "y \<in> A"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   150
  hence "y \<sqsubseteq> a" by (intro Upper_memD[OF aelem, THEN conjunct1] Acarr)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   151
  also note aa'
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   152
  finally
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   153
      show "y \<sqsubseteq> a'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   154
      by (simp add: a'carr subsetD[OF Acarr yA] subsetD[OF Upper_closed aelem])
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   155
qed
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   156
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   157
lemma (in weak_partial_order) Upper_cong:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   158
  assumes Acarr: "A \<subseteq> carrier L" and A'carr: "A' \<subseteq> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   159
    and AA': "A {.=} A'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   160
  shows "Upper L A = Upper L A'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   161
unfolding Upper_def
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   162
apply rule
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   163
 apply (rule, clarsimp) defer 1
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   164
 apply (rule, clarsimp) defer 1
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   165
proof -
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   166
  fix x a'
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   167
  assume carr: "x \<in> carrier L" "a' \<in> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   168
    and a'A': "a' \<in> A'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   169
  assume aLxCond[rule_format]: "\<forall>a. a \<in> A \<and> a \<in> carrier L \<longrightarrow> a \<sqsubseteq> x"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   170
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   171
  from AA' and a'A' have "\<exists>a\<in>A. a' .= a" by (rule set_eqD2)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   172
  from this obtain a
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   173
      where aA: "a \<in> A"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   174
      and a'a: "a' .= a"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   175
      by auto
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   176
  note [simp] = subsetD[OF Acarr aA] carr
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   177
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   178
  note a'a
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   179
  also have "a \<sqsubseteq> x" by (simp add: aLxCond aA)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   180
  finally show "a' \<sqsubseteq> x" by simp
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   181
next
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   182
  fix x a
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   183
  assume carr: "x \<in> carrier L" "a \<in> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   184
    and aA: "a \<in> A"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   185
  assume a'LxCond[rule_format]: "\<forall>a'. a' \<in> A' \<and> a' \<in> carrier L \<longrightarrow> a' \<sqsubseteq> x"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   186
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   187
  from AA' and aA have "\<exists>a'\<in>A'. a .= a'" by (rule set_eqD1)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   188
  from this obtain a'
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   189
      where a'A': "a' \<in> A'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   190
      and aa': "a .= a'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   191
      by auto
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   192
  note [simp] = subsetD[OF A'carr a'A'] carr
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   193
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   194
  note aa'
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   195
  also have "a' \<sqsubseteq> x" by (simp add: a'LxCond a'A')
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   196
  finally show "a \<sqsubseteq> x" by simp
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   197
qed
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   198
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   199
lemma Lower_closed [intro!, simp]:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   200
  "Lower L A \<subseteq> carrier L"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   201
  by (unfold Lower_def) clarify
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   202
27700
ef4b26efa8b6 Renamed theorems;
ballarin
parents: 26805
diff changeset
   203
lemma Lower_memD [dest]:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   204
  fixes L (structure)
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   205
  shows "[| l \<in> Lower L A; x \<in> A; A \<subseteq> carrier L |] ==> l \<sqsubseteq> x \<and> l \<in> carrier L"
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   206
  by (unfold Lower_def) blast
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   207
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   208
lemma Lower_memI:
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   209
  fixes L (structure)
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   210
  shows "[| !! y. y \<in> A ==> x \<sqsubseteq> y; x \<in> carrier L |] ==> x \<in> Lower L A"
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   211
  by (unfold Lower_def) blast
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   212
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   213
lemma Lower_antimono:
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   214
  "A \<subseteq> B ==> Lower L B \<subseteq> Lower L A"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   215
  by (unfold Lower_def) blast
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   216
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   217
lemma (in weak_partial_order) Lower_is_closed [simp]:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   218
  "A \<subseteq> carrier L \<Longrightarrow> is_closed (Lower L A)"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   219
  by (rule is_closedI) (blast intro: Lower_memI dest: sym)+
14651
02b8f3bcf7fe improved notation;
wenzelm
parents: 14577
diff changeset
   220
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   221
lemma (in weak_partial_order) Lower_mem_cong:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   222
  assumes a'carr: "a' \<in> carrier L" and Acarr: "A \<subseteq> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   223
    and aa': "a .= a'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   224
    and aelem: "a \<in> Lower L A"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   225
  shows "a' \<in> Lower L A"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   226
using assms Lower_closed[of L A]
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   227
by (intro Lower_memI) (blast intro: le_cong_l[OF aa'[symmetric]])
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   228
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   229
lemma (in weak_partial_order) Lower_cong:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   230
  assumes Acarr: "A \<subseteq> carrier L" and A'carr: "A' \<subseteq> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   231
    and AA': "A {.=} A'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   232
  shows "Lower L A = Lower L A'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   233
using Lower_memD[of y]
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   234
unfolding Lower_def
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   235
apply safe
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   236
 apply clarsimp defer 1
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   237
 apply clarsimp defer 1
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   238
proof -
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   239
  fix x a'
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   240
  assume carr: "x \<in> carrier L" "a' \<in> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   241
    and a'A': "a' \<in> A'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   242
  assume "\<forall>a. a \<in> A \<and> a \<in> carrier L \<longrightarrow> x \<sqsubseteq> a"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   243
  hence aLxCond: "\<And>a. \<lbrakk>a \<in> A; a \<in> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> a" by fast
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   244
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   245
  from AA' and a'A' have "\<exists>a\<in>A. a' .= a" by (rule set_eqD2)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   246
  from this obtain a
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   247
      where aA: "a \<in> A"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   248
      and a'a: "a' .= a"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   249
      by auto
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   250
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   251
  from aA and subsetD[OF Acarr aA]
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   252
      have "x \<sqsubseteq> a" by (rule aLxCond)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   253
  also note a'a[symmetric]
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   254
  finally
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   255
      show "x \<sqsubseteq> a'" by (simp add: carr subsetD[OF Acarr aA])
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   256
next
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   257
  fix x a
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   258
  assume carr: "x \<in> carrier L" "a \<in> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   259
    and aA: "a \<in> A"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   260
  assume "\<forall>a'. a' \<in> A' \<and> a' \<in> carrier L \<longrightarrow> x \<sqsubseteq> a'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   261
  hence a'LxCond: "\<And>a'. \<lbrakk>a' \<in> A'; a' \<in> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> a'" by fast+
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   262
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   263
  from AA' and aA have "\<exists>a'\<in>A'. a .= a'" by (rule set_eqD1)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   264
  from this obtain a'
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   265
      where a'A': "a' \<in> A'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   266
      and aa': "a .= a'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   267
      by auto
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   268
  from a'A' and subsetD[OF A'carr a'A']
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   269
      have "x \<sqsubseteq> a'" by (rule a'LxCond)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   270
  also note aa'[symmetric]
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   271
  finally show "x \<sqsubseteq> a" by (simp add: carr subsetD[OF A'carr a'A'])
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   272
qed
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   273
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   274
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   275
subsubsection {* Least and greatest, as predicate *}
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   276
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   277
constdefs (structure L)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   278
  least :: "[_, 'a, 'a set] => bool"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   279
  "least L l A == A \<subseteq> carrier L & l \<in> A & (ALL x : A. l \<sqsubseteq> x)"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   280
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   281
  greatest :: "[_, 'a, 'a set] => bool"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   282
  "greatest L g A == A \<subseteq> carrier L & g \<in> A & (ALL x : A. x \<sqsubseteq> g)"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   283
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   284
text {* Could weaken these to @{term [locale=weak_partial_order] "l \<in> carrier L \<and> l
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   285
  .\<in> A"} and @{term [locale=weak_partial_order] "g \<in> carrier L \<and> g .\<in> A"}. *}
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   286
27700
ef4b26efa8b6 Renamed theorems;
ballarin
parents: 26805
diff changeset
   287
lemma least_closed [intro, simp]:
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   288
  "least L l A ==> l \<in> carrier L"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   289
  by (unfold least_def) fast
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   290
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   291
lemma least_mem:
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   292
  "least L l A ==> l \<in> A"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   293
  by (unfold least_def) fast
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   294
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   295
lemma (in weak_partial_order) weak_least_unique:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   296
  "[| least L x A; least L y A |] ==> x .= y"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   297
  by (unfold least_def) blast
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   298
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   299
lemma least_le:
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   300
  fixes L (structure)
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   301
  shows "[| least L x A; a \<in> A |] ==> x \<sqsubseteq> a"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   302
  by (unfold least_def) fast
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   303
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   304
lemma (in weak_partial_order) least_cong:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   305
  "[| x .= x'; x \<in> carrier L; x' \<in> carrier L; is_closed A |] ==> least L x A = least L x' A"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   306
  by (unfold least_def) (auto dest: sym)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   307
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   308
text {* @{const least} is not congruent in the second parameter for 
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   309
  @{term [locale=weak_partial_order] "A {.=} A'"} *}
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   310
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   311
lemma (in weak_partial_order) least_Upper_cong_l:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   312
  assumes "x .= x'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   313
    and "x \<in> carrier L" "x' \<in> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   314
    and "A \<subseteq> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   315
  shows "least L x (Upper L A) = least L x' (Upper L A)"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   316
  apply (rule least_cong) using assms by auto
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   317
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   318
lemma (in weak_partial_order) least_Upper_cong_r:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   319
  assumes Acarrs: "A \<subseteq> carrier L" "A' \<subseteq> carrier L" (* unneccessary with current Upper? *)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   320
    and AA': "A {.=} A'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   321
  shows "least L x (Upper L A) = least L x (Upper L A')"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   322
apply (subgoal_tac "Upper L A = Upper L A'", simp)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   323
by (rule Upper_cong) fact+
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   324
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   325
lemma least_UpperI:
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   326
  fixes L (structure)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   327
  assumes above: "!! x. x \<in> A ==> x \<sqsubseteq> s"
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   328
    and below: "!! y. y \<in> Upper L A ==> s \<sqsubseteq> y"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   329
    and L: "A \<subseteq> carrier L"  "s \<in> carrier L"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   330
  shows "least L s (Upper L A)"
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   331
proof -
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   332
  have "Upper L A \<subseteq> carrier L" by simp
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   333
  moreover from above L have "s \<in> Upper L A" by (simp add: Upper_def)
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   334
  moreover from below have "ALL x : Upper L A. s \<sqsubseteq> x" by fast
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   335
  ultimately show ?thesis by (simp add: least_def)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   336
qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   337
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   338
lemma least_Upper_above:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   339
  fixes L (structure)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   340
  shows "[| least L s (Upper L A); x \<in> A; A \<subseteq> carrier L |] ==> x \<sqsubseteq> s"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   341
  by (unfold least_def) blast
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   342
27700
ef4b26efa8b6 Renamed theorems;
ballarin
parents: 26805
diff changeset
   343
lemma greatest_closed [intro, simp]:
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   344
  "greatest L l A ==> l \<in> carrier L"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   345
  by (unfold greatest_def) fast
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   346
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   347
lemma greatest_mem:
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   348
  "greatest L l A ==> l \<in> A"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   349
  by (unfold greatest_def) fast
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   350
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   351
lemma (in weak_partial_order) weak_greatest_unique:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   352
  "[| greatest L x A; greatest L y A |] ==> x .= y"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   353
  by (unfold greatest_def) blast
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   354
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   355
lemma greatest_le:
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   356
  fixes L (structure)
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   357
  shows "[| greatest L x A; a \<in> A |] ==> a \<sqsubseteq> x"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   358
  by (unfold greatest_def) fast
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   359
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   360
lemma (in weak_partial_order) greatest_cong:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   361
  "[| x .= x'; x \<in> carrier L; x' \<in> carrier L; is_closed A |] ==>
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   362
  greatest L x A = greatest L x' A"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   363
  by (unfold greatest_def) (auto dest: sym)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   364
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   365
text {* @{const greatest} is not congruent in the second parameter for 
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   366
  @{term [locale=weak_partial_order] "A {.=} A'"} *}
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   367
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   368
lemma (in weak_partial_order) greatest_Lower_cong_l:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   369
  assumes "x .= x'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   370
    and "x \<in> carrier L" "x' \<in> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   371
    and "A \<subseteq> carrier L" (* unneccessary with current Lower *)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   372
  shows "greatest L x (Lower L A) = greatest L x' (Lower L A)"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   373
  apply (rule greatest_cong) using assms by auto
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   374
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   375
lemma (in weak_partial_order) greatest_Lower_cong_r:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   376
  assumes Acarrs: "A \<subseteq> carrier L" "A' \<subseteq> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   377
    and AA': "A {.=} A'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   378
  shows "greatest L x (Lower L A) = greatest L x (Lower L A')"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   379
apply (subgoal_tac "Lower L A = Lower L A'", simp)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   380
by (rule Lower_cong) fact+
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   381
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   382
lemma greatest_LowerI:
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   383
  fixes L (structure)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   384
  assumes below: "!! x. x \<in> A ==> i \<sqsubseteq> x"
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   385
    and above: "!! y. y \<in> Lower L A ==> y \<sqsubseteq> i"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   386
    and L: "A \<subseteq> carrier L"  "i \<in> carrier L"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   387
  shows "greatest L i (Lower L A)"
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   388
proof -
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   389
  have "Lower L A \<subseteq> carrier L" by simp
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   390
  moreover from below L have "i \<in> Lower L A" by (simp add: Lower_def)
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   391
  moreover from above have "ALL x : Lower L A. x \<sqsubseteq> i" by fast
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   392
  ultimately show ?thesis by (simp add: greatest_def)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   393
qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   394
27700
ef4b26efa8b6 Renamed theorems;
ballarin
parents: 26805
diff changeset
   395
lemma greatest_Lower_below:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   396
  fixes L (structure)
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   397
  shows "[| greatest L i (Lower L A); x \<in> A; A \<subseteq> carrier L |] ==> i \<sqsubseteq> x"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   398
  by (unfold greatest_def) blast
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   399
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   400
text {* Supremum and infimum *}
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   401
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   402
constdefs (structure L)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   403
  sup :: "[_, 'a set] => 'a" ("\<Squnion>\<index>_" [90] 90)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   404
  "\<Squnion>A == SOME x. least L x (Upper L A)"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   405
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   406
  inf :: "[_, 'a set] => 'a" ("\<Sqinter>\<index>_" [90] 90)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   407
  "\<Sqinter>A == SOME x. greatest L x (Lower L A)"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   408
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   409
  join :: "[_, 'a, 'a] => 'a" (infixl "\<squnion>\<index>" 65)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   410
  "x \<squnion> y == \<Squnion> {x, y}"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   411
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   412
  meet :: "[_, 'a, 'a] => 'a" (infixl "\<sqinter>\<index>" 70)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   413
  "x \<sqinter> y == \<Sqinter> {x, y}"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   414
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   415
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   416
subsection {* Lattices *}
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   417
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   418
locale weak_upper_semilattice = weak_partial_order +
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   419
  assumes sup_of_two_exists:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   420
    "[| x \<in> carrier L; y \<in> carrier L |] ==> EX s. least L s (Upper L {x, y})"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   421
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   422
locale weak_lower_semilattice = weak_partial_order +
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   423
  assumes inf_of_two_exists:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   424
    "[| x \<in> carrier L; y \<in> carrier L |] ==> EX s. greatest L s (Lower L {x, y})"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   425
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   426
locale weak_lattice = weak_upper_semilattice + weak_lower_semilattice
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   427
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   428
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   429
subsubsection {* Supremum *}
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   430
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   431
lemma (in weak_upper_semilattice) joinI:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   432
  "[| !!l. least L l (Upper L {x, y}) ==> P l; x \<in> carrier L; y \<in> carrier L |]
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   433
  ==> P (x \<squnion> y)"
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   434
proof (unfold join_def sup_def)
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   435
  assume L: "x \<in> carrier L"  "y \<in> carrier L"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   436
    and P: "!!l. least L l (Upper L {x, y}) ==> P l"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   437
  with sup_of_two_exists obtain s where "least L s (Upper L {x, y})" by fast
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   438
  with L show "P (SOME l. least L l (Upper L {x, y}))"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   439
    by (fast intro: someI2 P)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   440
qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   441
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   442
lemma (in weak_upper_semilattice) join_closed [simp]:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   443
  "[| x \<in> carrier L; y \<in> carrier L |] ==> x \<squnion> y \<in> carrier L"
27700
ef4b26efa8b6 Renamed theorems;
ballarin
parents: 26805
diff changeset
   444
  by (rule joinI) (rule least_closed)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   445
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   446
lemma (in weak_upper_semilattice) join_cong_l:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   447
  assumes carr: "x \<in> carrier L" "x' \<in> carrier L" "y \<in> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   448
    and xx': "x .= x'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   449
  shows "x \<squnion> y .= x' \<squnion> y"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   450
proof (rule joinI, rule joinI)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   451
  fix a b
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   452
  from xx' carr
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   453
      have seq: "{x, y} {.=} {x', y}" by (rule set_eq_pairI)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   454
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   455
  assume leasta: "least L a (Upper L {x, y})"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   456
  assume "least L b (Upper L {x', y})"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   457
  with carr
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   458
      have leastb: "least L b (Upper L {x, y})"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   459
      by (simp add: least_Upper_cong_r[OF _ _ seq])
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   460
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   461
  from leasta leastb
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   462
      show "a .= b" by (rule weak_least_unique)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   463
qed (rule carr)+
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   464
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   465
lemma (in weak_upper_semilattice) join_cong_r:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   466
  assumes carr: "x \<in> carrier L" "y \<in> carrier L" "y' \<in> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   467
    and yy': "y .= y'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   468
  shows "x \<squnion> y .= x \<squnion> y'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   469
proof (rule joinI, rule joinI)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   470
  fix a b
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   471
  have "{x, y} = {y, x}" by fast
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   472
  also from carr yy'
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   473
      have "{y, x} {.=} {y', x}" by (intro set_eq_pairI)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   474
  also have "{y', x} = {x, y'}" by fast
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   475
  finally
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   476
      have seq: "{x, y} {.=} {x, y'}" .
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   477
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   478
  assume leasta: "least L a (Upper L {x, y})"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   479
  assume "least L b (Upper L {x, y'})"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   480
  with carr
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   481
      have leastb: "least L b (Upper L {x, y})"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   482
      by (simp add: least_Upper_cong_r[OF _ _ seq])
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   483
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   484
  from leasta leastb
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   485
      show "a .= b" by (rule weak_least_unique)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   486
qed (rule carr)+
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   487
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   488
lemma (in weak_partial_order) sup_of_singletonI:      (* only reflexivity needed ? *)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   489
  "x \<in> carrier L ==> least L x (Upper L {x})"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   490
  by (rule least_UpperI) auto
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   491
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   492
lemma (in weak_partial_order) weak_sup_of_singleton [simp]:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   493
  "x \<in> carrier L ==> \<Squnion>{x} .= x"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   494
  unfolding sup_def
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   495
  by (rule someI2) (auto intro: weak_least_unique sup_of_singletonI)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   496
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   497
lemma (in weak_partial_order) sup_of_singleton_closed [simp]:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   498
  "x \<in> carrier L \<Longrightarrow> \<Squnion>{x} \<in> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   499
  unfolding sup_def
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   500
  by (rule someI2) (auto intro: sup_of_singletonI)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   501
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   502
text {* Condition on @{text A}: supremum exists. *}
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   503
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   504
lemma (in weak_upper_semilattice) sup_insertI:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   505
  "[| !!s. least L s (Upper L (insert x A)) ==> P s;
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   506
  least L a (Upper L A); x \<in> carrier L; A \<subseteq> carrier L |]
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   507
  ==> P (\<Squnion>(insert x A))"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   508
proof (unfold sup_def)
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   509
  assume L: "x \<in> carrier L"  "A \<subseteq> carrier L"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   510
    and P: "!!l. least L l (Upper L (insert x A)) ==> P l"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   511
    and least_a: "least L a (Upper L A)"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   512
  from L least_a have La: "a \<in> carrier L" by simp
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   513
  from L sup_of_two_exists least_a
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   514
  obtain s where least_s: "least L s (Upper L {a, x})" by blast
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   515
  show "P (SOME l. least L l (Upper L (insert x A)))"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   516
  proof (rule someI2)
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   517
    show "least L s (Upper L (insert x A))"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   518
    proof (rule least_UpperI)
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   519
      fix z
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   520
      assume "z \<in> insert x A"
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   521
      then show "z \<sqsubseteq> s"
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   522
      proof
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   523
        assume "z = x" then show ?thesis
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   524
          by (simp add: least_Upper_above [OF least_s] L La)
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   525
      next
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   526
        assume "z \<in> A"
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   527
        with L least_s least_a show ?thesis
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   528
          by (rule_tac le_trans [where y = a]) (auto dest: least_Upper_above)
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   529
      qed
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   530
    next
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   531
      fix y
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   532
      assume y: "y \<in> Upper L (insert x A)"
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   533
      show "s \<sqsubseteq> y"
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   534
      proof (rule least_le [OF least_s], rule Upper_memI)
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   535
	fix z
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   536
	assume z: "z \<in> {a, x}"
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   537
	then show "z \<sqsubseteq> y"
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   538
	proof
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   539
          have y': "y \<in> Upper L A"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   540
            apply (rule subsetD [where A = "Upper L (insert x A)"])
23463
9953ff53cc64 tuned proofs -- avoid implicit prems;
wenzelm
parents: 23350
diff changeset
   541
             apply (rule Upper_antimono)
9953ff53cc64 tuned proofs -- avoid implicit prems;
wenzelm
parents: 23350
diff changeset
   542
	     apply blast
9953ff53cc64 tuned proofs -- avoid implicit prems;
wenzelm
parents: 23350
diff changeset
   543
	    apply (rule y)
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   544
            done
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   545
          assume "z = a"
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   546
          with y' least_a show ?thesis by (fast dest: least_le)
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   547
	next
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   548
	  assume "z \<in> {x}"  (* FIXME "z = x"; declare specific elim rule for "insert x {}" (!?) *)
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   549
          with y L show ?thesis by blast
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   550
	qed
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   551
      qed (rule Upper_closed [THEN subsetD, OF y])
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   552
    next
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   553
      from L show "insert x A \<subseteq> carrier L" by simp
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   554
      from least_s show "s \<in> carrier L" by simp
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   555
    qed
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   556
  qed (rule P)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   557
qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   558
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   559
lemma (in weak_upper_semilattice) finite_sup_least:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   560
  "[| finite A; A \<subseteq> carrier L; A ~= {} |] ==> least L (\<Squnion>A) (Upper L A)"
22265
3c5c6bdf61de Adapted to changes in Finite_Set theory.
berghofe
parents: 22063
diff changeset
   561
proof (induct set: finite)
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   562
  case empty
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   563
  then show ?case by simp
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   564
next
15328
35951e6a7855 mod because of change in finite set induction
nipkow
parents: 14751
diff changeset
   565
  case (insert x A)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   566
  show ?case
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   567
  proof (cases "A = {}")
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   568
    case True
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   569
    with insert show ?thesis
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   570
      by simp (simp add: least_cong [OF weak_sup_of_singleton]
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   571
	sup_of_singleton_closed sup_of_singletonI)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   572
	(* The above step is hairy; least_cong can make simp loop.
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   573
	Would want special version of simp to apply least_cong. *)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   574
  next
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   575
    case False
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   576
    with insert have "least L (\<Squnion>A) (Upper L A)" by simp
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   577
    with _ show ?thesis
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   578
      by (rule sup_insertI) (simp_all add: insert [simplified])
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   579
  qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   580
qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   581
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   582
lemma (in weak_upper_semilattice) finite_sup_insertI:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   583
  assumes P: "!!l. least L l (Upper L (insert x A)) ==> P l"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   584
    and xA: "finite A"  "x \<in> carrier L"  "A \<subseteq> carrier L"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   585
  shows "P (\<Squnion> (insert x A))"
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   586
proof (cases "A = {}")
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   587
  case True with P and xA show ?thesis
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   588
    by (simp add: finite_sup_least)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   589
next
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   590
  case False with P and xA show ?thesis
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   591
    by (simp add: sup_insertI finite_sup_least)
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   592
qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   593
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   594
lemma (in weak_upper_semilattice) finite_sup_closed [simp]:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   595
  "[| finite A; A \<subseteq> carrier L; A ~= {} |] ==> \<Squnion>A \<in> carrier L"
22265
3c5c6bdf61de Adapted to changes in Finite_Set theory.
berghofe
parents: 22063
diff changeset
   596
proof (induct set: finite)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   597
  case empty then show ?case by simp
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   598
next
15328
35951e6a7855 mod because of change in finite set induction
nipkow
parents: 14751
diff changeset
   599
  case insert then show ?case
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   600
    by - (rule finite_sup_insertI, simp_all)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   601
qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   602
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   603
lemma (in weak_upper_semilattice) join_left:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   604
  "[| x \<in> carrier L; y \<in> carrier L |] ==> x \<sqsubseteq> x \<squnion> y"
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   605
  by (rule joinI [folded join_def]) (blast dest: least_mem)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   606
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   607
lemma (in weak_upper_semilattice) join_right:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   608
  "[| x \<in> carrier L; y \<in> carrier L |] ==> y \<sqsubseteq> x \<squnion> y"
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   609
  by (rule joinI [folded join_def]) (blast dest: least_mem)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   610
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   611
lemma (in weak_upper_semilattice) sup_of_two_least:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   612
  "[| x \<in> carrier L; y \<in> carrier L |] ==> least L (\<Squnion>{x, y}) (Upper L {x, y})"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   613
proof (unfold sup_def)
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   614
  assume L: "x \<in> carrier L"  "y \<in> carrier L"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   615
  with sup_of_two_exists obtain s where "least L s (Upper L {x, y})" by fast
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   616
  with L show "least L (SOME z. least L z (Upper L {x, y})) (Upper L {x, y})"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   617
  by (fast intro: someI2 weak_least_unique)  (* blast fails *)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   618
qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   619
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   620
lemma (in weak_upper_semilattice) join_le:
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   621
  assumes sub: "x \<sqsubseteq> z"  "y \<sqsubseteq> z"
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   622
    and x: "x \<in> carrier L" and y: "y \<in> carrier L" and z: "z \<in> carrier L"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   623
  shows "x \<squnion> y \<sqsubseteq> z"
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   624
proof (rule joinI [OF _ x y])
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   625
  fix s
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   626
  assume "least L s (Upper L {x, y})"
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   627
  with sub z show "s \<sqsubseteq> z" by (fast elim: least_le intro: Upper_memI)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   628
qed
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   629
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   630
lemma (in weak_upper_semilattice) weak_join_assoc_lemma:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   631
  assumes L: "x \<in> carrier L"  "y \<in> carrier L"  "z \<in> carrier L"
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   632
  shows "x \<squnion> (y \<squnion> z) .= \<Squnion>{x, y, z}"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   633
proof (rule finite_sup_insertI)
14651
02b8f3bcf7fe improved notation;
wenzelm
parents: 14577
diff changeset
   634
  -- {* The textbook argument in Jacobson I, p 457 *}
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   635
  fix s
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   636
  assume sup: "least L s (Upper L {x, y, z})"
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   637
  show "x \<squnion> (y \<squnion> z) .= s"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   638
  proof (rule weak_le_anti_sym)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   639
    from sup L show "x \<squnion> (y \<squnion> z) \<sqsubseteq> s"
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   640
      by (fastsimp intro!: join_le elim: least_Upper_above)
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   641
  next
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   642
    from sup L show "s \<sqsubseteq> x \<squnion> (y \<squnion> z)"
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   643
    by (erule_tac least_le)
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   644
      (blast intro!: Upper_memI intro: le_trans join_left join_right join_closed)
27700
ef4b26efa8b6 Renamed theorems;
ballarin
parents: 26805
diff changeset
   645
  qed (simp_all add: L least_closed [OF sup])
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   646
qed (simp_all add: L)
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   647
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   648
text {* Commutativity holds for @{text "="}. *}
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   649
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   650
lemma join_comm:
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   651
  fixes L (structure)
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   652
  shows "x \<squnion> y = y \<squnion> x"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   653
  by (unfold join_def) (simp add: insert_commute)
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   654
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   655
lemma (in weak_upper_semilattice) weak_join_assoc:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   656
  assumes L: "x \<in> carrier L"  "y \<in> carrier L"  "z \<in> carrier L"
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   657
  shows "(x \<squnion> y) \<squnion> z .= x \<squnion> (y \<squnion> z)"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   658
proof -
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   659
  (* FIXME: could be simplified by improved simp: uniform use of .=,
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   660
     omit [symmetric] in last step. *)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   661
  have "(x \<squnion> y) \<squnion> z = z \<squnion> (x \<squnion> y)" by (simp only: join_comm)
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   662
  also from L have "... .= \<Squnion>{z, x, y}" by (simp add: weak_join_assoc_lemma)
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   663
  also from L have "... = \<Squnion>{x, y, z}" by (simp add: insert_commute)
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   664
  also from L have "... .= x \<squnion> (y \<squnion> z)" by (simp add: weak_join_assoc_lemma [symmetric])
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   665
  finally show ?thesis by (simp add: L)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   666
qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   667
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   668
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   669
subsubsection {* Infimum *}
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   670
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   671
lemma (in weak_lower_semilattice) meetI:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   672
  "[| !!i. greatest L i (Lower L {x, y}) ==> P i;
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   673
  x \<in> carrier L; y \<in> carrier L |]
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   674
  ==> P (x \<sqinter> y)"
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   675
proof (unfold meet_def inf_def)
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   676
  assume L: "x \<in> carrier L"  "y \<in> carrier L"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   677
    and P: "!!g. greatest L g (Lower L {x, y}) ==> P g"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   678
  with inf_of_two_exists obtain i where "greatest L i (Lower L {x, y})" by fast
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   679
  with L show "P (SOME g. greatest L g (Lower L {x, y}))"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   680
  by (fast intro: someI2 weak_greatest_unique P)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   681
qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   682
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   683
lemma (in weak_lower_semilattice) meet_closed [simp]:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   684
  "[| x \<in> carrier L; y \<in> carrier L |] ==> x \<sqinter> y \<in> carrier L"
27700
ef4b26efa8b6 Renamed theorems;
ballarin
parents: 26805
diff changeset
   685
  by (rule meetI) (rule greatest_closed)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   686
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   687
lemma (in weak_lower_semilattice) meet_cong_l:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   688
  assumes carr: "x \<in> carrier L" "x' \<in> carrier L" "y \<in> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   689
    and xx': "x .= x'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   690
  shows "x \<sqinter> y .= x' \<sqinter> y"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   691
proof (rule meetI, rule meetI)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   692
  fix a b
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   693
  from xx' carr
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   694
      have seq: "{x, y} {.=} {x', y}" by (rule set_eq_pairI)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   695
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   696
  assume greatesta: "greatest L a (Lower L {x, y})"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   697
  assume "greatest L b (Lower L {x', y})"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   698
  with carr
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   699
      have greatestb: "greatest L b (Lower L {x, y})"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   700
      by (simp add: greatest_Lower_cong_r[OF _ _ seq])
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   701
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   702
  from greatesta greatestb
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   703
      show "a .= b" by (rule weak_greatest_unique)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   704
qed (rule carr)+
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   705
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   706
lemma (in weak_lower_semilattice) meet_cong_r:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   707
  assumes carr: "x \<in> carrier L" "y \<in> carrier L" "y' \<in> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   708
    and yy': "y .= y'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   709
  shows "x \<sqinter> y .= x \<sqinter> y'"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   710
proof (rule meetI, rule meetI)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   711
  fix a b
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   712
  have "{x, y} = {y, x}" by fast
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   713
  also from carr yy'
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   714
      have "{y, x} {.=} {y', x}" by (intro set_eq_pairI)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   715
  also have "{y', x} = {x, y'}" by fast
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   716
  finally
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   717
      have seq: "{x, y} {.=} {x, y'}" .
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   718
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   719
  assume greatesta: "greatest L a (Lower L {x, y})"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   720
  assume "greatest L b (Lower L {x, y'})"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   721
  with carr
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   722
      have greatestb: "greatest L b (Lower L {x, y})"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   723
      by (simp add: greatest_Lower_cong_r[OF _ _ seq])
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   724
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   725
  from greatesta greatestb
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   726
      show "a .= b" by (rule weak_greatest_unique)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   727
qed (rule carr)+
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   728
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   729
lemma (in weak_partial_order) inf_of_singletonI:      (* only reflexivity needed ? *)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   730
  "x \<in> carrier L ==> greatest L x (Lower L {x})"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   731
  by (rule greatest_LowerI) auto
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   732
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   733
lemma (in weak_partial_order) weak_inf_of_singleton [simp]:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   734
  "x \<in> carrier L ==> \<Sqinter>{x} .= x"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   735
  unfolding inf_def
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   736
  by (rule someI2) (auto intro: weak_greatest_unique inf_of_singletonI)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   737
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   738
lemma (in weak_partial_order) inf_of_singleton_closed:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   739
  "x \<in> carrier L ==> \<Sqinter>{x} \<in> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   740
  unfolding inf_def
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   741
  by (rule someI2) (auto intro: inf_of_singletonI)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   742
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   743
text {* Condition on @{text A}: infimum exists. *}
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   744
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   745
lemma (in weak_lower_semilattice) inf_insertI:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   746
  "[| !!i. greatest L i (Lower L (insert x A)) ==> P i;
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   747
  greatest L a (Lower L A); x \<in> carrier L; A \<subseteq> carrier L |]
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   748
  ==> P (\<Sqinter>(insert x A))"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   749
proof (unfold inf_def)
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   750
  assume L: "x \<in> carrier L"  "A \<subseteq> carrier L"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   751
    and P: "!!g. greatest L g (Lower L (insert x A)) ==> P g"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   752
    and greatest_a: "greatest L a (Lower L A)"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   753
  from L greatest_a have La: "a \<in> carrier L" by simp
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   754
  from L inf_of_two_exists greatest_a
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   755
  obtain i where greatest_i: "greatest L i (Lower L {a, x})" by blast
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   756
  show "P (SOME g. greatest L g (Lower L (insert x A)))"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   757
  proof (rule someI2)
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   758
    show "greatest L i (Lower L (insert x A))"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   759
    proof (rule greatest_LowerI)
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   760
      fix z
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   761
      assume "z \<in> insert x A"
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   762
      then show "i \<sqsubseteq> z"
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   763
      proof
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   764
        assume "z = x" then show ?thesis
27700
ef4b26efa8b6 Renamed theorems;
ballarin
parents: 26805
diff changeset
   765
          by (simp add: greatest_Lower_below [OF greatest_i] L La)
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   766
      next
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   767
        assume "z \<in> A"
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   768
        with L greatest_i greatest_a show ?thesis
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   769
          by (rule_tac le_trans [where y = a]) (auto dest: greatest_Lower_below)
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   770
      qed
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   771
    next
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   772
      fix y
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   773
      assume y: "y \<in> Lower L (insert x A)"
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   774
      show "y \<sqsubseteq> i"
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   775
      proof (rule greatest_le [OF greatest_i], rule Lower_memI)
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   776
	fix z
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   777
	assume z: "z \<in> {a, x}"
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   778
	then show "y \<sqsubseteq> z"
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   779
	proof
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   780
          have y': "y \<in> Lower L A"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   781
            apply (rule subsetD [where A = "Lower L (insert x A)"])
23463
9953ff53cc64 tuned proofs -- avoid implicit prems;
wenzelm
parents: 23350
diff changeset
   782
            apply (rule Lower_antimono)
9953ff53cc64 tuned proofs -- avoid implicit prems;
wenzelm
parents: 23350
diff changeset
   783
	     apply blast
9953ff53cc64 tuned proofs -- avoid implicit prems;
wenzelm
parents: 23350
diff changeset
   784
	    apply (rule y)
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   785
            done
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   786
          assume "z = a"
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   787
          with y' greatest_a show ?thesis by (fast dest: greatest_le)
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   788
	next
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   789
          assume "z \<in> {x}"
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   790
          with y L show ?thesis by blast
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   791
	qed
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   792
      qed (rule Lower_closed [THEN subsetD, OF y])
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   793
    next
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   794
      from L show "insert x A \<subseteq> carrier L" by simp
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   795
      from greatest_i show "i \<in> carrier L" by simp
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   796
    qed
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   797
  qed (rule P)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   798
qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   799
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   800
lemma (in weak_lower_semilattice) finite_inf_greatest:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   801
  "[| finite A; A \<subseteq> carrier L; A ~= {} |] ==> greatest L (\<Sqinter>A) (Lower L A)"
22265
3c5c6bdf61de Adapted to changes in Finite_Set theory.
berghofe
parents: 22063
diff changeset
   802
proof (induct set: finite)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   803
  case empty then show ?case by simp
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   804
next
15328
35951e6a7855 mod because of change in finite set induction
nipkow
parents: 14751
diff changeset
   805
  case (insert x A)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   806
  show ?case
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   807
  proof (cases "A = {}")
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   808
    case True
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   809
    with insert show ?thesis
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   810
      by simp (simp add: greatest_cong [OF weak_inf_of_singleton]
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   811
	inf_of_singleton_closed inf_of_singletonI)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   812
  next
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   813
    case False
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   814
    from insert show ?thesis
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   815
    proof (rule_tac inf_insertI)
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   816
      from False insert show "greatest L (\<Sqinter>A) (Lower L A)" by simp
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   817
    qed simp_all
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   818
  qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   819
qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   820
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   821
lemma (in weak_lower_semilattice) finite_inf_insertI:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   822
  assumes P: "!!i. greatest L i (Lower L (insert x A)) ==> P i"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   823
    and xA: "finite A"  "x \<in> carrier L"  "A \<subseteq> carrier L"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   824
  shows "P (\<Sqinter> (insert x A))"
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   825
proof (cases "A = {}")
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   826
  case True with P and xA show ?thesis
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   827
    by (simp add: finite_inf_greatest)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   828
next
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   829
  case False with P and xA show ?thesis
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   830
    by (simp add: inf_insertI finite_inf_greatest)
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   831
qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   832
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   833
lemma (in weak_lower_semilattice) finite_inf_closed [simp]:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   834
  "[| finite A; A \<subseteq> carrier L; A ~= {} |] ==> \<Sqinter>A \<in> carrier L"
22265
3c5c6bdf61de Adapted to changes in Finite_Set theory.
berghofe
parents: 22063
diff changeset
   835
proof (induct set: finite)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   836
  case empty then show ?case by simp
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   837
next
15328
35951e6a7855 mod because of change in finite set induction
nipkow
parents: 14751
diff changeset
   838
  case insert then show ?case
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   839
    by (rule_tac finite_inf_insertI) (simp_all)
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   840
qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   841
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   842
lemma (in weak_lower_semilattice) meet_left:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   843
  "[| x \<in> carrier L; y \<in> carrier L |] ==> x \<sqinter> y \<sqsubseteq> x"
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   844
  by (rule meetI [folded meet_def]) (blast dest: greatest_mem)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   845
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   846
lemma (in weak_lower_semilattice) meet_right:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   847
  "[| x \<in> carrier L; y \<in> carrier L |] ==> x \<sqinter> y \<sqsubseteq> y"
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   848
  by (rule meetI [folded meet_def]) (blast dest: greatest_mem)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   849
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   850
lemma (in weak_lower_semilattice) inf_of_two_greatest:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   851
  "[| x \<in> carrier L; y \<in> carrier L |] ==>
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   852
  greatest L (\<Sqinter> {x, y}) (Lower L {x, y})"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   853
proof (unfold inf_def)
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   854
  assume L: "x \<in> carrier L"  "y \<in> carrier L"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   855
  with inf_of_two_exists obtain s where "greatest L s (Lower L {x, y})" by fast
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   856
  with L
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   857
  show "greatest L (SOME z. greatest L z (Lower L {x, y})) (Lower L {x, y})"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   858
  by (fast intro: someI2 weak_greatest_unique)  (* blast fails *)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   859
qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   860
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   861
lemma (in weak_lower_semilattice) meet_le:
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   862
  assumes sub: "z \<sqsubseteq> x"  "z \<sqsubseteq> y"
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   863
    and x: "x \<in> carrier L" and y: "y \<in> carrier L" and z: "z \<in> carrier L"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   864
  shows "z \<sqsubseteq> x \<sqinter> y"
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   865
proof (rule meetI [OF _ x y])
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   866
  fix i
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   867
  assume "greatest L i (Lower L {x, y})"
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   868
  with sub z show "z \<sqsubseteq> i" by (fast elim: greatest_le intro: Lower_memI)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   869
qed
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   870
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   871
lemma (in weak_lower_semilattice) weak_meet_assoc_lemma:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   872
  assumes L: "x \<in> carrier L"  "y \<in> carrier L"  "z \<in> carrier L"
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   873
  shows "x \<sqinter> (y \<sqinter> z) .= \<Sqinter>{x, y, z}"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   874
proof (rule finite_inf_insertI)
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   875
  txt {* The textbook argument in Jacobson I, p 457 *}
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   876
  fix i
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   877
  assume inf: "greatest L i (Lower L {x, y, z})"
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   878
  show "x \<sqinter> (y \<sqinter> z) .= i"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   879
  proof (rule weak_le_anti_sym)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   880
    from inf L show "i \<sqsubseteq> x \<sqinter> (y \<sqinter> z)"
27700
ef4b26efa8b6 Renamed theorems;
ballarin
parents: 26805
diff changeset
   881
      by (fastsimp intro!: meet_le elim: greatest_Lower_below)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   882
  next
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   883
    from inf L show "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> i"
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   884
    by (erule_tac greatest_le)
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   885
      (blast intro!: Lower_memI intro: le_trans meet_left meet_right meet_closed)
27700
ef4b26efa8b6 Renamed theorems;
ballarin
parents: 26805
diff changeset
   886
  qed (simp_all add: L greatest_closed [OF inf])
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   887
qed (simp_all add: L)
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   888
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   889
lemma meet_comm:
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   890
  fixes L (structure)
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   891
  shows "x \<sqinter> y = y \<sqinter> x"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   892
  by (unfold meet_def) (simp add: insert_commute)
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   893
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   894
lemma (in weak_lower_semilattice) weak_meet_assoc:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   895
  assumes L: "x \<in> carrier L"  "y \<in> carrier L"  "z \<in> carrier L"
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   896
  shows "(x \<sqinter> y) \<sqinter> z .= x \<sqinter> (y \<sqinter> z)"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   897
proof -
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   898
  (* FIXME: improved simp, see weak_join_assoc above *)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   899
  have "(x \<sqinter> y) \<sqinter> z = z \<sqinter> (x \<sqinter> y)" by (simp only: meet_comm)
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   900
  also from L have "... .= \<Sqinter> {z, x, y}" by (simp add: weak_meet_assoc_lemma)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   901
  also from L have "... = \<Sqinter> {x, y, z}" by (simp add: insert_commute)
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   902
  also from L have "... .= x \<sqinter> (y \<sqinter> z)" by (simp add: weak_meet_assoc_lemma [symmetric])
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   903
  finally show ?thesis by (simp add: L)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   904
qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   905
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   906
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   907
subsection {* Total Orders *}
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   908
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   909
locale weak_total_order = weak_partial_order +
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   910
  assumes total: "[| x \<in> carrier L; y \<in> carrier L |] ==> x \<sqsubseteq> y | y \<sqsubseteq> x"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   911
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   912
text {* Introduction rule: the usual definition of total order *}
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   913
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   914
lemma (in weak_partial_order) weak_total_orderI:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   915
  assumes total: "!!x y. [| x \<in> carrier L; y \<in> carrier L |] ==> x \<sqsubseteq> y | y \<sqsubseteq> x"
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   916
  shows "weak_total_order L"
24087
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   917
  by unfold_locales (rule total)
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   918
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   919
text {* Total orders are lattices. *}
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   920
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   921
interpretation weak_total_order < weak_lattice
24087
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   922
proof unfold_locales
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   923
  fix x y
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   924
  assume L: "x \<in> carrier L"  "y \<in> carrier L"
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   925
  show "EX s. least L s (Upper L {x, y})"
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   926
  proof -
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   927
    note total L
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   928
    moreover
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   929
    {
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   930
      assume "x \<sqsubseteq> y"
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   931
      with L have "least L y (Upper L {x, y})"
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   932
        by (rule_tac least_UpperI) auto
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   933
    }
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   934
    moreover
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   935
    {
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   936
      assume "y \<sqsubseteq> x"
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   937
      with L have "least L x (Upper L {x, y})"
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   938
        by (rule_tac least_UpperI) auto
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   939
    }
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   940
    ultimately show ?thesis by blast
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   941
  qed
24087
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   942
next
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   943
  fix x y
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   944
  assume L: "x \<in> carrier L"  "y \<in> carrier L"
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   945
  show "EX i. greatest L i (Lower L {x, y})"
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   946
  proof -
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   947
    note total L
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   948
    moreover
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   949
    {
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   950
      assume "y \<sqsubseteq> x"
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   951
      with L have "greatest L y (Lower L {x, y})"
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   952
        by (rule_tac greatest_LowerI) auto
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   953
    }
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   954
    moreover
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   955
    {
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   956
      assume "x \<sqsubseteq> y"
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   957
      with L have "greatest L x (Lower L {x, y})"
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   958
        by (rule_tac greatest_LowerI) auto
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   959
    }
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   960
    ultimately show ?thesis by blast
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   961
  qed
eb025d149a34 Proper interpretation of total orders in lattices.
ballarin
parents: 23463
diff changeset
   962
qed
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   963
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
   964
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   965
subsection {* Complete lattices *}
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   966
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   967
locale weak_complete_lattice = weak_lattice +
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   968
  assumes sup_exists:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   969
    "[| A \<subseteq> carrier L |] ==> EX s. least L s (Upper L A)"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   970
    and inf_exists:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   971
    "[| A \<subseteq> carrier L |] ==> EX i. greatest L i (Lower L A)"
21041
60e418260b4d Order and lattice structures no longer based on records.
ballarin
parents: 20318
diff changeset
   972
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   973
text {* Introduction rule: the usual definition of complete lattice *}
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   974
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   975
lemma (in weak_partial_order) weak_complete_latticeI:
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   976
  assumes sup_exists:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   977
    "!!A. [| A \<subseteq> carrier L |] ==> EX s. least L s (Upper L A)"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   978
    and inf_exists:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   979
    "!!A. [| A \<subseteq> carrier L |] ==> EX i. greatest L i (Lower L A)"
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   980
  shows "weak_complete_lattice L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   981
  by unfold_locales (auto intro: sup_exists inf_exists)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   982
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   983
constdefs (structure L)
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   984
  top :: "_ => 'a" ("\<top>\<index>")
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   985
  "\<top> == sup L (carrier L)"
21041
60e418260b4d Order and lattice structures no longer based on records.
ballarin
parents: 20318
diff changeset
   986
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   987
  bottom :: "_ => 'a" ("\<bottom>\<index>")
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   988
  "\<bottom> == inf L (carrier L)"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   989
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   990
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   991
lemma (in weak_complete_lattice) supI:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   992
  "[| !!l. least L l (Upper L A) ==> P l; A \<subseteq> carrier L |]
14651
02b8f3bcf7fe improved notation;
wenzelm
parents: 14577
diff changeset
   993
  ==> P (\<Squnion>A)"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
   994
proof (unfold sup_def)
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   995
  assume L: "A \<subseteq> carrier L"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   996
    and P: "!!l. least L l (Upper L A) ==> P l"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
   997
  with sup_exists obtain s where "least L s (Upper L A)" by blast
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   998
  with L show "P (SOME l. least L l (Upper L A))"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
   999
  by (fast intro: someI2 weak_least_unique P)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1000
qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1001
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1002
lemma (in weak_complete_lattice) sup_closed [simp]:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1003
  "A \<subseteq> carrier L ==> \<Squnion>A \<in> carrier L"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1004
  by (rule supI) simp_all
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1005
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1006
lemma (in weak_complete_lattice) top_closed [simp, intro]:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1007
  "\<top> \<in> carrier L"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1008
  by (unfold top_def) simp
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1009
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1010
lemma (in weak_complete_lattice) infI:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1011
  "[| !!i. greatest L i (Lower L A) ==> P i; A \<subseteq> carrier L |]
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
  1012
  ==> P (\<Sqinter>A)"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1013
proof (unfold inf_def)
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1014
  assume L: "A \<subseteq> carrier L"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1015
    and P: "!!l. greatest L l (Lower L A) ==> P l"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1016
  with inf_exists obtain s where "greatest L s (Lower L A)" by blast
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1017
  with L show "P (SOME l. greatest L l (Lower L A))"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1018
  by (fast intro: someI2 weak_greatest_unique P)
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1019
qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1020
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1021
lemma (in weak_complete_lattice) inf_closed [simp]:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1022
  "A \<subseteq> carrier L ==> \<Sqinter>A \<in> carrier L"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1023
  by (rule infI) simp_all
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1024
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1025
lemma (in weak_complete_lattice) bottom_closed [simp, intro]:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1026
  "\<bottom> \<in> carrier L"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1027
  by (unfold bottom_def) simp
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1028
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1029
text {* Jacobson: Theorem 8.1 *}
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1030
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1031
lemma Lower_empty [simp]:
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1032
  "Lower L {} = carrier L"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1033
  by (unfold Lower_def) simp
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1034
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1035
lemma Upper_empty [simp]:
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1036
  "Upper L {} = carrier L"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1037
  by (unfold Upper_def) simp
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1038
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1039
theorem (in weak_partial_order) weak_complete_lattice_criterion1:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1040
  assumes top_exists: "EX g. greatest L g (carrier L)"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1041
    and inf_exists:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1042
      "!!A. [| A \<subseteq> carrier L; A ~= {} |] ==> EX i. greatest L i (Lower L A)"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1043
  shows "weak_complete_lattice L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1044
proof (rule weak_complete_latticeI)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1045
  from top_exists obtain top where top: "greatest L top (carrier L)" ..
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1046
  fix A
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1047
  assume L: "A \<subseteq> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1048
  let ?B = "Upper L A"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1049
  from L top have "top \<in> ?B" by (fast intro!: Upper_memI intro: greatest_le)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1050
  then have B_non_empty: "?B ~= {}" by fast
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1051
  have B_L: "?B \<subseteq> carrier L" by simp
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1052
  from inf_exists [OF B_L B_non_empty]
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1053
  obtain b where b_inf_B: "greatest L b (Lower L ?B)" ..
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1054
  have "least L b (Upper L A)"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1055
apply (rule least_UpperI)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1056
   apply (rule greatest_le [where A = "Lower L ?B"])
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1057
    apply (rule b_inf_B)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1058
   apply (rule Lower_memI)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1059
    apply (erule Upper_memD [THEN conjunct1])
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1060
     apply assumption
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1061
    apply (rule L)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1062
   apply (fast intro: L [THEN subsetD])
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1063
  apply (erule greatest_Lower_below [OF b_inf_B])
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1064
  apply simp
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1065
 apply (rule L)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1066
apply (rule greatest_closed [OF b_inf_B])
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1067
done
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1068
  then show "EX s. least L s (Upper L A)" ..
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1069
next
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1070
  fix A
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1071
  assume L: "A \<subseteq> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1072
  show "EX i. greatest L i (Lower L A)"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1073
  proof (cases "A = {}")
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1074
    case True then show ?thesis
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1075
      by (simp add: top_exists)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1076
  next
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1077
    case False with L show ?thesis
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1078
      by (rule inf_exists)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1079
  qed
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1080
qed
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1081
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1082
(* TODO: prove dual version *)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1083
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1084
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1085
subsection {* Orders and Lattices where @{text eq} is the Equality *}
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1086
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1087
locale partial_order = weak_partial_order +
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1088
  assumes eq_is_equal: "op .= = op ="
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1089
begin
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1090
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1091
declare weak_le_anti_sym [rule del]
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1092
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1093
lemma le_anti_sym [intro]:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1094
  "[| x \<sqsubseteq> y; y \<sqsubseteq> x; x \<in> carrier L; y \<in> carrier L |] ==> x = y"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1095
  using weak_le_anti_sym unfolding eq_is_equal .
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1096
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1097
lemma lless_eq:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1098
  "x \<sqsubset> y \<longleftrightarrow> x \<sqsubseteq> y & x \<noteq> y"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1099
  unfolding lless_def by (simp add: eq_is_equal)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1100
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1101
lemma lless_asym:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1102
  assumes "a \<in> carrier L" "b \<in> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1103
    and "a \<sqsubset> b" "b \<sqsubset> a"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1104
  shows "P"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1105
  using assms unfolding lless_eq by auto
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1106
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1107
lemma lless_trans [trans]:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1108
  assumes "a \<sqsubset> b" "b \<sqsubset> c"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1109
    and carr[simp]: "a \<in> carrier L" "b \<in> carrier L" "c \<in> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1110
  shows "a \<sqsubset> c"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1111
  using assms unfolding lless_eq by (blast dest: le_trans intro: sym)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1112
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1113
end
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1114
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1115
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1116
subsubsection {* Upper and lower bounds of a set *}
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1117
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1118
(* all relevant lemmas are global and already proved above *)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1119
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1120
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1121
subsubsection {* Least and greatest, as predicate *}
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1122
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1123
lemma (in partial_order) least_unique:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1124
  "[| least L x A; least L y A |] ==> x = y"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1125
  using weak_least_unique unfolding eq_is_equal .
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1126
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1127
lemma (in partial_order) greatest_unique:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1128
  "[| greatest L x A; greatest L y A |] ==> x = y"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1129
  using weak_greatest_unique unfolding eq_is_equal .
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1130
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1131
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1132
subsection {* Lattices *}
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1133
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1134
locale upper_semilattice = partial_order +
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1135
  assumes sup_of_two_exists:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1136
    "[| x \<in> carrier L; y \<in> carrier L |] ==> EX s. least L s (Upper L {x, y})"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1137
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1138
interpretation upper_semilattice < weak_upper_semilattice
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1139
  by unfold_locales (rule sup_of_two_exists)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1140
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1141
locale lower_semilattice = partial_order +
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1142
  assumes inf_of_two_exists:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1143
    "[| x \<in> carrier L; y \<in> carrier L |] ==> EX s. greatest L s (Lower L {x, y})"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1144
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1145
interpretation lower_semilattice < weak_lower_semilattice
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1146
  by unfold_locales (rule inf_of_two_exists)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1147
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1148
locale lattice = upper_semilattice + lower_semilattice
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1149
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1150
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1151
subsubsection {* Supremum *}
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1152
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1153
context partial_order begin
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1154
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1155
declare weak_sup_of_singleton [simp del]
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1156
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1157
lemma sup_of_singleton [simp]:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1158
  "x \<in> carrier L ==> \<Squnion>{x} = x"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1159
  using weak_sup_of_singleton unfolding eq_is_equal .
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1160
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1161
end
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1162
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1163
context upper_semilattice begin
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1164
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1165
lemma join_assoc_lemma:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1166
  assumes L: "x \<in> carrier L"  "y \<in> carrier L"  "z \<in> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1167
  shows "x \<squnion> (y \<squnion> z) = \<Squnion>{x, y, z}"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1168
  using weak_join_assoc_lemma unfolding eq_is_equal .
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1169
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1170
end
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1171
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1172
lemma (in upper_semilattice) join_assoc:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1173
  assumes L: "x \<in> carrier L"  "y \<in> carrier L"  "z \<in> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1174
  shows "(x \<squnion> y) \<squnion> z = x \<squnion> (y \<squnion> z)"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1175
  using weak_join_assoc unfolding eq_is_equal .
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1176
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1177
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1178
subsubsection {* Infimum *}
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1179
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1180
context partial_order begin
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1181
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1182
declare weak_inf_of_singleton [simp del]
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1183
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1184
lemma inf_of_singleton [simp]:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1185
  "x \<in> carrier L ==> \<Sqinter>{x} = x"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1186
  using weak_inf_of_singleton unfolding eq_is_equal .
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1187
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1188
end
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1189
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1190
context lower_semilattice begin
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1191
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1192
text {* Condition on @{text A}: infimum exists. *}
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1193
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1194
lemma meet_assoc_lemma:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1195
  assumes L: "x \<in> carrier L"  "y \<in> carrier L"  "z \<in> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1196
  shows "x \<sqinter> (y \<sqinter> z) = \<Sqinter>{x, y, z}"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1197
  using weak_meet_assoc_lemma unfolding eq_is_equal .
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1198
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1199
end
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1200
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1201
lemma (in lower_semilattice) meet_assoc:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1202
  assumes L: "x \<in> carrier L"  "y \<in> carrier L"  "z \<in> carrier L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1203
  shows "(x \<sqinter> y) \<sqinter> z = x \<sqinter> (y \<sqinter> z)"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1204
  using weak_meet_assoc unfolding eq_is_equal .
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1205
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1206
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1207
subsection {* Total Orders *}
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1208
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1209
locale total_order = partial_order +
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1210
  assumes total: "[| x \<in> carrier L; y \<in> carrier L |] ==> x \<sqsubseteq> y | y \<sqsubseteq> x"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1211
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1212
interpretation total_order < weak_total_order
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1213
  by unfold_locales (rule total)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1214
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1215
text {* Introduction rule: the usual definition of total order *}
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1216
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1217
lemma (in partial_order) total_orderI:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1218
  assumes total: "!!x y. [| x \<in> carrier L; y \<in> carrier L |] ==> x \<sqsubseteq> y | y \<sqsubseteq> x"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1219
  shows "total_order L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1220
  by unfold_locales (rule total)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1221
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1222
text {* Total orders are lattices. *}
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1223
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1224
interpretation total_order < lattice
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1225
  by unfold_locales (auto intro: sup_of_two_exists inf_of_two_exists)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1226
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1227
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1228
subsection {* Complete lattices *}
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1229
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1230
locale complete_lattice = lattice +
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1231
  assumes sup_exists:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1232
    "[| A \<subseteq> carrier L |] ==> EX s. least L s (Upper L A)"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1233
    and inf_exists:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1234
    "[| A \<subseteq> carrier L |] ==> EX i. greatest L i (Lower L A)"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1235
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1236
interpretation complete_lattice < weak_complete_lattice
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1237
  by unfold_locales (auto intro: sup_exists inf_exists)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1238
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1239
text {* Introduction rule: the usual definition of complete lattice *}
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1240
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1241
lemma (in partial_order) complete_latticeI:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1242
  assumes sup_exists:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1243
    "!!A. [| A \<subseteq> carrier L |] ==> EX s. least L s (Upper L A)"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1244
    and inf_exists:
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1245
    "!!A. [| A \<subseteq> carrier L |] ==> EX i. greatest L i (Lower L A)"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1246
  shows "complete_lattice L"
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1247
  by unfold_locales (auto intro: sup_exists inf_exists)
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1248
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1249
theorem (in partial_order) complete_lattice_criterion1:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1250
  assumes top_exists: "EX g. greatest L g (carrier L)"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1251
    and inf_exists:
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1252
      "!!A. [| A \<subseteq> carrier L; A ~= {} |] ==> EX i. greatest L i (Lower L A)"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1253
  shows "complete_lattice L"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1254
proof (rule complete_latticeI)
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1255
  from top_exists obtain top where top: "greatest L top (carrier L)" ..
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1256
  fix A
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1257
  assume L: "A \<subseteq> carrier L"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1258
  let ?B = "Upper L A"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1259
  from L top have "top \<in> ?B" by (fast intro!: Upper_memI intro: greatest_le)
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1260
  then have B_non_empty: "?B ~= {}" by fast
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1261
  have B_L: "?B \<subseteq> carrier L" by simp
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1262
  from inf_exists [OF B_L B_non_empty]
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1263
  obtain b where b_inf_B: "greatest L b (Lower L ?B)" ..
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1264
  have "least L b (Upper L A)"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1265
apply (rule least_UpperI)
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1266
   apply (rule greatest_le [where A = "Lower L ?B"])
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1267
    apply (rule b_inf_B)
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1268
   apply (rule Lower_memI)
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1269
    apply (erule Upper_memD [THEN conjunct1])
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1270
     apply assumption
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1271
    apply (rule L)
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1272
   apply (fast intro: L [THEN subsetD])
27700
ef4b26efa8b6 Renamed theorems;
ballarin
parents: 26805
diff changeset
  1273
  apply (erule greatest_Lower_below [OF b_inf_B])
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1274
  apply simp
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1275
 apply (rule L)
27700
ef4b26efa8b6 Renamed theorems;
ballarin
parents: 26805
diff changeset
  1276
apply (rule greatest_closed [OF b_inf_B])
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1277
done
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1278
  then show "EX s. least L s (Upper L A)" ..
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1279
next
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1280
  fix A
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1281
  assume L: "A \<subseteq> carrier L"
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1282
  show "EX i. greatest L i (Lower L A)"
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1283
  proof (cases "A = {}")
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1284
    case True then show ?thesis
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1285
      by (simp add: top_exists)
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1286
  next
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1287
    case False with L show ?thesis
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1288
      by (rule inf_exists)
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1289
  qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1290
qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1291
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1292
(* TODO: prove dual version *)
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1293
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19984
diff changeset
  1294
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1295
subsection {* Examples *}
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1296
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 19984
diff changeset
  1297
subsubsection {* Powerset of a Set is a Complete Lattice *}
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1298
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1299
theorem powerset_is_complete_lattice:
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1300
  "complete_lattice (| carrier = Pow A, eq = op =, le = op \<subseteq> |)"
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1301
  (is "complete_lattice ?L")
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1302
proof (rule partial_order.complete_latticeI)
22063
717425609192 Reverted to structure representation with records.
ballarin
parents: 21896
diff changeset
  1303
  show "partial_order ?L"
27713
95b36bfe7fc4 New locales for orders and lattices where the equivalence relation is not restricted to equality.
ballarin
parents: 27700
diff changeset
  1304
    by unfold_locales auto
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1305
next
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1306
  fix B
26805
27941d7d9a11 Replaced forward proofs of existential statements by backward proofs
berghofe
parents: 24087
diff changeset
  1307
  assume B: "B \<subseteq> carrier ?L"
27941d7d9a11 Replaced forward proofs of existential statements by backward proofs
berghofe
parents: 24087
diff changeset
  1308
  show "EX s. least ?L s (Upper ?L B)"
27941d7d9a11 Replaced forward proofs of existential statements by backward proofs
berghofe
parents: 24087
diff changeset
  1309
  proof
27941d7d9a11 Replaced forward proofs of existential statements by backward proofs
berghofe
parents: 24087
diff changeset
  1310
    from B show "least ?L (\<Union> B) (Upper ?L B)"
27941d7d9a11 Replaced forward proofs of existential statements by backward proofs
berghofe
parents: 24087
diff changeset
  1311
      by (fastsimp intro!: least_UpperI simp: Upper_def)
27941d7d9a11 Replaced forward proofs of existential statements by backward proofs
berghofe
parents: 24087
diff changeset
  1312
  qed
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1313
next
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1314
  fix B
26805
27941d7d9a11 Replaced forward proofs of existential statements by backward proofs
berghofe
parents: 24087
diff changeset
  1315
  assume B: "B \<subseteq> carrier ?L"
27941d7d9a11 Replaced forward proofs of existential statements by backward proofs
berghofe
parents: 24087
diff changeset
  1316
  show "EX i. greatest ?L i (Lower ?L B)"
27941d7d9a11 Replaced forward proofs of existential statements by backward proofs
berghofe
parents: 24087
diff changeset
  1317
  proof
27941d7d9a11 Replaced forward proofs of existential statements by backward proofs
berghofe
parents: 24087
diff changeset
  1318
    from B show "greatest ?L (\<Inter> B \<inter> A) (Lower ?L B)"
27941d7d9a11 Replaced forward proofs of existential statements by backward proofs
berghofe
parents: 24087
diff changeset
  1319
      txt {* @{term "\<Inter> B"} is not the infimum of @{term B}:
27941d7d9a11 Replaced forward proofs of existential statements by backward proofs
berghofe
parents: 24087
diff changeset
  1320
	@{term "\<Inter> {} = UNIV"} which is in general bigger than @{term "A"}! *}
27941d7d9a11 Replaced forward proofs of existential statements by backward proofs
berghofe
parents: 24087
diff changeset
  1321
      by (fastsimp intro!: greatest_LowerI simp: Lower_def)
27941d7d9a11 Replaced forward proofs of existential statements by backward proofs
berghofe
parents: 24087
diff changeset
  1322
  qed
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1323
qed
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1324
14751
0d7850e27fed Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents: 14706
diff changeset
  1325
text {* An other example, that of the lattice of subgroups of a group,
0d7850e27fed Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents: 14706
diff changeset
  1326
  can be found in Group theory (Section~\ref{sec:subgroup-lattice}). *}
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents:
diff changeset
  1327
14693
4deda204e1d8 improved syntax;
wenzelm
parents: 14666
diff changeset
  1328
end