src/HOL/Algebra/Order.thy
author haftmann
Wed, 05 Apr 2017 13:47:38 +0200
changeset 65388 a8d868477bc0
parent 65099 30d0b2f1df76
child 66453 cc19f7ca2ed6
permissions -rw-r--r--
more on lists
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
     1
(*  Title:      HOL/Algebra/Order.thy
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
     2
    Author:     Clemens Ballarin, started 7 November 2003
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
     3
    Copyright:  Clemens Ballarin
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
     4
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
     5
Most congruence rules by Stephan Hohe.
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
     6
With additional contributions from Alasdair Armstrong and Simon Foster.
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
     7
*)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
     8
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
     9
theory Order
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    10
imports 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    11
  "~~/src/HOL/Library/FuncSet"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    12
  Congruence
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    13
begin
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    14
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    15
section \<open>Orders\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    16
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    17
subsection \<open>Partial Orders\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    18
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    19
record 'a gorder = "'a eq_object" +
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    20
  le :: "['a, 'a] => bool" (infixl "\<sqsubseteq>\<index>" 50)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    21
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    22
abbreviation inv_gorder :: "_ \<Rightarrow> 'a gorder" where
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    23
  "inv_gorder L \<equiv>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    24
   \<lparr> carrier = carrier L,
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    25
     eq = op .=\<^bsub>L\<^esub>,
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    26
     le = (\<lambda> x y. y \<sqsubseteq>\<^bsub>L \<^esub>x) \<rparr>"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    27
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    28
lemma inv_gorder_inv:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    29
  "inv_gorder (inv_gorder L) = L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    30
  by simp
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    31
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    32
locale weak_partial_order = equivalence L for L (structure) +
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    33
  assumes le_refl [intro, simp]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    34
      "x \<in> carrier L ==> x \<sqsubseteq> x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    35
    and weak_le_antisym [intro]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    36
      "[| x \<sqsubseteq> y; y \<sqsubseteq> x; x \<in> carrier L; y \<in> carrier L |] ==> x .= y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    37
    and le_trans [trans]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    38
      "[| x \<sqsubseteq> y; y \<sqsubseteq> z; x \<in> carrier L; y \<in> carrier L; z \<in> carrier L |] ==> x \<sqsubseteq> z"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    39
    and le_cong:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    40
      "\<lbrakk> x .= y; z .= w; x \<in> carrier L; y \<in> carrier L; z \<in> carrier L; w \<in> carrier L \<rbrakk> \<Longrightarrow>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    41
      x \<sqsubseteq> z \<longleftrightarrow> y \<sqsubseteq> w"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    42
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    43
definition
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    44
  lless :: "[_, 'a, 'a] => bool" (infixl "\<sqsubset>\<index>" 50)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    45
  where "x \<sqsubset>\<^bsub>L\<^esub> y \<longleftrightarrow> x \<sqsubseteq>\<^bsub>L\<^esub> y & x .\<noteq>\<^bsub>L\<^esub> y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    46
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    47
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    48
subsubsection \<open>The order relation\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    49
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    50
context weak_partial_order
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    51
begin
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    52
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    53
lemma le_cong_l [intro, trans]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    54
  "\<lbrakk> x .= y; y \<sqsubseteq> z; x \<in> carrier L; y \<in> carrier L; z \<in> carrier L \<rbrakk> \<Longrightarrow> x \<sqsubseteq> z"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    55
  by (auto intro: le_cong [THEN iffD2])
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    56
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    57
lemma le_cong_r [intro, trans]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    58
  "\<lbrakk> x \<sqsubseteq> y; y .= z; x \<in> carrier L; y \<in> carrier L; z \<in> carrier L \<rbrakk> \<Longrightarrow> x \<sqsubseteq> z"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    59
  by (auto intro: le_cong [THEN iffD1])
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    60
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    61
lemma weak_refl [intro, simp]: "\<lbrakk> x .= y; x \<in> carrier L; y \<in> carrier L \<rbrakk> \<Longrightarrow> x \<sqsubseteq> y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    62
  by (simp add: le_cong_l)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    63
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    64
end
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    65
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    66
lemma weak_llessI:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    67
  fixes R (structure)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    68
  assumes "x \<sqsubseteq> y" and "~(x .= y)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    69
  shows "x \<sqsubset> y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    70
  using assms unfolding lless_def by simp
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    71
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    72
lemma lless_imp_le:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    73
  fixes R (structure)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    74
  assumes "x \<sqsubset> y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    75
  shows "x \<sqsubseteq> y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    76
  using assms unfolding lless_def by simp
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    77
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    78
lemma weak_lless_imp_not_eq:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    79
  fixes R (structure)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    80
  assumes "x \<sqsubset> y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    81
  shows "\<not> (x .= y)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    82
  using assms unfolding lless_def by simp
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    83
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    84
lemma weak_llessE:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    85
  fixes R (structure)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    86
  assumes p: "x \<sqsubset> y" and e: "\<lbrakk>x \<sqsubseteq> y; \<not> (x .= y)\<rbrakk> \<Longrightarrow> P"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    87
  shows "P"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    88
  using p by (blast dest: lless_imp_le weak_lless_imp_not_eq e)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    89
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    90
lemma (in weak_partial_order) lless_cong_l [trans]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    91
  assumes xx': "x .= x'"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    92
    and xy: "x' \<sqsubset> y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    93
    and carr: "x \<in> carrier L" "x' \<in> carrier L" "y \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    94
  shows "x \<sqsubset> y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    95
  using assms unfolding lless_def by (auto intro: trans sym)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    96
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    97
lemma (in weak_partial_order) lless_cong_r [trans]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    98
  assumes xy: "x \<sqsubset> y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    99
    and  yy': "y .= y'"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   100
    and carr: "x \<in> carrier L" "y \<in> carrier L" "y' \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   101
  shows "x \<sqsubset> y'"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   102
  using assms unfolding lless_def by (auto intro: trans sym)  (*slow*)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   103
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   104
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   105
lemma (in weak_partial_order) lless_antisym:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   106
  assumes "a \<in> carrier L" "b \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   107
    and "a \<sqsubset> b" "b \<sqsubset> a"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   108
  shows "P"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   109
  using assms
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   110
  by (elim weak_llessE) auto
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   111
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   112
lemma (in weak_partial_order) lless_trans [trans]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   113
  assumes "a \<sqsubset> b" "b \<sqsubset> c"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   114
    and carr[simp]: "a \<in> carrier L" "b \<in> carrier L" "c \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   115
  shows "a \<sqsubset> c"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   116
  using assms unfolding lless_def by (blast dest: le_trans intro: sym)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   117
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   118
lemma weak_partial_order_subset:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   119
  assumes "weak_partial_order L" "A \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   120
  shows "weak_partial_order (L\<lparr> carrier := A \<rparr>)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   121
proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   122
  interpret L: weak_partial_order L
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   123
    by (simp add: assms)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   124
  interpret equivalence "(L\<lparr> carrier := A \<rparr>)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   125
    by (simp add: L.equivalence_axioms assms(2) equivalence_subset)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   126
  show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   127
    apply (unfold_locales, simp_all)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   128
    using assms(2) apply auto[1]
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   129
    using assms(2) apply auto[1]
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   130
    apply (meson L.le_trans assms(2) contra_subsetD)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   131
    apply (meson L.le_cong assms(2) subsetCE)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   132
  done
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   133
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   134
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   135
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   136
subsubsection \<open>Upper and lower bounds of a set\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   137
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   138
definition
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   139
  Upper :: "[_, 'a set] => 'a set"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   140
  where "Upper L A = {u. (ALL x. x \<in> A \<inter> carrier L --> x \<sqsubseteq>\<^bsub>L\<^esub> u)} \<inter> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   141
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   142
definition
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   143
  Lower :: "[_, 'a set] => 'a set"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   144
  where "Lower L A = {l. (ALL x. x \<in> A \<inter> carrier L --> l \<sqsubseteq>\<^bsub>L\<^esub> x)} \<inter> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   145
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   146
lemma Upper_closed [intro!, simp]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   147
  "Upper L A \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   148
  by (unfold Upper_def) clarify
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   149
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   150
lemma Upper_memD [dest]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   151
  fixes L (structure)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   152
  shows "[| u \<in> Upper L A; x \<in> A; A \<subseteq> carrier L |] ==> x \<sqsubseteq> u \<and> u \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   153
  by (unfold Upper_def) blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   154
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   155
lemma (in weak_partial_order) Upper_elemD [dest]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   156
  "[| u .\<in> Upper L A; u \<in> carrier L; x \<in> A; A \<subseteq> carrier L |] ==> x \<sqsubseteq> u"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   157
  unfolding Upper_def elem_def
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   158
  by (blast dest: sym)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   159
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   160
lemma Upper_memI:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   161
  fixes L (structure)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   162
  shows "[| !! y. y \<in> A ==> y \<sqsubseteq> x; x \<in> carrier L |] ==> x \<in> Upper L A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   163
  by (unfold Upper_def) blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   164
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   165
lemma (in weak_partial_order) Upper_elemI:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   166
  "[| !! y. y \<in> A ==> y \<sqsubseteq> x; x \<in> carrier L |] ==> x .\<in> Upper L A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   167
  unfolding Upper_def by blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   168
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   169
lemma Upper_antimono:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   170
  "A \<subseteq> B ==> Upper L B \<subseteq> Upper L A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   171
  by (unfold Upper_def) blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   172
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   173
lemma (in weak_partial_order) Upper_is_closed [simp]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   174
  "A \<subseteq> carrier L ==> is_closed (Upper L A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   175
  by (rule is_closedI) (blast intro: Upper_memI)+
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   176
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   177
lemma (in weak_partial_order) Upper_mem_cong:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   178
  assumes a'carr: "a' \<in> carrier L" and Acarr: "A \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   179
    and aa': "a .= a'"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   180
    and aelem: "a \<in> Upper L A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   181
  shows "a' \<in> Upper L A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   182
proof (rule Upper_memI[OF _ a'carr])
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   183
  fix y
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   184
  assume yA: "y \<in> A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   185
  hence "y \<sqsubseteq> a" by (intro Upper_memD[OF aelem, THEN conjunct1] Acarr)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   186
  also note aa'
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   187
  finally
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   188
      show "y \<sqsubseteq> a'"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   189
      by (simp add: a'carr subsetD[OF Acarr yA] subsetD[OF Upper_closed aelem])
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   190
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   191
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   192
lemma (in weak_partial_order) Upper_cong:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   193
  assumes Acarr: "A \<subseteq> carrier L" and A'carr: "A' \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   194
    and AA': "A {.=} A'"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   195
  shows "Upper L A = Upper L A'"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   196
unfolding Upper_def
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   197
apply rule
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   198
 apply (rule, clarsimp) defer 1
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   199
 apply (rule, clarsimp) defer 1
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   200
proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   201
  fix x a'
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   202
  assume carr: "x \<in> carrier L" "a' \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   203
    and a'A': "a' \<in> A'"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   204
  assume aLxCond[rule_format]: "\<forall>a. a \<in> A \<and> a \<in> carrier L \<longrightarrow> a \<sqsubseteq> x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   205
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   206
  from AA' and a'A' have "\<exists>a\<in>A. a' .= a" by (rule set_eqD2)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   207
  from this obtain a
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   208
      where aA: "a \<in> A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   209
      and a'a: "a' .= a"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   210
      by auto
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   211
  note [simp] = subsetD[OF Acarr aA] carr
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   212
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   213
  note a'a
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   214
  also have "a \<sqsubseteq> x" by (simp add: aLxCond aA)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   215
  finally show "a' \<sqsubseteq> x" by simp
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   216
next
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   217
  fix x a
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   218
  assume carr: "x \<in> carrier L" "a \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   219
    and aA: "a \<in> A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   220
  assume a'LxCond[rule_format]: "\<forall>a'. a' \<in> A' \<and> a' \<in> carrier L \<longrightarrow> a' \<sqsubseteq> x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   221
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   222
  from AA' and aA have "\<exists>a'\<in>A'. a .= a'" by (rule set_eqD1)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   223
  from this obtain a'
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   224
      where a'A': "a' \<in> A'"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   225
      and aa': "a .= a'"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   226
      by auto
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   227
  note [simp] = subsetD[OF A'carr a'A'] carr
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   228
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   229
  note aa'
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   230
  also have "a' \<sqsubseteq> x" by (simp add: a'LxCond a'A')
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   231
  finally show "a \<sqsubseteq> x" by simp
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   232
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   233
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   234
lemma Lower_closed [intro!, simp]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   235
  "Lower L A \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   236
  by (unfold Lower_def) clarify
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   237
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   238
lemma Lower_memD [dest]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   239
  fixes L (structure)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   240
  shows "[| l \<in> Lower L A; x \<in> A; A \<subseteq> carrier L |] ==> l \<sqsubseteq> x \<and> l \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   241
  by (unfold Lower_def) blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   242
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   243
lemma Lower_memI:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   244
  fixes L (structure)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   245
  shows "[| !! y. y \<in> A ==> x \<sqsubseteq> y; x \<in> carrier L |] ==> x \<in> Lower L A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   246
  by (unfold Lower_def) blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   247
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   248
lemma Lower_antimono:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   249
  "A \<subseteq> B ==> Lower L B \<subseteq> Lower L A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   250
  by (unfold Lower_def) blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   251
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   252
lemma (in weak_partial_order) Lower_is_closed [simp]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   253
  "A \<subseteq> carrier L \<Longrightarrow> is_closed (Lower L A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   254
  by (rule is_closedI) (blast intro: Lower_memI dest: sym)+
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   255
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   256
lemma (in weak_partial_order) Lower_mem_cong:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   257
  assumes a'carr: "a' \<in> carrier L" and Acarr: "A \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   258
    and aa': "a .= a'"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   259
    and aelem: "a \<in> Lower L A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   260
  shows "a' \<in> Lower L A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   261
using assms Lower_closed[of L A]
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   262
by (intro Lower_memI) (blast intro: le_cong_l[OF aa'[symmetric]])
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   263
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   264
lemma (in weak_partial_order) Lower_cong:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   265
  assumes Acarr: "A \<subseteq> carrier L" and A'carr: "A' \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   266
    and AA': "A {.=} A'"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   267
  shows "Lower L A = Lower L A'"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   268
unfolding Lower_def
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   269
apply rule
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   270
 apply clarsimp defer 1
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   271
 apply clarsimp defer 1
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   272
proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   273
  fix x a'
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   274
  assume carr: "x \<in> carrier L" "a' \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   275
    and a'A': "a' \<in> A'"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   276
  assume "\<forall>a. a \<in> A \<and> a \<in> carrier L \<longrightarrow> x \<sqsubseteq> a"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   277
  hence aLxCond: "\<And>a. \<lbrakk>a \<in> A; a \<in> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> a" by fast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   278
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   279
  from AA' and a'A' have "\<exists>a\<in>A. a' .= a" by (rule set_eqD2)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   280
  from this obtain a
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   281
      where aA: "a \<in> A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   282
      and a'a: "a' .= a"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   283
      by auto
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   284
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   285
  from aA and subsetD[OF Acarr aA]
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   286
      have "x \<sqsubseteq> a" by (rule aLxCond)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   287
  also note a'a[symmetric]
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   288
  finally
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   289
      show "x \<sqsubseteq> a'" by (simp add: carr subsetD[OF Acarr aA])
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   290
next
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   291
  fix x a
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   292
  assume carr: "x \<in> carrier L" "a \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   293
    and aA: "a \<in> A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   294
  assume "\<forall>a'. a' \<in> A' \<and> a' \<in> carrier L \<longrightarrow> x \<sqsubseteq> a'"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   295
  hence a'LxCond: "\<And>a'. \<lbrakk>a' \<in> A'; a' \<in> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> a'" by fast+
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   296
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   297
  from AA' and aA have "\<exists>a'\<in>A'. a .= a'" by (rule set_eqD1)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   298
  from this obtain a'
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   299
      where a'A': "a' \<in> A'"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   300
      and aa': "a .= a'"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   301
      by auto
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   302
  from a'A' and subsetD[OF A'carr a'A']
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   303
      have "x \<sqsubseteq> a'" by (rule a'LxCond)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   304
  also note aa'[symmetric]
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   305
  finally show "x \<sqsubseteq> a" by (simp add: carr subsetD[OF A'carr a'A'])
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   306
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   307
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   308
text \<open>Jacobson: Theorem 8.1\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   309
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   310
lemma Lower_empty [simp]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   311
  "Lower L {} = carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   312
  by (unfold Lower_def) simp
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   313
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   314
lemma Upper_empty [simp]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   315
  "Upper L {} = carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   316
  by (unfold Upper_def) simp
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   317
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   318
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   319
subsubsection \<open>Least and greatest, as predicate\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   320
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   321
definition
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   322
  least :: "[_, 'a, 'a set] => bool"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   323
  where "least L l A \<longleftrightarrow> A \<subseteq> carrier L & l \<in> A & (ALL x : A. l \<sqsubseteq>\<^bsub>L\<^esub> x)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   324
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   325
definition
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   326
  greatest :: "[_, 'a, 'a set] => bool"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   327
  where "greatest L g A \<longleftrightarrow> A \<subseteq> carrier L & g \<in> A & (ALL x : A. x \<sqsubseteq>\<^bsub>L\<^esub> g)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   328
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   329
text (in weak_partial_order) \<open>Could weaken these to @{term "l \<in> carrier L \<and> l
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   330
  .\<in> A"} and @{term "g \<in> carrier L \<and> g .\<in> A"}.\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   331
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   332
lemma least_closed [intro, simp]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   333
  "least L l A ==> l \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   334
  by (unfold least_def) fast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   335
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   336
lemma least_mem:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   337
  "least L l A ==> l \<in> A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   338
  by (unfold least_def) fast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   339
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   340
lemma (in weak_partial_order) weak_least_unique:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   341
  "[| least L x A; least L y A |] ==> x .= y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   342
  by (unfold least_def) blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   343
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   344
lemma least_le:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   345
  fixes L (structure)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   346
  shows "[| least L x A; a \<in> A |] ==> x \<sqsubseteq> a"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   347
  by (unfold least_def) fast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   348
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   349
lemma (in weak_partial_order) least_cong:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   350
  "[| x .= x'; x \<in> carrier L; x' \<in> carrier L; is_closed A |] ==> least L x A = least L x' A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   351
  by (unfold least_def) (auto dest: sym)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   352
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   353
abbreviation is_lub :: "[_, 'a, 'a set] => bool"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   354
where "is_lub L x A \<equiv> least L x (Upper L A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   355
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   356
text (in weak_partial_order) \<open>@{const least} is not congruent in the second parameter for
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   357
  @{term "A {.=} A'"}\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   358
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   359
lemma (in weak_partial_order) least_Upper_cong_l:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   360
  assumes "x .= x'"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   361
    and "x \<in> carrier L" "x' \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   362
    and "A \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   363
  shows "least L x (Upper L A) = least L x' (Upper L A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   364
  apply (rule least_cong) using assms by auto
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   365
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   366
lemma (in weak_partial_order) least_Upper_cong_r:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   367
  assumes Acarrs: "A \<subseteq> carrier L" "A' \<subseteq> carrier L" (* unneccessary with current Upper? *)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   368
    and AA': "A {.=} A'"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   369
  shows "least L x (Upper L A) = least L x (Upper L A')"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   370
apply (subgoal_tac "Upper L A = Upper L A'", simp)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   371
by (rule Upper_cong) fact+
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   372
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   373
lemma least_UpperI:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   374
  fixes L (structure)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   375
  assumes above: "!! x. x \<in> A ==> x \<sqsubseteq> s"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   376
    and below: "!! y. y \<in> Upper L A ==> s \<sqsubseteq> y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   377
    and L: "A \<subseteq> carrier L"  "s \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   378
  shows "least L s (Upper L A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   379
proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   380
  have "Upper L A \<subseteq> carrier L" by simp
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   381
  moreover from above L have "s \<in> Upper L A" by (simp add: Upper_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   382
  moreover from below have "ALL x : Upper L A. s \<sqsubseteq> x" by fast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   383
  ultimately show ?thesis by (simp add: least_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   384
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   385
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   386
lemma least_Upper_above:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   387
  fixes L (structure)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   388
  shows "[| least L s (Upper L A); x \<in> A; A \<subseteq> carrier L |] ==> x \<sqsubseteq> s"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   389
  by (unfold least_def) blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   390
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   391
lemma greatest_closed [intro, simp]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   392
  "greatest L l A ==> l \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   393
  by (unfold greatest_def) fast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   394
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   395
lemma greatest_mem:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   396
  "greatest L l A ==> l \<in> A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   397
  by (unfold greatest_def) fast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   398
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   399
lemma (in weak_partial_order) weak_greatest_unique:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   400
  "[| greatest L x A; greatest L y A |] ==> x .= y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   401
  by (unfold greatest_def) blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   402
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   403
lemma greatest_le:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   404
  fixes L (structure)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   405
  shows "[| greatest L x A; a \<in> A |] ==> a \<sqsubseteq> x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   406
  by (unfold greatest_def) fast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   407
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   408
lemma (in weak_partial_order) greatest_cong:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   409
  "[| x .= x'; x \<in> carrier L; x' \<in> carrier L; is_closed A |] ==>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   410
  greatest L x A = greatest L x' A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   411
  by (unfold greatest_def) (auto dest: sym)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   412
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   413
abbreviation is_glb :: "[_, 'a, 'a set] => bool"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   414
where "is_glb L x A \<equiv> greatest L x (Lower L A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   415
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   416
text (in weak_partial_order) \<open>@{const greatest} is not congruent in the second parameter for
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   417
  @{term "A {.=} A'"} \<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   418
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   419
lemma (in weak_partial_order) greatest_Lower_cong_l:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   420
  assumes "x .= x'"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   421
    and "x \<in> carrier L" "x' \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   422
    and "A \<subseteq> carrier L" (* unneccessary with current Lower *)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   423
  shows "greatest L x (Lower L A) = greatest L x' (Lower L A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   424
  apply (rule greatest_cong) using assms by auto
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   425
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   426
lemma (in weak_partial_order) greatest_Lower_cong_r:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   427
  assumes Acarrs: "A \<subseteq> carrier L" "A' \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   428
    and AA': "A {.=} A'"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   429
  shows "greatest L x (Lower L A) = greatest L x (Lower L A')"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   430
apply (subgoal_tac "Lower L A = Lower L A'", simp)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   431
by (rule Lower_cong) fact+
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   432
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   433
lemma greatest_LowerI:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   434
  fixes L (structure)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   435
  assumes below: "!! x. x \<in> A ==> i \<sqsubseteq> x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   436
    and above: "!! y. y \<in> Lower L A ==> y \<sqsubseteq> i"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   437
    and L: "A \<subseteq> carrier L"  "i \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   438
  shows "greatest L i (Lower L A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   439
proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   440
  have "Lower L A \<subseteq> carrier L" by simp
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   441
  moreover from below L have "i \<in> Lower L A" by (simp add: Lower_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   442
  moreover from above have "ALL x : Lower L A. x \<sqsubseteq> i" by fast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   443
  ultimately show ?thesis by (simp add: greatest_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   444
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   445
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   446
lemma greatest_Lower_below:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   447
  fixes L (structure)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   448
  shows "[| greatest L i (Lower L A); x \<in> A; A \<subseteq> carrier L |] ==> i \<sqsubseteq> x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   449
  by (unfold greatest_def) blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   450
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   451
lemma Lower_dual [simp]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   452
  "Lower (inv_gorder L) A = Upper L A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   453
  by (simp add:Upper_def Lower_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   454
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   455
lemma Upper_dual [simp]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   456
  "Upper (inv_gorder L) A = Lower L A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   457
  by (simp add:Upper_def Lower_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   458
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   459
lemma least_dual [simp]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   460
  "least (inv_gorder L) x A = greatest L x A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   461
  by (simp add:least_def greatest_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   462
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   463
lemma greatest_dual [simp]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   464
  "greatest (inv_gorder L) x A = least L x A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   465
  by (simp add:least_def greatest_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   466
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   467
lemma (in weak_partial_order) dual_weak_order:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   468
  "weak_partial_order (inv_gorder L)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   469
  apply (unfold_locales)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   470
  apply (simp_all)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   471
  apply (metis sym)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   472
  apply (metis trans)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   473
  apply (metis weak_le_antisym)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   474
  apply (metis le_trans)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   475
  apply (metis le_cong_l le_cong_r sym)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   476
done
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   477
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   478
lemma dual_weak_order_iff:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   479
  "weak_partial_order (inv_gorder A) \<longleftrightarrow> weak_partial_order A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   480
proof
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   481
  assume "weak_partial_order (inv_gorder A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   482
  then interpret dpo: weak_partial_order "inv_gorder A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   483
  rewrites "carrier (inv_gorder A) = carrier A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   484
  and   "le (inv_gorder A)      = (\<lambda> x y. le A y x)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   485
  and   "eq (inv_gorder A)      = eq A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   486
    by (simp_all)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   487
  show "weak_partial_order A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   488
    by (unfold_locales, auto intro: dpo.sym dpo.trans dpo.le_trans)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   489
next
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   490
  assume "weak_partial_order A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   491
  thus "weak_partial_order (inv_gorder A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   492
    by (metis weak_partial_order.dual_weak_order)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   493
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   494
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   495
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   496
subsubsection \<open>Intervals\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   497
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   498
definition
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   499
  at_least_at_most :: "('a, 'c) gorder_scheme \<Rightarrow> 'a => 'a => 'a set" ("(1\<lbrace>_.._\<rbrace>\<index>)")
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   500
  where "\<lbrace>l..u\<rbrace>\<^bsub>A\<^esub> = {x \<in> carrier A. l \<sqsubseteq>\<^bsub>A\<^esub> x \<and> x \<sqsubseteq>\<^bsub>A\<^esub> u}"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   501
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   502
context weak_partial_order
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   503
begin
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   504
  
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   505
  lemma at_least_at_most_upper [dest]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   506
    "x \<in> \<lbrace>a..b\<rbrace> \<Longrightarrow> x \<sqsubseteq> b"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   507
    by (simp add: at_least_at_most_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   508
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   509
  lemma at_least_at_most_lower [dest]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   510
    "x \<in> \<lbrace>a..b\<rbrace> \<Longrightarrow> a \<sqsubseteq> x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   511
    by (simp add: at_least_at_most_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   512
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   513
  lemma at_least_at_most_closed: "\<lbrace>a..b\<rbrace> \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   514
    by (auto simp add: at_least_at_most_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   515
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   516
  lemma at_least_at_most_member [intro]: 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   517
    "\<lbrakk> x \<in> carrier L; a \<sqsubseteq> x; x \<sqsubseteq> b \<rbrakk> \<Longrightarrow> x \<in> \<lbrace>a..b\<rbrace>"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   518
    by (simp add: at_least_at_most_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   519
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   520
end
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   521
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   522
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   523
subsubsection \<open>Isotone functions\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   524
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   525
definition isotone :: "('a, 'c) gorder_scheme \<Rightarrow> ('b, 'd) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   526
  where
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   527
  "isotone A B f \<equiv>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   528
   weak_partial_order A \<and> weak_partial_order B \<and>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   529
   (\<forall>x\<in>carrier A. \<forall>y\<in>carrier A. x \<sqsubseteq>\<^bsub>A\<^esub> y \<longrightarrow> f x \<sqsubseteq>\<^bsub>B\<^esub> f y)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   530
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   531
lemma isotoneI [intro?]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   532
  fixes f :: "'a \<Rightarrow> 'b"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   533
  assumes "weak_partial_order L1"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   534
          "weak_partial_order L2"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   535
          "(\<And>x y. \<lbrakk> x \<in> carrier L1; y \<in> carrier L1; x \<sqsubseteq>\<^bsub>L1\<^esub> y \<rbrakk> 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   536
                   \<Longrightarrow> f x \<sqsubseteq>\<^bsub>L2\<^esub> f y)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   537
  shows "isotone L1 L2 f"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   538
  using assms by (auto simp add:isotone_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   539
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   540
abbreviation Monotone :: "('a, 'b) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool" ("Mono\<index>")
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   541
  where "Monotone L f \<equiv> isotone L L f"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   542
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   543
lemma use_iso1:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   544
  "\<lbrakk>isotone A A f; x \<in> carrier A; y \<in> carrier A; x \<sqsubseteq>\<^bsub>A\<^esub> y\<rbrakk> \<Longrightarrow>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   545
   f x \<sqsubseteq>\<^bsub>A\<^esub> f y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   546
  by (simp add: isotone_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   547
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   548
lemma use_iso2:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   549
  "\<lbrakk>isotone A B f; x \<in> carrier A; y \<in> carrier A; x \<sqsubseteq>\<^bsub>A\<^esub> y\<rbrakk> \<Longrightarrow>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   550
   f x \<sqsubseteq>\<^bsub>B\<^esub> f y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   551
  by (simp add: isotone_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   552
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   553
lemma iso_compose:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   554
  "\<lbrakk>f \<in> carrier A \<rightarrow> carrier B; isotone A B f; g \<in> carrier B \<rightarrow> carrier C; isotone B C g\<rbrakk> \<Longrightarrow>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   555
   isotone A C (g \<circ> f)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   556
  by (simp add: isotone_def, safe, metis Pi_iff)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   557
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   558
lemma (in weak_partial_order) inv_isotone [simp]: 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   559
  "isotone (inv_gorder A) (inv_gorder B) f = isotone A B f"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   560
  by (auto simp add:isotone_def dual_weak_order dual_weak_order_iff)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   561
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   562
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   563
subsubsection \<open>Idempotent functions\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   564
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   565
definition idempotent :: 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   566
  "('a, 'b) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool" ("Idem\<index>") where
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   567
  "idempotent L f \<equiv> \<forall>x\<in>carrier L. f (f x) .=\<^bsub>L\<^esub> f x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   568
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   569
lemma (in weak_partial_order) idempotent:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   570
  "\<lbrakk> Idem f; x \<in> carrier L \<rbrakk> \<Longrightarrow> f (f x) .= f x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   571
  by (auto simp add: idempotent_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   572
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   573
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   574
subsubsection \<open>Order embeddings\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   575
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   576
definition order_emb :: "('a, 'c) gorder_scheme \<Rightarrow> ('b, 'd) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   577
  where
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   578
  "order_emb A B f \<equiv> weak_partial_order A 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   579
                   \<and> weak_partial_order B 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   580
                   \<and> (\<forall>x\<in>carrier A. \<forall>y\<in>carrier A. f x \<sqsubseteq>\<^bsub>B\<^esub> f y \<longleftrightarrow> x \<sqsubseteq>\<^bsub>A\<^esub> y )"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   581
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   582
lemma order_emb_isotone: "order_emb A B f \<Longrightarrow> isotone A B f"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   583
  by (auto simp add: isotone_def order_emb_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   584
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   585
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   586
subsubsection \<open>Commuting functions\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   587
    
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   588
definition commuting :: "('a, 'c) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool" where
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   589
"commuting A f g = (\<forall>x\<in>carrier A. (f \<circ> g) x .=\<^bsub>A\<^esub> (g \<circ> f) x)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   590
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   591
subsection \<open>Partial orders where \<open>eq\<close> is the Equality\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   592
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   593
locale partial_order = weak_partial_order +
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   594
  assumes eq_is_equal: "op .= = op ="
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   595
begin
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   596
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   597
declare weak_le_antisym [rule del]
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   598
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   599
lemma le_antisym [intro]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   600
  "[| x \<sqsubseteq> y; y \<sqsubseteq> x; x \<in> carrier L; y \<in> carrier L |] ==> x = y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   601
  using weak_le_antisym unfolding eq_is_equal .
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   602
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   603
lemma lless_eq:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   604
  "x \<sqsubset> y \<longleftrightarrow> x \<sqsubseteq> y & x \<noteq> y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   605
  unfolding lless_def by (simp add: eq_is_equal)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   606
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   607
lemma set_eq_is_eq: "A {.=} B \<longleftrightarrow> A = B"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   608
  by (auto simp add: set_eq_def elem_def eq_is_equal)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   609
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   610
end
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   611
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   612
lemma (in partial_order) dual_order:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   613
  "partial_order (inv_gorder L)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   614
proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   615
  interpret dwo: weak_partial_order "inv_gorder L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   616
    by (metis dual_weak_order)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   617
  show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   618
    by (unfold_locales, simp add:eq_is_equal)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   619
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   620
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   621
lemma dual_order_iff:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   622
  "partial_order (inv_gorder A) \<longleftrightarrow> partial_order A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   623
proof
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   624
  assume assm:"partial_order (inv_gorder A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   625
  then interpret po: partial_order "inv_gorder A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   626
  rewrites "carrier (inv_gorder A) = carrier A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   627
  and   "le (inv_gorder A)      = (\<lambda> x y. le A y x)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   628
  and   "eq (inv_gorder A)      = eq A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   629
    by (simp_all)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   630
  show "partial_order A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   631
    apply (unfold_locales, simp_all)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   632
    apply (metis po.sym, metis po.trans)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   633
    apply (metis po.weak_le_antisym, metis po.le_trans)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   634
    apply (metis (full_types) po.eq_is_equal, metis po.eq_is_equal)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   635
  done
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   636
next
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   637
  assume "partial_order A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   638
  thus "partial_order (inv_gorder A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   639
    by (metis partial_order.dual_order)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   640
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   641
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   642
text \<open>Least and greatest, as predicate\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   643
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   644
lemma (in partial_order) least_unique:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   645
  "[| least L x A; least L y A |] ==> x = y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   646
  using weak_least_unique unfolding eq_is_equal .
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   647
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   648
lemma (in partial_order) greatest_unique:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   649
  "[| greatest L x A; greatest L y A |] ==> x = y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   650
  using weak_greatest_unique unfolding eq_is_equal .
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   651
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   652
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   653
subsection \<open>Bounded Orders\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   654
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   655
definition
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   656
  top :: "_ => 'a" ("\<top>\<index>") where
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   657
  "\<top>\<^bsub>L\<^esub> = (SOME x. greatest L x (carrier L))"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   658
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   659
definition
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   660
  bottom :: "_ => 'a" ("\<bottom>\<index>") where
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   661
  "\<bottom>\<^bsub>L\<^esub> = (SOME x. least L x (carrier L))"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   662
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   663
locale weak_partial_order_bottom = weak_partial_order L for L (structure) +
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   664
  assumes bottom_exists: "\<exists> x. least L x (carrier L)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   665
begin
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   666
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   667
lemma bottom_least: "least L \<bottom> (carrier L)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   668
proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   669
  obtain x where "least L x (carrier L)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   670
    by (metis bottom_exists)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   671
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   672
  thus ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   673
    by (auto intro:someI2 simp add: bottom_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   674
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   675
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   676
lemma bottom_closed [simp, intro]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   677
  "\<bottom> \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   678
  by (metis bottom_least least_mem)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   679
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   680
lemma bottom_lower [simp, intro]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   681
  "x \<in> carrier L \<Longrightarrow> \<bottom> \<sqsubseteq> x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   682
  by (metis bottom_least least_le)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   683
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   684
end
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   685
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   686
locale weak_partial_order_top = weak_partial_order L for L (structure) +
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   687
  assumes top_exists: "\<exists> x. greatest L x (carrier L)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   688
begin
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   689
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   690
lemma top_greatest: "greatest L \<top> (carrier L)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   691
proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   692
  obtain x where "greatest L x (carrier L)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   693
    by (metis top_exists)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   694
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   695
  thus ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   696
    by (auto intro:someI2 simp add: top_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   697
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   698
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   699
lemma top_closed [simp, intro]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   700
  "\<top> \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   701
  by (metis greatest_mem top_greatest)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   702
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   703
lemma top_higher [simp, intro]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   704
  "x \<in> carrier L \<Longrightarrow> x \<sqsubseteq> \<top>"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   705
  by (metis greatest_le top_greatest)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   706
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   707
end
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   708
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   709
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   710
subsection \<open>Total Orders\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   711
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   712
locale weak_total_order = weak_partial_order +
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   713
  assumes total: "\<lbrakk> x \<in> carrier L; y \<in> carrier L \<rbrakk> \<Longrightarrow> x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   714
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   715
text \<open>Introduction rule: the usual definition of total order\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   716
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   717
lemma (in weak_partial_order) weak_total_orderI:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   718
  assumes total: "!!x y. \<lbrakk> x \<in> carrier L; y \<in> carrier L \<rbrakk> \<Longrightarrow> x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   719
  shows "weak_total_order L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   720
  by unfold_locales (rule total)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   721
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   722
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   723
subsection \<open>Total orders where \<open>eq\<close> is the Equality\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   724
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   725
locale total_order = partial_order +
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   726
  assumes total_order_total: "\<lbrakk> x \<in> carrier L; y \<in> carrier L \<rbrakk> \<Longrightarrow> x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   727
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   728
sublocale total_order < weak?: weak_total_order
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   729
  by unfold_locales (rule total_order_total)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   730
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   731
text \<open>Introduction rule: the usual definition of total order\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   732
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   733
lemma (in partial_order) total_orderI:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   734
  assumes total: "!!x y. \<lbrakk> x \<in> carrier L; y \<in> carrier L \<rbrakk> \<Longrightarrow> x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   735
  shows "total_order L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   736
  by unfold_locales (rule total)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   737
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   738
end