src/HOL/Algebra/Complete_Lattice.thy
author wenzelm
Thu, 05 Jan 2023 21:33:49 +0100
changeset 76923 8a66a88cd5dc
parent 69712 dc85b5b3a532
permissions -rw-r--r--
isabelle update -u path_cartouches;
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/Complete_Lattice.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 Complete_Lattice
66579
2db3fe23fdaf Revert 5a42eddc11c1.
ballarin
parents: 66501
diff changeset
    10
imports Lattice
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    11
begin
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    12
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    13
section \<open>Complete Lattices\<close>
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
locale weak_complete_lattice = weak_partial_order +
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    16
  assumes sup_exists:
67091
1393c2340eec more symbols;
wenzelm
parents: 66580
diff changeset
    17
    "[| A \<subseteq> carrier L |] ==> \<exists>s. least L s (Upper L A)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    18
    and inf_exists:
67091
1393c2340eec more symbols;
wenzelm
parents: 66580
diff changeset
    19
    "[| A \<subseteq> carrier L |] ==> \<exists>i. greatest L i (Lower L A)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    20
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    21
sublocale weak_complete_lattice \<subseteq> weak_lattice
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    22
proof
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    23
  fix x y
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    24
  assume a: "x \<in> carrier L" "y \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    25
  thus "\<exists>s. is_lub L s {x, y}"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    26
    by (rule_tac sup_exists[of "{x, y}"], auto)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    27
  from a show "\<exists>s. is_glb L s {x, y}"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    28
    by (rule_tac inf_exists[of "{x, y}"], auto)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    29
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    30
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    31
text \<open>Introduction rule: the usual definition of complete lattice\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    32
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    33
lemma (in weak_partial_order) weak_complete_latticeI:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    34
  assumes sup_exists:
67091
1393c2340eec more symbols;
wenzelm
parents: 66580
diff changeset
    35
    "!!A. [| A \<subseteq> carrier L |] ==> \<exists>s. least L s (Upper L A)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    36
    and inf_exists:
67091
1393c2340eec more symbols;
wenzelm
parents: 66580
diff changeset
    37
    "!!A. [| A \<subseteq> carrier L |] ==> \<exists>i. greatest L i (Lower L A)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    38
  shows "weak_complete_lattice L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    39
  by standard (auto intro: sup_exists inf_exists)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    40
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    41
lemma (in weak_complete_lattice) dual_weak_complete_lattice:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    42
  "weak_complete_lattice (inv_gorder L)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    43
proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    44
  interpret dual: weak_lattice "inv_gorder L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    45
    by (metis dual_weak_lattice)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    46
  show ?thesis
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
    47
    by (unfold_locales) (simp_all add:inf_exists sup_exists)
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    48
qed
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
lemma (in weak_complete_lattice) supI:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    51
  "[| !!l. least L l (Upper L A) ==> P l; A \<subseteq> carrier L |]
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    52
  ==> P (\<Squnion>A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    53
proof (unfold sup_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    54
  assume L: "A \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    55
    and P: "!!l. least L l (Upper L A) ==> P l"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    56
  with sup_exists obtain s where "least L s (Upper L A)" by blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    57
  with L show "P (SOME l. least L l (Upper L A))"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    58
  by (fast intro: someI2 weak_least_unique P)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    59
qed
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 (in weak_complete_lattice) sup_closed [simp]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    62
  "A \<subseteq> carrier L ==> \<Squnion>A \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    63
  by (rule supI) simp_all
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    64
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    65
lemma (in weak_complete_lattice) sup_cong:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    66
  assumes "A \<subseteq> carrier L" "B \<subseteq> carrier L" "A {.=} B"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    67
  shows "\<Squnion> A .= \<Squnion> B"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    68
proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    69
  have "\<And> x. is_lub L x A \<longleftrightarrow> is_lub L x B"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    70
    by (rule least_Upper_cong_r, simp_all add: assms)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    71
  moreover have "\<Squnion> B \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    72
    by (simp add: assms(2))
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    73
  ultimately show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    74
    by (simp add: sup_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    75
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    76
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    77
sublocale weak_complete_lattice \<subseteq> weak_bounded_lattice
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    78
  apply (unfold_locales)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    79
  apply (metis Upper_empty empty_subsetI sup_exists)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    80
  apply (metis Lower_empty empty_subsetI inf_exists)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    81
done
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    82
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    83
lemma (in weak_complete_lattice) infI:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    84
  "[| !!i. greatest L i (Lower L A) ==> P i; A \<subseteq> carrier L |]
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    85
  ==> P (\<Sqinter>A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    86
proof (unfold inf_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    87
  assume L: "A \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    88
    and P: "!!l. greatest L l (Lower L A) ==> P l"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    89
  with inf_exists obtain s where "greatest L s (Lower L A)" by blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    90
  with L show "P (SOME l. greatest L l (Lower L A))"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    91
  by (fast intro: someI2 weak_greatest_unique P)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    92
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    93
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    94
lemma (in weak_complete_lattice) inf_closed [simp]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    95
  "A \<subseteq> carrier L ==> \<Sqinter>A \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    96
  by (rule infI) simp_all
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    97
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    98
lemma (in weak_complete_lattice) inf_cong:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
    99
  assumes "A \<subseteq> carrier L" "B \<subseteq> carrier L" "A {.=} B"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   100
  shows "\<Sqinter> A .= \<Sqinter> B"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   101
proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   102
  have "\<And> x. is_glb L x A \<longleftrightarrow> is_glb L x B"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   103
    by (rule greatest_Lower_cong_r, simp_all add: assms)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   104
  moreover have "\<Sqinter> B \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   105
    by (simp add: assms(2))
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   106
  ultimately show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   107
    by (simp add: inf_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   108
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   109
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   110
theorem (in weak_partial_order) weak_complete_lattice_criterion1:
67091
1393c2340eec more symbols;
wenzelm
parents: 66580
diff changeset
   111
  assumes top_exists: "\<exists>g. greatest L g (carrier L)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   112
    and inf_exists:
67091
1393c2340eec more symbols;
wenzelm
parents: 66580
diff changeset
   113
      "\<And>A. [| A \<subseteq> carrier L; A \<noteq> {} |] ==> \<exists>i. greatest L i (Lower L A)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   114
  shows "weak_complete_lattice L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   115
proof (rule weak_complete_latticeI)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   116
  from top_exists obtain top where top: "greatest L top (carrier L)" ..
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   117
  fix A
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   118
  assume L: "A \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   119
  let ?B = "Upper L A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   120
  from L top have "top \<in> ?B" by (fast intro!: Upper_memI intro: greatest_le)
67091
1393c2340eec more symbols;
wenzelm
parents: 66580
diff changeset
   121
  then have B_non_empty: "?B \<noteq> {}" by fast
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   122
  have B_L: "?B \<subseteq> carrier L" by simp
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   123
  from inf_exists [OF B_L B_non_empty]
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   124
  obtain b where b_inf_B: "greatest L b (Lower L ?B)" ..
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   125
  then have bcarr: "b \<in> carrier L"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   126
    by auto
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   127
  have "least L b (Upper L A)"
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   128
  proof (rule least_UpperI)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   129
    show "\<And>x. x \<in> A \<Longrightarrow> x \<sqsubseteq> b"
69712
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   130
      by (meson L Lower_memI Upper_memD b_inf_B greatest_le subsetD)
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   131
    show "\<And>y. y \<in> Upper L A \<Longrightarrow> b \<sqsubseteq> y"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   132
      by (meson B_L b_inf_B greatest_Lower_below)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   133
  qed (use bcarr L in auto)
67091
1393c2340eec more symbols;
wenzelm
parents: 66580
diff changeset
   134
  then show "\<exists>s. least L s (Upper L A)" ..
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   135
next
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   136
  fix A
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   137
  assume L: "A \<subseteq> carrier L"
67091
1393c2340eec more symbols;
wenzelm
parents: 66580
diff changeset
   138
  show "\<exists>i. greatest L i (Lower L A)"
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   139
    by (metis L Lower_empty inf_exists top_exists)
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   140
qed
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
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   143
text \<open>Supremum\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   144
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   145
declare (in partial_order) weak_sup_of_singleton [simp del]
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   146
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   147
lemma (in partial_order) sup_of_singleton [simp]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   148
  "x \<in> carrier L ==> \<Squnion>{x} = x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   149
  using weak_sup_of_singleton unfolding eq_is_equal .
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   150
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   151
lemma (in upper_semilattice) join_assoc_lemma:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   152
  assumes L: "x \<in> carrier L"  "y \<in> carrier L"  "z \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   153
  shows "x \<squnion> (y \<squnion> z) = \<Squnion>{x, y, z}"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   154
  using weak_join_assoc_lemma L unfolding eq_is_equal .
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   155
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   156
lemma (in upper_semilattice) join_assoc:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   157
  assumes L: "x \<in> carrier L"  "y \<in> carrier L"  "z \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   158
  shows "(x \<squnion> y) \<squnion> z = x \<squnion> (y \<squnion> z)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   159
  using weak_join_assoc L unfolding eq_is_equal .
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   160
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   161
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   162
text \<open>Infimum\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   163
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   164
declare (in partial_order) weak_inf_of_singleton [simp del]
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   165
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   166
lemma (in partial_order) inf_of_singleton [simp]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   167
  "x \<in> carrier L ==> \<Sqinter>{x} = x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   168
  using weak_inf_of_singleton unfolding eq_is_equal .
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   169
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   170
text \<open>Condition on \<open>A\<close>: infimum exists.\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   171
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   172
lemma (in lower_semilattice) meet_assoc_lemma:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   173
  assumes L: "x \<in> carrier L"  "y \<in> carrier L"  "z \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   174
  shows "x \<sqinter> (y \<sqinter> z) = \<Sqinter>{x, y, z}"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   175
  using weak_meet_assoc_lemma L unfolding eq_is_equal .
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 lower_semilattice) meet_assoc:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   178
  assumes L: "x \<in> carrier L"  "y \<in> carrier L"  "z \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   179
  shows "(x \<sqinter> y) \<sqinter> z = x \<sqinter> (y \<sqinter> z)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   180
  using weak_meet_assoc L unfolding eq_is_equal .
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   181
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   182
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   183
subsection \<open>Infimum Laws\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   184
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   185
context weak_complete_lattice
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   186
begin
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   187
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   188
lemma inf_glb: 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   189
  assumes "A \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   190
  shows "greatest L (\<Sqinter>A) (Lower L A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   191
proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   192
  obtain i where "greatest L i (Lower L A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   193
    by (metis assms inf_exists)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   194
  thus ?thesis
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   195
    by (metis inf_def someI_ex)
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   196
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   197
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   198
lemma inf_lower:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   199
  assumes "A \<subseteq> carrier L" "x \<in> A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   200
  shows "\<Sqinter>A \<sqsubseteq> x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   201
  by (metis assms greatest_Lower_below inf_glb)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   202
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   203
lemma inf_greatest: 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   204
  assumes "A \<subseteq> carrier L" "z \<in> carrier L" 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   205
          "(\<And>x. x \<in> A \<Longrightarrow> z \<sqsubseteq> x)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   206
  shows "z \<sqsubseteq> \<Sqinter>A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   207
  by (metis Lower_memI assms greatest_le inf_glb)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   208
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   209
lemma weak_inf_empty [simp]: "\<Sqinter>{} .= \<top>"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   210
  by (metis Lower_empty empty_subsetI inf_glb top_greatest weak_greatest_unique)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   211
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   212
lemma weak_inf_carrier [simp]: "\<Sqinter>carrier L .= \<bottom>"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   213
  by (metis bottom_weak_eq inf_closed inf_lower subset_refl)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   214
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   215
lemma weak_inf_insert [simp]: 
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   216
  assumes "a \<in> carrier L" "A \<subseteq> carrier L"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   217
  shows "\<Sqinter>insert a A .= a \<sqinter> \<Sqinter>A"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   218
proof (rule weak_le_antisym)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   219
  show "\<Sqinter>insert a A \<sqsubseteq> a \<sqinter> \<Sqinter>A"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   220
    by (simp add: assms inf_lower local.inf_greatest meet_le)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   221
  show aA: "a \<sqinter> \<Sqinter>A \<in> carrier L"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   222
    using assms by simp
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   223
  show "a \<sqinter> \<Sqinter>A \<sqsubseteq> \<Sqinter>insert a A"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   224
    apply (rule inf_greatest)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   225
    using assms apply (simp_all add: aA)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   226
    by (meson aA inf_closed inf_lower local.le_trans meet_left meet_right subsetCE)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   227
  show "\<Sqinter>insert a A \<in> carrier L"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   228
    using assms by (force intro: le_trans inf_closed meet_right meet_left inf_lower)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   229
qed
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   230
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   231
subsection \<open>Supremum Laws\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   232
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   233
lemma sup_lub: 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   234
  assumes "A \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   235
  shows "least L (\<Squnion>A) (Upper L A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   236
    by (metis Upper_is_closed assms least_closed least_cong supI sup_closed sup_exists weak_least_unique)
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 sup_upper: 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   239
  assumes "A \<subseteq> carrier L" "x \<in> A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   240
  shows "x \<sqsubseteq> \<Squnion>A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   241
  by (metis assms least_Upper_above supI)
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 sup_least:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   244
  assumes "A \<subseteq> carrier L" "z \<in> carrier L" 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   245
          "(\<And>x. x \<in> A \<Longrightarrow> x \<sqsubseteq> z)" 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   246
  shows "\<Squnion>A \<sqsubseteq> z"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   247
  by (metis Upper_memI assms least_le sup_lub)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   248
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   249
lemma weak_sup_empty [simp]: "\<Squnion>{} .= \<bottom>"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   250
  by (metis Upper_empty bottom_least empty_subsetI sup_lub weak_least_unique)
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 weak_sup_carrier [simp]: "\<Squnion>carrier L .= \<top>"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   253
  by (metis Lower_closed Lower_empty sup_closed sup_upper top_closed top_higher weak_le_antisym)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   254
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   255
lemma weak_sup_insert [simp]: 
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   256
  assumes "a \<in> carrier L" "A \<subseteq> carrier L"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   257
  shows "\<Squnion>insert a A .= a \<squnion> \<Squnion>A"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   258
proof (rule weak_le_antisym)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   259
  show aA: "a \<squnion> \<Squnion>A \<in> carrier L"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   260
    using assms by simp
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   261
  show "\<Squnion>insert a A \<sqsubseteq> a \<squnion> \<Squnion>A"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   262
    apply (rule sup_least)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   263
    using assms apply (simp_all add: aA)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   264
    by (meson aA join_left join_right local.le_trans subsetCE sup_closed sup_upper)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   265
  show "a \<squnion> \<Squnion>A \<sqsubseteq> \<Squnion>insert a A"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   266
    by (simp add: assms join_le local.sup_least sup_upper)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   267
  show "\<Squnion>insert a A \<in> carrier L"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   268
    using assms by (force intro: le_trans inf_closed meet_right meet_left inf_lower)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   269
qed
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   270
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   271
end
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   272
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   273
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   274
subsection \<open>Fixed points of a lattice\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   275
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   276
definition "fps L f = {x \<in> carrier L. f x .=\<^bsub>L\<^esub> x}"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   277
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   278
abbreviation "fpl L f \<equiv> L\<lparr>carrier := fps L f\<rparr>"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   279
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   280
lemma (in weak_partial_order) 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   281
  use_fps: "x \<in> fps L f \<Longrightarrow> f x .= x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   282
  by (simp add: fps_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   283
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   284
lemma fps_carrier [simp]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   285
  "fps L f \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   286
  by (auto simp add: fps_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   287
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   288
lemma (in weak_complete_lattice) fps_sup_image: 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   289
  assumes "f \<in> carrier L \<rightarrow> carrier L" "A \<subseteq> fps L f" 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   290
  shows "\<Squnion> (f ` A) .= \<Squnion> A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   291
proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   292
  from assms(2) have AL: "A \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   293
    by (auto simp add: fps_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   294
  show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   295
  proof (rule sup_cong, simp_all add: AL)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   296
    from assms(1) AL show "f ` A \<subseteq> carrier L"
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   297
      by auto
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   298
    then have *: "\<And>b. \<lbrakk>A \<subseteq> {x \<in> carrier L. f x .= x}; b \<in> A\<rbrakk> \<Longrightarrow> \<exists>a\<in>f ` A. b .= a"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   299
      by (meson AL assms(2) image_eqI local.sym subsetCE use_fps)
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   300
    from assms(2) show "f ` A {.=} A"
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   301
      by (auto simp add: fps_def intro: set_eqI2 [OF _ *])
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   302
  qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   303
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   304
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   305
lemma (in weak_complete_lattice) fps_idem:
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   306
  assumes "f \<in> carrier L \<rightarrow> carrier L" "Idem f"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   307
  shows "fps L f {.=} f ` carrier L"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   308
proof (rule set_eqI2)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   309
  show "\<And>a. a \<in> fps L f \<Longrightarrow> \<exists>b\<in>f ` carrier L. a .= b"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   310
    using assms by (force simp add: fps_def intro: local.sym)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   311
  show "\<And>b. b \<in> f ` carrier L \<Longrightarrow> \<exists>a\<in>fps L f. b .= a"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   312
    using assms by (force simp add: idempotent_def fps_def)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   313
qed
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   314
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   315
context weak_complete_lattice
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   316
begin
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
lemma weak_sup_pre_fixed_point: 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   319
  assumes "f \<in> carrier L \<rightarrow> carrier L" "isotone L L f" "A \<subseteq> fps L f"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   320
  shows "(\<Squnion>\<^bsub>L\<^esub> A) \<sqsubseteq>\<^bsub>L\<^esub> f (\<Squnion>\<^bsub>L\<^esub> A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   321
proof (rule sup_least)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   322
  from assms(3) show AL: "A \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   323
    by (auto simp add: fps_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   324
  thus fA: "f (\<Squnion>A) \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   325
    by (simp add: assms funcset_carrier[of f L L])
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   326
  fix x
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   327
  assume xA: "x \<in> A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   328
  hence "x \<in> fps L f"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   329
    using assms subsetCE by blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   330
  hence "f x .=\<^bsub>L\<^esub> x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   331
    by (auto simp add: fps_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   332
  moreover have "f x \<sqsubseteq>\<^bsub>L\<^esub> f (\<Squnion>\<^bsub>L\<^esub>A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   333
    by (meson AL assms(2) subsetCE sup_closed sup_upper use_iso1 xA)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   334
  ultimately show "x \<sqsubseteq>\<^bsub>L\<^esub> f (\<Squnion>\<^bsub>L\<^esub>A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   335
    by (meson AL fA assms(1) funcset_carrier le_cong local.refl subsetCE xA)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   336
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   337
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   338
lemma weak_sup_post_fixed_point: 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   339
  assumes "f \<in> carrier L \<rightarrow> carrier L" "isotone L L f" "A \<subseteq> fps L f"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   340
  shows "f (\<Sqinter>\<^bsub>L\<^esub> A) \<sqsubseteq>\<^bsub>L\<^esub> (\<Sqinter>\<^bsub>L\<^esub> A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   341
proof (rule inf_greatest)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   342
  from assms(3) show AL: "A \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   343
    by (auto simp add: fps_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   344
  thus fA: "f (\<Sqinter>A) \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   345
    by (simp add: assms funcset_carrier[of f L L])
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   346
  fix x
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   347
  assume xA: "x \<in> A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   348
  hence "x \<in> fps L f"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   349
    using assms subsetCE by blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   350
  hence "f x .=\<^bsub>L\<^esub> x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   351
    by (auto simp add: fps_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   352
  moreover have "f (\<Sqinter>\<^bsub>L\<^esub>A) \<sqsubseteq>\<^bsub>L\<^esub> f x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   353
    by (meson AL assms(2) inf_closed inf_lower subsetCE use_iso1 xA)   
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   354
  ultimately show "f (\<Sqinter>\<^bsub>L\<^esub>A) \<sqsubseteq>\<^bsub>L\<^esub> x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   355
    by (meson AL assms(1) fA funcset_carrier le_cong_r subsetCE xA)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   356
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   357
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
subsubsection \<open>Least fixed points\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   360
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   361
lemma LFP_closed [intro, simp]:
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   362
  "LFP f \<in> carrier L"
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   363
  by (metis (lifting) LEAST_FP_def inf_closed mem_Collect_eq subsetI)
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   364
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   365
lemma LFP_lowerbound: 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   366
  assumes "x \<in> carrier L" "f x \<sqsubseteq> x" 
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   367
  shows "LFP f \<sqsubseteq> x"
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   368
  by (auto intro:inf_lower assms simp add:LEAST_FP_def)
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   369
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   370
lemma LFP_greatest: 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   371
  assumes "x \<in> carrier L" 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   372
          "(\<And>u. \<lbrakk> u \<in> carrier L; f u \<sqsubseteq> u \<rbrakk> \<Longrightarrow> x \<sqsubseteq> u)"
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   373
  shows "x \<sqsubseteq> LFP f"
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   374
  by (auto simp add:LEAST_FP_def intro:inf_greatest assms)
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   375
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   376
lemma LFP_lemma2: 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   377
  assumes "Mono f" "f \<in> carrier L \<rightarrow> carrier L"
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   378
  shows "f (LFP f) \<sqsubseteq> LFP f"
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   379
proof (rule LFP_greatest)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   380
  have f: "\<And>x. x \<in> carrier L \<Longrightarrow> f x \<in> carrier L"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   381
    using assms by (auto simp add: Pi_def)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   382
  with assms show "f (LFP f) \<in> carrier L"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   383
    by blast
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   384
  show "\<And>u. \<lbrakk>u \<in> carrier L; f u \<sqsubseteq> u\<rbrakk> \<Longrightarrow> f (LFP f) \<sqsubseteq> u"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   385
    by (meson LFP_closed LFP_lowerbound assms(1) f local.le_trans use_iso1)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   386
qed
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   387
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   388
lemma LFP_lemma3: 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   389
  assumes "Mono f" "f \<in> carrier L \<rightarrow> carrier L"
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   390
  shows "LFP f \<sqsubseteq> f (LFP f)"
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   391
  using assms by (simp add: Pi_def) (metis LFP_closed LFP_lemma2 LFP_lowerbound assms(2) use_iso2)
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   392
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   393
lemma LFP_weak_unfold: 
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   394
  "\<lbrakk> Mono f; f \<in> carrier L \<rightarrow> carrier L \<rbrakk> \<Longrightarrow> LFP f .= f (LFP f)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   395
  by (auto intro: LFP_lemma2 LFP_lemma3 funcset_mem)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   396
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   397
lemma LFP_fixed_point [intro]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   398
  assumes "Mono f" "f \<in> carrier L \<rightarrow> carrier L"
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   399
  shows "LFP f \<in> fps L f"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   400
proof -
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   401
  have "f (LFP f) \<in> carrier L"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   402
    using assms(2) by blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   403
  with assms show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   404
    by (simp add: LFP_weak_unfold fps_def local.sym)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   405
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   406
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   407
lemma LFP_least_fixed_point:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   408
  assumes "Mono f" "f \<in> carrier L \<rightarrow> carrier L" "x \<in> fps L f"
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   409
  shows "LFP f \<sqsubseteq> x"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   410
  using assms by (force intro: LFP_lowerbound simp add: fps_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   411
  
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   412
lemma LFP_idem: 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   413
  assumes "f \<in> carrier L \<rightarrow> carrier L" "Mono f" "Idem f"
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   414
  shows "LFP f .= (f \<bottom>)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   415
proof (rule weak_le_antisym)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   416
  from assms(1) show fb: "f \<bottom> \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   417
    by (rule funcset_mem, simp)
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   418
  from assms show mf: "LFP f \<in> carrier L"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   419
    by blast
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   420
  show "LFP f \<sqsubseteq> f \<bottom>"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   421
  proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   422
    have "f (f \<bottom>) .= f \<bottom>"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   423
      by (auto simp add: fps_def fb assms(3) idempotent)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   424
    moreover have "f (f \<bottom>) \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   425
      by (rule funcset_mem[of f "carrier L"], simp_all add: assms fb)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   426
    ultimately show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   427
      by (auto intro: LFP_lowerbound simp add: fb)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   428
  qed
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   429
  show "f \<bottom> \<sqsubseteq> LFP f"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   430
  proof -
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   431
    have "f \<bottom> \<sqsubseteq> f (LFP f)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   432
      by (auto intro: use_iso1[of _ f] simp add: assms)
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   433
    moreover have "... .= LFP f"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   434
      using assms(1) assms(2) fps_def by force
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   435
    moreover from assms(1) have "f (LFP f) \<in> carrier L"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   436
      by (auto)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   437
    ultimately show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   438
      using fb by blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   439
  qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   440
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   441
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   442
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   443
subsubsection \<open>Greatest fixed points\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   444
  
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   445
lemma GFP_closed [intro, simp]:
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   446
  "GFP f \<in> carrier L"
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   447
  by (auto intro:sup_closed simp add:GREATEST_FP_def)
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   448
  
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   449
lemma GFP_upperbound:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   450
  assumes "x \<in> carrier L" "x \<sqsubseteq> f x"
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   451
  shows "x \<sqsubseteq> GFP f"
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   452
  by (auto intro:sup_upper assms simp add:GREATEST_FP_def)
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   453
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   454
lemma GFP_least: 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   455
  assumes "x \<in> carrier L" 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   456
          "(\<And>u. \<lbrakk> u \<in> carrier L; u \<sqsubseteq> f u \<rbrakk> \<Longrightarrow> u \<sqsubseteq> x)"
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   457
  shows "GFP f \<sqsubseteq> x"
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   458
  by (auto simp add:GREATEST_FP_def intro:sup_least assms)
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   459
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   460
lemma GFP_lemma2:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   461
  assumes "Mono f" "f \<in> carrier L \<rightarrow> carrier L"
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   462
  shows "GFP f \<sqsubseteq> f (GFP f)"
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   463
proof (rule GFP_least)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   464
  have f: "\<And>x. x \<in> carrier L \<Longrightarrow> f x \<in> carrier L"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   465
    using assms by (auto simp add: Pi_def)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   466
  with assms show "f (GFP f) \<in> carrier L"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   467
    by blast
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   468
  show "\<And>u. \<lbrakk>u \<in> carrier L; u \<sqsubseteq> f u\<rbrakk> \<Longrightarrow> u \<sqsubseteq> f (GFP f)"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   469
    by (meson GFP_closed GFP_upperbound le_trans assms(1) f local.le_trans use_iso1)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   470
qed
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   471
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   472
lemma GFP_lemma3:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   473
  assumes "Mono f" "f \<in> carrier L \<rightarrow> carrier L"
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   474
  shows "f (GFP f) \<sqsubseteq> GFP f"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   475
  by (metis GFP_closed GFP_lemma2 GFP_upperbound assms funcset_mem use_iso2)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   476
  
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   477
lemma GFP_weak_unfold: 
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   478
  "\<lbrakk> Mono f; f \<in> carrier L \<rightarrow> carrier L \<rbrakk> \<Longrightarrow> GFP f .= f (GFP f)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   479
  by (auto intro: GFP_lemma2 GFP_lemma3 funcset_mem)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   480
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   481
lemma (in weak_complete_lattice) GFP_fixed_point [intro]:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   482
  assumes "Mono f" "f \<in> carrier L \<rightarrow> carrier L"
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   483
  shows "GFP f \<in> fps L f"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   484
  using assms
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   485
proof -
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   486
  have "f (GFP f) \<in> carrier L"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   487
    using assms(2) by blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   488
  with assms show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   489
    by (simp add: GFP_weak_unfold fps_def local.sym)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   490
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   491
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   492
lemma GFP_greatest_fixed_point:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   493
  assumes "Mono f" "f \<in> carrier L \<rightarrow> carrier L" "x \<in> fps L f"
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   494
  shows "x \<sqsubseteq> GFP f"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   495
  using assms 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   496
  by (rule_tac GFP_upperbound, auto simp add: fps_def, meson PiE local.sym weak_refl)
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
lemma GFP_idem: 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   499
  assumes "f \<in> carrier L \<rightarrow> carrier L" "Mono f" "Idem f"
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   500
  shows "GFP f .= (f \<top>)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   501
proof (rule weak_le_antisym)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   502
  from assms(1) show fb: "f \<top> \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   503
    by (rule funcset_mem, simp)
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   504
  from assms show mf: "GFP f \<in> carrier L"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   505
    by blast
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   506
  show "f \<top> \<sqsubseteq> GFP f"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   507
  proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   508
    have "f (f \<top>) .= f \<top>"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   509
      by (auto simp add: fps_def fb assms(3) idempotent)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   510
    moreover have "f (f \<top>) \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   511
      by (rule funcset_mem[of f "carrier L"], simp_all add: assms fb)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   512
    ultimately show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   513
      by (rule_tac GFP_upperbound, simp_all add: fb local.sym)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   514
  qed
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   515
  show "GFP f \<sqsubseteq> f \<top>"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   516
  proof -
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   517
    have "GFP f \<sqsubseteq> f (GFP f)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   518
      by (simp add: GFP_lemma2 assms(1) assms(2))
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   519
    moreover have "... \<sqsubseteq> f \<top>"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   520
      by (auto intro: use_iso1[of _ f] simp add: assms)
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   521
    moreover from assms(1) have "f (GFP f) \<in> carrier L"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   522
      by (auto)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   523
    ultimately show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   524
      using fb local.le_trans by blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   525
  qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   526
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   527
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   528
end
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   529
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   530
67226
ec32cdaab97b isabelle update_cartouches -c -t;
wenzelm
parents: 67091
diff changeset
   531
subsection \<open>Complete lattices where \<open>eq\<close> is the Equality\<close>
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   532
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   533
locale complete_lattice = partial_order +
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   534
  assumes sup_exists:
67091
1393c2340eec more symbols;
wenzelm
parents: 66580
diff changeset
   535
    "[| A \<subseteq> carrier L |] ==> \<exists>s. least L s (Upper L A)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   536
    and inf_exists:
67091
1393c2340eec more symbols;
wenzelm
parents: 66580
diff changeset
   537
    "[| A \<subseteq> carrier L |] ==> \<exists>i. greatest L i (Lower L A)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   538
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   539
sublocale complete_lattice \<subseteq> lattice
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   540
proof
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   541
  fix x y
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   542
  assume a: "x \<in> carrier L" "y \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   543
  thus "\<exists>s. is_lub L s {x, y}"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   544
    by (rule_tac sup_exists[of "{x, y}"], auto)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   545
  from a show "\<exists>s. is_glb L s {x, y}"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   546
    by (rule_tac inf_exists[of "{x, y}"], auto)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   547
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   548
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   549
sublocale complete_lattice \<subseteq> weak?: weak_complete_lattice
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   550
  by standard (auto intro: sup_exists inf_exists)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   551
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   552
lemma complete_lattice_lattice [simp]: 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   553
  assumes "complete_lattice X"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   554
  shows "lattice X"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   555
proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   556
  interpret c: complete_lattice X
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   557
    by (simp add: assms)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   558
  show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   559
    by (unfold_locales)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   560
qed
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
text \<open>Introduction rule: the usual definition of complete lattice\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   563
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   564
lemma (in partial_order) complete_latticeI:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   565
  assumes sup_exists:
67091
1393c2340eec more symbols;
wenzelm
parents: 66580
diff changeset
   566
    "!!A. [| A \<subseteq> carrier L |] ==> \<exists>s. least L s (Upper L A)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   567
    and inf_exists:
67091
1393c2340eec more symbols;
wenzelm
parents: 66580
diff changeset
   568
    "!!A. [| A \<subseteq> carrier L |] ==> \<exists>i. greatest L i (Lower L A)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   569
  shows "complete_lattice L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   570
  by standard (auto intro: sup_exists inf_exists)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   571
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   572
theorem (in partial_order) complete_lattice_criterion1:
67091
1393c2340eec more symbols;
wenzelm
parents: 66580
diff changeset
   573
  assumes top_exists: "\<exists>g. greatest L g (carrier L)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   574
    and inf_exists:
67091
1393c2340eec more symbols;
wenzelm
parents: 66580
diff changeset
   575
      "!!A. [| A \<subseteq> carrier L; A \<noteq> {} |] ==> \<exists>i. greatest L i (Lower L A)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   576
  shows "complete_lattice L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   577
proof (rule complete_latticeI)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   578
  from top_exists obtain top where top: "greatest L top (carrier L)" ..
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   579
  fix A
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   580
  assume L: "A \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   581
  let ?B = "Upper L A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   582
  from L top have "top \<in> ?B" by (fast intro!: Upper_memI intro: greatest_le)
67091
1393c2340eec more symbols;
wenzelm
parents: 66580
diff changeset
   583
  then have B_non_empty: "?B \<noteq> {}" by fast
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   584
  have B_L: "?B \<subseteq> carrier L" by simp
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   585
  from inf_exists [OF B_L B_non_empty]
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   586
  obtain b where b_inf_B: "greatest L b (Lower L ?B)" ..
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   587
  then have bcarr: "b \<in> carrier L"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   588
    by blast
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   589
  have "least L b (Upper L A)"
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   590
  proof (rule least_UpperI)
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   591
    show "\<And>x. x \<in> A \<Longrightarrow> x \<sqsubseteq> b"
69712
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   592
      by (meson L Lower_memI Upper_memD b_inf_B greatest_le rev_subsetD)
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   593
    show "\<And>y. y \<in> Upper L A \<Longrightarrow> b \<sqsubseteq> y"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   594
      by (auto elim: greatest_Lower_below [OF b_inf_B])
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   595
  qed (use L bcarr in auto)
67091
1393c2340eec more symbols;
wenzelm
parents: 66580
diff changeset
   596
  then show "\<exists>s. least L s (Upper L A)" ..
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   597
next
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   598
  fix A
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   599
  assume L: "A \<subseteq> carrier L"
67091
1393c2340eec more symbols;
wenzelm
parents: 66580
diff changeset
   600
  show "\<exists>i. greatest L i (Lower L A)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   601
  proof (cases "A = {}")
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   602
    case True then show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   603
      by (simp add: top_exists)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   604
  next
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   605
    case False with L show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   606
      by (rule inf_exists)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   607
  qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   608
qed
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
(* TODO: prove dual version *)
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
subsection \<open>Fixed points\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   613
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   614
context complete_lattice
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   615
begin
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   616
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   617
lemma LFP_unfold: 
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   618
  "\<lbrakk> Mono f; f \<in> carrier L \<rightarrow> carrier L \<rbrakk> \<Longrightarrow> LFP f = f (LFP f)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   619
  using eq_is_equal weak.LFP_weak_unfold by auto
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 LFP_const:
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   622
  "t \<in> carrier L \<Longrightarrow> LFP (\<lambda> x. t) = t"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   623
  by (simp add: local.le_antisym weak.LFP_greatest weak.LFP_lowerbound)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   624
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   625
lemma LFP_id:
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   626
  "LFP id = \<bottom>"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   627
  by (simp add: local.le_antisym weak.LFP_lowerbound)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   628
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   629
lemma GFP_unfold:
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   630
  "\<lbrakk> Mono f; f \<in> carrier L \<rightarrow> carrier L \<rbrakk> \<Longrightarrow> GFP f = f (GFP f)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   631
  using eq_is_equal weak.GFP_weak_unfold by auto
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   632
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   633
lemma GFP_const:
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   634
  "t \<in> carrier L \<Longrightarrow> GFP (\<lambda> x. t) = t"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   635
  by (simp add: local.le_antisym weak.GFP_least weak.GFP_upperbound)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   636
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   637
lemma GFP_id:
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   638
  "GFP id = \<top>"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   639
  using weak.GFP_upperbound by auto
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   640
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   641
end
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   642
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
subsection \<open>Interval complete lattices\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   645
  
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   646
context weak_complete_lattice
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   647
begin
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   648
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   649
  lemma at_least_at_most_Sup: "\<lbrakk> a \<in> carrier L; b \<in> carrier L; a \<sqsubseteq> b \<rbrakk> \<Longrightarrow> \<Squnion> \<lbrace>a..b\<rbrace> .= b"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   650
    by (rule weak_le_antisym [OF sup_least sup_upper]) (auto simp add: at_least_at_most_closed)
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   651
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   652
  lemma at_least_at_most_Inf: "\<lbrakk> a \<in> carrier L; b \<in> carrier L; a \<sqsubseteq> b \<rbrakk> \<Longrightarrow> \<Sqinter> \<lbrace>a..b\<rbrace> .= a"
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   653
    by (rule weak_le_antisym [OF inf_lower inf_greatest]) (auto simp add: at_least_at_most_closed)
65099
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
end
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   656
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   657
lemma weak_complete_lattice_interval:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   658
  assumes "weak_complete_lattice L" "a \<in> carrier L" "b \<in> carrier L" "a \<sqsubseteq>\<^bsub>L\<^esub> b"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   659
  shows "weak_complete_lattice (L \<lparr> carrier := \<lbrace>a..b\<rbrace>\<^bsub>L\<^esub> \<rparr>)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   660
proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   661
  interpret L: weak_complete_lattice L
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   662
    by (simp add: assms)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   663
  interpret weak_partial_order "L \<lparr> carrier := \<lbrace>a..b\<rbrace>\<^bsub>L\<^esub> \<rparr>"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   664
  proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   665
    have "\<lbrace>a..b\<rbrace>\<^bsub>L\<^esub> \<subseteq> carrier L"
68488
dfbd80c3d180 more modernisaton and de-applying
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   666
      by (auto simp add: at_least_at_most_def)
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   667
    thus "weak_partial_order (L\<lparr>carrier := \<lbrace>a..b\<rbrace>\<^bsub>L\<^esub>\<rparr>)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   668
      by (simp add: L.weak_partial_order_axioms weak_partial_order_subset)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   669
  qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   670
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   671
  show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   672
  proof
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   673
    fix A
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   674
    assume a: "A \<subseteq> carrier (L\<lparr>carrier := \<lbrace>a..b\<rbrace>\<^bsub>L\<^esub>\<rparr>)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   675
    show "\<exists>s. is_lub (L\<lparr>carrier := \<lbrace>a..b\<rbrace>\<^bsub>L\<^esub>\<rparr>) s A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   676
    proof (cases "A = {}")
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   677
      case True
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   678
      thus ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   679
        by (rule_tac x="a" in exI, auto simp add: least_def assms)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   680
    next
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   681
      case False
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   682
      show ?thesis
68684
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   683
      proof (intro exI least_UpperI, simp_all)
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   684
        show b:"\<And> x. x \<in> A \<Longrightarrow> x \<sqsubseteq>\<^bsub>L\<^esub> \<Squnion>\<^bsub>L\<^esub>A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   685
          using a by (auto intro: L.sup_upper, meson L.at_least_at_most_closed L.sup_upper subset_trans)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   686
        show "\<And>y. y \<in> Upper (L\<lparr>carrier := \<lbrace>a..b\<rbrace>\<^bsub>L\<^esub>\<rparr>) A \<Longrightarrow> \<Squnion>\<^bsub>L\<^esub>A \<sqsubseteq>\<^bsub>L\<^esub> y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   687
          using a L.at_least_at_most_closed by (rule_tac L.sup_least, auto intro: funcset_mem simp add: Upper_def)
68684
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   688
        from a show *: "A \<subseteq> \<lbrace>a..b\<rbrace>\<^bsub>L\<^esub>"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   689
          by auto
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   690
        show "\<Squnion>\<^bsub>L\<^esub>A \<in> \<lbrace>a..b\<rbrace>\<^bsub>L\<^esub>"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   691
        proof (rule_tac L.at_least_at_most_member)
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   692
          show 1: "\<Squnion>\<^bsub>L\<^esub>A \<in> carrier L"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   693
            by (meson L.at_least_at_most_closed L.sup_closed subset_trans *)
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   694
          show "a \<sqsubseteq>\<^bsub>L\<^esub> \<Squnion>\<^bsub>L\<^esub>A"
69712
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   695
            by (meson "*" False L.at_least_at_most_closed L.at_least_at_most_lower L.le_trans L.sup_upper 1 all_not_in_conv assms(2) subsetD subset_trans)
68684
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   696
          show "\<Squnion>\<^bsub>L\<^esub>A \<sqsubseteq>\<^bsub>L\<^esub> b"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   697
          proof (rule L.sup_least)
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   698
            show "A \<subseteq> carrier L" "\<And>x. x \<in> A \<Longrightarrow> x \<sqsubseteq>\<^bsub>L\<^esub> b"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   699
              using * L.at_least_at_most_closed by blast+
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   700
          qed (simp add: assms)
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   701
        qed
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   702
      qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   703
    qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   704
    show "\<exists>s. is_glb (L\<lparr>carrier := \<lbrace>a..b\<rbrace>\<^bsub>L\<^esub>\<rparr>) s A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   705
    proof (cases "A = {}")
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   706
      case True
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   707
      thus ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   708
        by (rule_tac x="b" in exI, auto simp add: greatest_def assms)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   709
    next
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   710
      case False
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   711
      show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   712
      proof (rule_tac x="\<Sqinter>\<^bsub>L\<^esub> A" in exI, rule greatest_LowerI, simp_all)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   713
        show b:"\<And>x. x \<in> A \<Longrightarrow> \<Sqinter>\<^bsub>L\<^esub>A \<sqsubseteq>\<^bsub>L\<^esub> x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   714
          using a L.at_least_at_most_closed by (force intro!: L.inf_lower)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   715
        show "\<And>y. y \<in> Lower (L\<lparr>carrier := \<lbrace>a..b\<rbrace>\<^bsub>L\<^esub>\<rparr>) A \<Longrightarrow> y \<sqsubseteq>\<^bsub>L\<^esub> \<Sqinter>\<^bsub>L\<^esub>A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   716
           using a L.at_least_at_most_closed by (rule_tac L.inf_greatest, auto intro: funcset_carrier' simp add: Lower_def)
68684
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   717
        from a show *: "A \<subseteq> \<lbrace>a..b\<rbrace>\<^bsub>L\<^esub>"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   718
          by auto
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   719
        show "\<Sqinter>\<^bsub>L\<^esub>A \<in> \<lbrace>a..b\<rbrace>\<^bsub>L\<^esub>"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   720
        proof (rule_tac L.at_least_at_most_member)
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   721
          show 1: "\<Sqinter>\<^bsub>L\<^esub>A \<in> carrier L"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   722
            by (meson "*" L.at_least_at_most_closed L.inf_closed subset_trans)
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   723
          show "a \<sqsubseteq>\<^bsub>L\<^esub> \<Sqinter>\<^bsub>L\<^esub>A"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   724
            by (meson "*" L.at_least_at_most_closed L.at_least_at_most_lower L.inf_greatest assms(2) subsetD subset_trans)
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   725
          show "\<Sqinter>\<^bsub>L\<^esub>A \<sqsubseteq>\<^bsub>L\<^esub> b"
69712
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   726
            by (meson * 1 False L.at_least_at_most_closed L.at_least_at_most_upper L.inf_lower L.le_trans all_not_in_conv assms(3) subsetD subset_trans)
68684
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   727
        qed
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   728
      qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   729
    qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   730
  qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   731
qed
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
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   734
subsection \<open>Knaster-Tarski theorem and variants\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   735
  
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   736
text \<open>The set of fixed points of a complete lattice is itself a complete lattice\<close>
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
theorem Knaster_Tarski:
68684
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   739
  assumes "weak_complete_lattice L" and f: "f \<in> carrier L \<rightarrow> carrier L" and "isotone L L f"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   740
  shows "weak_complete_lattice (fpl L f)" (is "weak_complete_lattice ?L'")
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   741
proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   742
  interpret L: weak_complete_lattice L
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   743
    by (simp add: assms)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   744
  interpret weak_partial_order ?L'
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   745
  proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   746
    have "{x \<in> carrier L. f x .=\<^bsub>L\<^esub> x} \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   747
      by (auto)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   748
    thus "weak_partial_order ?L'"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   749
      by (simp add: L.weak_partial_order_axioms weak_partial_order_subset)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   750
  qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   751
  show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   752
  proof (unfold_locales, simp_all)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   753
    fix A
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   754
    assume A: "A \<subseteq> fps L f"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   755
    show "\<exists>s. is_lub (fpl L f) s A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   756
    proof
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   757
      from A have AL: "A \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   758
        by (meson fps_carrier subset_eq)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   759
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   760
      let ?w = "\<Squnion>\<^bsub>L\<^esub> A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   761
      have w: "f (\<Squnion>\<^bsub>L\<^esub>A) \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   762
        by (rule funcset_mem[of f "carrier L"], simp_all add: AL assms(2))
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   763
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   764
      have pf_w: "(\<Squnion>\<^bsub>L\<^esub> A) \<sqsubseteq>\<^bsub>L\<^esub> f (\<Squnion>\<^bsub>L\<^esub> A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   765
        by (simp add: A L.weak_sup_pre_fixed_point assms(2) assms(3))
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   766
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   767
      have f_top_chain: "f ` \<lbrace>?w..\<top>\<^bsub>L\<^esub>\<rbrace>\<^bsub>L\<^esub> \<subseteq> \<lbrace>?w..\<top>\<^bsub>L\<^esub>\<rbrace>\<^bsub>L\<^esub>"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   768
      proof (auto simp add: at_least_at_most_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   769
        fix x
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   770
        assume b: "x \<in> carrier L" "\<Squnion>\<^bsub>L\<^esub>A \<sqsubseteq>\<^bsub>L\<^esub> x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   771
        from b show fx: "f x \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   772
          using assms(2) by blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   773
        show "\<Squnion>\<^bsub>L\<^esub>A \<sqsubseteq>\<^bsub>L\<^esub> f x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   774
        proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   775
          have "?w \<sqsubseteq>\<^bsub>L\<^esub> f ?w"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   776
          proof (rule_tac L.sup_least, simp_all add: AL w)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   777
            fix y
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   778
            assume c: "y \<in> A" 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   779
            hence y: "y \<in> fps L f"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   780
              using A subsetCE by blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   781
            with assms have "y .=\<^bsub>L\<^esub> f y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   782
            proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   783
              from y have "y \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   784
                by (simp add: fps_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   785
              moreover hence "f y \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   786
                by (rule_tac funcset_mem[of f "carrier L"], simp_all add: assms)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   787
              ultimately show ?thesis using y
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   788
                by (rule_tac L.sym, simp_all add: L.use_fps)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   789
            qed              
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   790
            moreover have "y \<sqsubseteq>\<^bsub>L\<^esub> \<Squnion>\<^bsub>L\<^esub>A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   791
              by (simp add: AL L.sup_upper c(1))
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   792
            ultimately show "y \<sqsubseteq>\<^bsub>L\<^esub> f (\<Squnion>\<^bsub>L\<^esub>A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   793
              by (meson fps_def AL funcset_mem L.refl L.weak_complete_lattice_axioms assms(2) assms(3) c(1) isotone_def rev_subsetD weak_complete_lattice.sup_closed weak_partial_order.le_cong)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   794
          qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   795
          thus ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   796
            by (meson AL funcset_mem L.le_trans L.sup_closed assms(2) assms(3) b(1) b(2) use_iso2)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   797
        qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   798
   
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   799
        show "f x \<sqsubseteq>\<^bsub>L\<^esub> \<top>\<^bsub>L\<^esub>"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   800
          by (simp add: fx)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   801
      qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   802
  
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   803
      let ?L' = "L\<lparr> carrier := \<lbrace>?w..\<top>\<^bsub>L\<^esub>\<rbrace>\<^bsub>L\<^esub> \<rparr>"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   804
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   805
      interpret L': weak_complete_lattice ?L'
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   806
        by (auto intro: weak_complete_lattice_interval simp add: L.weak_complete_lattice_axioms AL)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   807
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   808
      let ?L'' = "L\<lparr> carrier := fps L f \<rparr>"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   809
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   810
      show "is_lub ?L'' (LFP\<^bsub>?L'\<^esub> f) A"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   811
      proof (rule least_UpperI, simp_all)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   812
        fix x
68684
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   813
        assume x: "x \<in> Upper ?L'' A"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   814
        have "LFP\<^bsub>?L'\<^esub> f \<sqsubseteq>\<^bsub>?L'\<^esub> x"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   815
        proof (rule L'.LFP_lowerbound, simp_all)
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   816
          show "x \<in> \<lbrace>\<Squnion>\<^bsub>L\<^esub>A..\<top>\<^bsub>L\<^esub>\<rbrace>\<^bsub>L\<^esub>"
69712
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   817
            using x by (auto simp add: Upper_def A AL L.at_least_at_most_member L.sup_least rev_subsetD)    
68684
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   818
          with x show "f x \<sqsubseteq>\<^bsub>L\<^esub> x"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   819
            by (simp add: Upper_def) (meson L.at_least_at_most_closed L.use_fps L.weak_refl subsetD f_top_chain imageI)
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   820
        qed
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   821
        thus " LFP\<^bsub>?L'\<^esub> f \<sqsubseteq>\<^bsub>L\<^esub> x"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   822
          by (simp)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   823
      next
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   824
        fix x
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   825
        assume xA: "x \<in> A"
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   826
        show "x \<sqsubseteq>\<^bsub>L\<^esub> LFP\<^bsub>?L'\<^esub> f"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   827
        proof -
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   828
          have "LFP\<^bsub>?L'\<^esub> f \<in> carrier ?L'"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   829
            by blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   830
          thus ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   831
            by (simp, meson AL L.at_least_at_most_closed L.at_least_at_most_lower L.le_trans L.sup_closed L.sup_upper xA subsetCE)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   832
        qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   833
      next
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   834
        show "A \<subseteq> fps L f"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   835
          by (simp add: A)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   836
      next
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   837
        show "LFP\<^bsub>?L'\<^esub> f \<in> fps L f"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   838
        proof (auto simp add: fps_def)
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   839
          have "LFP\<^bsub>?L'\<^esub> f \<in> carrier ?L'"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   840
            by (rule L'.LFP_closed)
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   841
          thus c:"LFP\<^bsub>?L'\<^esub> f \<in> carrier L"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   842
             by (auto simp add: at_least_at_most_def)
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   843
          have "LFP\<^bsub>?L'\<^esub> f .=\<^bsub>?L'\<^esub> f (LFP\<^bsub>?L'\<^esub> f)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   844
          proof (rule "L'.LFP_weak_unfold", simp_all)
68684
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   845
            have "\<And>x. \<lbrakk>x \<in> carrier L; \<Squnion>\<^bsub>L\<^esub>A \<sqsubseteq>\<^bsub>L\<^esub> x\<rbrakk> \<Longrightarrow> \<Squnion>\<^bsub>L\<^esub>A \<sqsubseteq>\<^bsub>L\<^esub> f x"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   846
              by (meson AL funcset_mem L.le_trans L.sup_closed assms(2) assms(3) pf_w use_iso2)
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   847
            with f show "f \<in> \<lbrace>\<Squnion>\<^bsub>L\<^esub>A..\<top>\<^bsub>L\<^esub>\<rbrace>\<^bsub>L\<^esub> \<rightarrow> \<lbrace>\<Squnion>\<^bsub>L\<^esub>A..\<top>\<^bsub>L\<^esub>\<rbrace>\<^bsub>L\<^esub>"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   848
              by (auto simp add: Pi_def at_least_at_most_def)
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   849
            show "Mono\<^bsub>L\<lparr>carrier := \<lbrace>\<Squnion>\<^bsub>L\<^esub>A..\<top>\<^bsub>L\<^esub>\<rbrace>\<^bsub>L\<^esub>\<rparr>\<^esub> f"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   850
              using L'.weak_partial_order_axioms assms(3) 
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   851
              by (auto simp add: isotone_def) (meson L.at_least_at_most_closed subsetCE)
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   852
          qed
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   853
          thus "f (LFP\<^bsub>?L'\<^esub> f) .=\<^bsub>L\<^esub> LFP\<^bsub>?L'\<^esub> f"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   854
            by (simp add: L.equivalence_axioms funcset_carrier' c assms(2) equivalence.sym) 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   855
        qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   856
      qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   857
    qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   858
    show "\<exists>i. is_glb (L\<lparr>carrier := fps L f\<rparr>) i A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   859
    proof
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   860
      from A have AL: "A \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   861
        by (meson fps_carrier subset_eq)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   862
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   863
      let ?w = "\<Sqinter>\<^bsub>L\<^esub> A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   864
      have w: "f (\<Sqinter>\<^bsub>L\<^esub>A) \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   865
        by (simp add: AL funcset_carrier' assms(2))
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   866
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   867
      have pf_w: "f (\<Sqinter>\<^bsub>L\<^esub> A) \<sqsubseteq>\<^bsub>L\<^esub> (\<Sqinter>\<^bsub>L\<^esub> A)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   868
        by (simp add: A L.weak_sup_post_fixed_point assms(2) assms(3))
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   869
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   870
      have f_bot_chain: "f ` \<lbrace>\<bottom>\<^bsub>L\<^esub>..?w\<rbrace>\<^bsub>L\<^esub> \<subseteq> \<lbrace>\<bottom>\<^bsub>L\<^esub>..?w\<rbrace>\<^bsub>L\<^esub>"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   871
      proof (auto simp add: at_least_at_most_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   872
        fix x
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   873
        assume b: "x \<in> carrier L" "x \<sqsubseteq>\<^bsub>L\<^esub> \<Sqinter>\<^bsub>L\<^esub>A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   874
        from b show fx: "f x \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   875
          using assms(2) by blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   876
        show "f x \<sqsubseteq>\<^bsub>L\<^esub> \<Sqinter>\<^bsub>L\<^esub>A"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   877
        proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   878
          have "f ?w \<sqsubseteq>\<^bsub>L\<^esub> ?w"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   879
          proof (rule_tac L.inf_greatest, simp_all add: AL w)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   880
            fix y
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   881
            assume c: "y \<in> A" 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   882
            with assms have "y .=\<^bsub>L\<^esub> f y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   883
              by (metis (no_types, lifting) A funcset_carrier'[OF assms(2)] L.sym fps_def mem_Collect_eq subset_eq)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   884
            moreover have "\<Sqinter>\<^bsub>L\<^esub>A \<sqsubseteq>\<^bsub>L\<^esub> y"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   885
              by (simp add: AL L.inf_lower c)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   886
            ultimately show "f (\<Sqinter>\<^bsub>L\<^esub>A) \<sqsubseteq>\<^bsub>L\<^esub> y"
69712
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   887
              by (meson AL L.inf_closed L.le_trans c pf_w rev_subsetD w)
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   888
          qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   889
          thus ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   890
            by (meson AL L.inf_closed L.le_trans assms(3) b(1) b(2) fx use_iso2 w)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   891
        qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   892
        show "\<bottom>\<^bsub>L\<^esub> \<sqsubseteq>\<^bsub>L\<^esub> f x"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   893
          by (simp add: fx)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   894
      qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   895
  
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   896
      let ?L' = "L\<lparr> carrier := \<lbrace>\<bottom>\<^bsub>L\<^esub>..?w\<rbrace>\<^bsub>L\<^esub> \<rparr>"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   897
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   898
      interpret L': weak_complete_lattice ?L'
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   899
        by (auto intro!: weak_complete_lattice_interval simp add: L.weak_complete_lattice_axioms AL)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   900
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   901
      let ?L'' = "L\<lparr> carrier := fps L f \<rparr>"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   902
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   903
      show "is_glb ?L'' (GFP\<^bsub>?L'\<^esub> f) A"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   904
      proof (rule greatest_LowerI, simp_all)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   905
        fix x
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   906
        assume "x \<in> Lower ?L'' A"
68684
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   907
        then have x: "\<forall>y. y \<in> A \<and> y \<in> fps L f \<longrightarrow> x \<sqsubseteq>\<^bsub>L\<^esub> y" "x \<in> fps L f"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   908
          by (auto simp add: Lower_def)
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   909
        have "x \<sqsubseteq>\<^bsub>?L'\<^esub> GFP\<^bsub>?L'\<^esub> f"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   910
          unfolding Lower_def
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   911
        proof (rule_tac L'.GFP_upperbound; simp)
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   912
          show "x \<in> \<lbrace>\<bottom>\<^bsub>L\<^esub>..\<Sqinter>\<^bsub>L\<^esub>A\<rbrace>\<^bsub>L\<^esub>"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   913
            by (meson x A AL L.at_least_at_most_member L.bottom_lower L.inf_greatest contra_subsetD fps_carrier)
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   914
          show "x \<sqsubseteq>\<^bsub>L\<^esub> f x"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   915
            using x by (simp add: funcset_carrier' L.sym assms(2) fps_def)
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   916
        qed
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   917
        thus "x \<sqsubseteq>\<^bsub>L\<^esub> GFP\<^bsub>?L'\<^esub> f"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   918
          by (simp)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   919
      next
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   920
        fix x
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   921
        assume xA: "x \<in> A"
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   922
        show "GFP\<^bsub>?L'\<^esub> f \<sqsubseteq>\<^bsub>L\<^esub> x"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   923
        proof -
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   924
          have "GFP\<^bsub>?L'\<^esub> f \<in> carrier ?L'"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   925
            by blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   926
          thus ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   927
            by (simp, meson AL L.at_least_at_most_closed L.at_least_at_most_upper L.inf_closed L.inf_lower L.le_trans subsetCE xA)     
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   928
        qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   929
      next
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   930
        show "A \<subseteq> fps L f"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   931
          by (simp add: A)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   932
      next
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   933
        show "GFP\<^bsub>?L'\<^esub> f \<in> fps L f"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   934
        proof (auto simp add: fps_def)
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   935
          have "GFP\<^bsub>?L'\<^esub> f \<in> carrier ?L'"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   936
            by (rule L'.GFP_closed)
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   937
          thus c:"GFP\<^bsub>?L'\<^esub> f \<in> carrier L"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   938
             by (auto simp add: at_least_at_most_def)
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   939
          have "GFP\<^bsub>?L'\<^esub> f .=\<^bsub>?L'\<^esub> f (GFP\<^bsub>?L'\<^esub> f)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   940
          proof (rule "L'.GFP_weak_unfold", simp_all)
68684
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   941
            have "\<And>x. \<lbrakk>x \<in> carrier L; x \<sqsubseteq>\<^bsub>L\<^esub> \<Sqinter>\<^bsub>L\<^esub>A\<rbrakk> \<Longrightarrow> f x \<sqsubseteq>\<^bsub>L\<^esub> \<Sqinter>\<^bsub>L\<^esub>A"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   942
              by (meson AL funcset_carrier L.inf_closed L.le_trans assms(2) assms(3) pf_w use_iso2)
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   943
            with assms(2) show "f \<in> \<lbrace>\<bottom>\<^bsub>L\<^esub>..?w\<rbrace>\<^bsub>L\<^esub> \<rightarrow> \<lbrace>\<bottom>\<^bsub>L\<^esub>..?w\<rbrace>\<^bsub>L\<^esub>"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   944
              by (auto simp add: Pi_def at_least_at_most_def)
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   945
            have "\<And>x y. \<lbrakk>x \<in> \<lbrace>\<bottom>\<^bsub>L\<^esub>..\<Sqinter>\<^bsub>L\<^esub>A\<rbrace>\<^bsub>L\<^esub>; y \<in> \<lbrace>\<bottom>\<^bsub>L\<^esub>..\<Sqinter>\<^bsub>L\<^esub>A\<rbrace>\<^bsub>L\<^esub>; x \<sqsubseteq>\<^bsub>L\<^esub> y\<rbrakk> \<Longrightarrow> f x \<sqsubseteq>\<^bsub>L\<^esub> f y"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   946
              by (meson L.at_least_at_most_closed subsetD use_iso1  assms(3)) 
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   947
            with L'.weak_partial_order_axioms show "Mono\<^bsub>L\<lparr>carrier := \<lbrace>\<bottom>\<^bsub>L\<^esub>..?w\<rbrace>\<^bsub>L\<^esub>\<rparr>\<^esub> f"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
   948
              by (auto simp add: isotone_def)
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   949
          qed
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   950
          thus "f (GFP\<^bsub>?L'\<^esub> f) .=\<^bsub>L\<^esub> GFP\<^bsub>?L'\<^esub> f"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   951
            by (simp add: L.equivalence_axioms funcset_carrier' c assms(2) equivalence.sym) 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   952
        qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   953
      qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   954
    qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   955
  qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   956
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   957
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   958
theorem Knaster_Tarski_top:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   959
  assumes "weak_complete_lattice L" "isotone L L f" "f \<in> carrier L \<rightarrow> carrier L"
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   960
  shows "\<top>\<^bsub>fpl L f\<^esub> .=\<^bsub>L\<^esub> GFP\<^bsub>L\<^esub> f"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   961
proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   962
  interpret L: weak_complete_lattice L
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   963
    by (simp add: assms)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   964
  interpret L': weak_complete_lattice "fpl L f"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   965
    by (rule Knaster_Tarski, simp_all add: assms)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   966
  show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   967
  proof (rule L.weak_le_antisym, simp_all)
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   968
    show "\<top>\<^bsub>fpl L f\<^esub> \<sqsubseteq>\<^bsub>L\<^esub> GFP\<^bsub>L\<^esub> f"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   969
      by (rule L.GFP_greatest_fixed_point, simp_all add: assms L'.top_closed[simplified])
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   970
    show "GFP\<^bsub>L\<^esub> f \<sqsubseteq>\<^bsub>L\<^esub> \<top>\<^bsub>fpl L f\<^esub>"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   971
    proof -
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   972
      have "GFP\<^bsub>L\<^esub> f \<in> fps L f"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   973
        by (rule L.GFP_fixed_point, simp_all add: assms)
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   974
      hence "GFP\<^bsub>L\<^esub> f \<in> carrier (fpl L f)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   975
        by simp
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   976
      hence "GFP\<^bsub>L\<^esub> f \<sqsubseteq>\<^bsub>fpl L f\<^esub> \<top>\<^bsub>fpl L f\<^esub>"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   977
        by (rule L'.top_higher)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   978
      thus ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   979
        by simp
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   980
    qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   981
    show "\<top>\<^bsub>fpl L f\<^esub> \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   982
    proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   983
      have "carrier (fpl L f) \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   984
        by (auto simp add: fps_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   985
      with L'.top_closed show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   986
        by blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   987
    qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   988
  qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   989
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   990
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   991
theorem Knaster_Tarski_bottom:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   992
  assumes "weak_complete_lattice L" "isotone L L f" "f \<in> carrier L \<rightarrow> carrier L"
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
   993
  shows "\<bottom>\<^bsub>fpl L f\<^esub> .=\<^bsub>L\<^esub> LFP\<^bsub>L\<^esub> f"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   994
proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   995
  interpret L: weak_complete_lattice L
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   996
    by (simp add: assms)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   997
  interpret L': weak_complete_lattice "fpl L f"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   998
    by (rule Knaster_Tarski, simp_all add: assms)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
   999
  show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1000
  proof (rule L.weak_le_antisym, simp_all)
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
  1001
    show "LFP\<^bsub>L\<^esub> f \<sqsubseteq>\<^bsub>L\<^esub> \<bottom>\<^bsub>fpl L f\<^esub>"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1002
      by (rule L.LFP_least_fixed_point, simp_all add: assms L'.bottom_closed[simplified])
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
  1003
    show "\<bottom>\<^bsub>fpl L f\<^esub> \<sqsubseteq>\<^bsub>L\<^esub> LFP\<^bsub>L\<^esub> f"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1004
    proof -
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
  1005
      have "LFP\<^bsub>L\<^esub> f \<in> fps L f"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1006
        by (rule L.LFP_fixed_point, simp_all add: assms)
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
  1007
      hence "LFP\<^bsub>L\<^esub> f \<in> carrier (fpl L f)"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1008
        by simp
66580
e5b1d4d55bf6 Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
ballarin
parents: 66579
diff changeset
  1009
      hence "\<bottom>\<^bsub>fpl L f\<^esub> \<sqsubseteq>\<^bsub>fpl L f\<^esub> LFP\<^bsub>L\<^esub> f"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1010
        by (rule L'.bottom_lower)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1011
      thus ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1012
        by simp
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1013
    qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1014
    show "\<bottom>\<^bsub>fpl L f\<^esub> \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1015
    proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1016
      have "carrier (fpl L f) \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1017
        by (auto simp add: fps_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1018
      with L'.bottom_closed show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1019
        by blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1020
    qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1021
  qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1022
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1023
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1024
text \<open>If a function is both idempotent and isotone then the image of the function forms a complete lattice\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1025
  
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1026
theorem Knaster_Tarski_idem:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1027
  assumes "complete_lattice L" "f \<in> carrier L \<rightarrow> carrier L" "isotone L L f" "idempotent L f"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1028
  shows "complete_lattice (L\<lparr>carrier := f ` carrier L\<rparr>)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1029
proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1030
  interpret L: complete_lattice L
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1031
    by (simp add: assms)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1032
  have "fps L f = f ` carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1033
    using L.weak.fps_idem[OF assms(2) assms(4)]
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1034
    by (simp add: L.set_eq_is_eq)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1035
  then interpret L': weak_complete_lattice "(L\<lparr>carrier := f ` carrier L\<rparr>)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1036
    by (metis Knaster_Tarski L.weak.weak_complete_lattice_axioms assms(2) assms(3))
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1037
  show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1038
    using L'.sup_exists L'.inf_exists
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1039
    by (unfold_locales, auto simp add: L.eq_is_equal)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1040
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1041
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1042
theorem Knaster_Tarski_idem_extremes:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1043
  assumes "weak_complete_lattice L" "isotone L L f" "idempotent L f" "f \<in> carrier L \<rightarrow> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1044
  shows "\<top>\<^bsub>fpl L f\<^esub> .=\<^bsub>L\<^esub> f (\<top>\<^bsub>L\<^esub>)" "\<bottom>\<^bsub>fpl L f\<^esub> .=\<^bsub>L\<^esub> f (\<bottom>\<^bsub>L\<^esub>)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1045
proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1046
  interpret L: weak_complete_lattice "L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1047
    by (simp_all add: assms)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1048
  interpret L': weak_complete_lattice "fpl L f"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1049
    by (rule Knaster_Tarski, simp_all add: assms)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1050
  have FA: "fps L f \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1051
    by (auto simp add: fps_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1052
  show "\<top>\<^bsub>fpl L f\<^esub> .=\<^bsub>L\<^esub> f (\<top>\<^bsub>L\<^esub>)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1053
  proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1054
    from FA have "\<top>\<^bsub>fpl L f\<^esub> \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1055
    proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1056
      have "\<top>\<^bsub>fpl L f\<^esub> \<in> fps L f"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1057
        using L'.top_closed by auto
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1058
      thus ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1059
        using FA by blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1060
    qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1061
    moreover with assms have "f \<top>\<^bsub>L\<^esub> \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1062
      by (auto)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1063
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1064
    ultimately show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1065
      using L.trans[OF Knaster_Tarski_top[of L f] L.GFP_idem[of f]]
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1066
      by (simp_all add: assms)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1067
  qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1068
  show "\<bottom>\<^bsub>fpl L f\<^esub> .=\<^bsub>L\<^esub> f (\<bottom>\<^bsub>L\<^esub>)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1069
  proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1070
    from FA have "\<bottom>\<^bsub>fpl L f\<^esub> \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1071
    proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1072
      have "\<bottom>\<^bsub>fpl L f\<^esub> \<in> fps L f"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1073
        using L'.bottom_closed by auto
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1074
      thus ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1075
        using FA by blast
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1076
    qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1077
    moreover with assms have "f \<bottom>\<^bsub>L\<^esub> \<in> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1078
      by (auto)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1079
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1080
    ultimately show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1081
      using L.trans[OF Knaster_Tarski_bottom[of L f] L.LFP_idem[of f]]
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1082
      by (simp_all add: assms)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1083
  qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1084
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1085
66187
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1086
theorem Knaster_Tarski_idem_inf_eq:
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1087
  assumes "weak_complete_lattice L" "isotone L L f" "idempotent L f" "f \<in> carrier L \<rightarrow> carrier L"
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1088
          "A \<subseteq> fps L f"
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1089
  shows "\<Sqinter>\<^bsub>fpl L f\<^esub> A .=\<^bsub>L\<^esub> f (\<Sqinter>\<^bsub>L\<^esub> A)"
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1090
proof -
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1091
  interpret L: weak_complete_lattice "L"
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1092
    by (simp_all add: assms)
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1093
  interpret L': weak_complete_lattice "fpl L f"
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1094
    by (rule Knaster_Tarski, simp_all add: assms)
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1095
  have FA: "fps L f \<subseteq> carrier L"
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1096
    by (auto simp add: fps_def)
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1097
  have A: "A \<subseteq> carrier L"
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1098
    using FA assms(5) by blast
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1099
  have fA: "f (\<Sqinter>\<^bsub>L\<^esub>A) \<in> fps L f"
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1100
    by (metis (no_types, lifting) A L.idempotent L.inf_closed PiE assms(3) assms(4) fps_def mem_Collect_eq)
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1101
  have infA: "\<Sqinter>\<^bsub>fpl L f\<^esub>A \<in> fps L f"
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1102
    by (rule L'.inf_closed[simplified], simp add: assms)
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1103
  show ?thesis
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1104
  proof (rule L.weak_le_antisym)
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1105
    show ic: "\<Sqinter>\<^bsub>fpl L f\<^esub>A \<in> carrier L"
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1106
      using FA infA by blast
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1107
    show fc: "f (\<Sqinter>\<^bsub>L\<^esub>A) \<in> carrier L"
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1108
      using FA fA by blast
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1109
    show "f (\<Sqinter>\<^bsub>L\<^esub>A) \<sqsubseteq>\<^bsub>L\<^esub> \<Sqinter>\<^bsub>fpl L f\<^esub>A"
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1110
    proof -
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1111
      have "\<And>x. x \<in> A \<Longrightarrow> f (\<Sqinter>\<^bsub>L\<^esub>A) \<sqsubseteq>\<^bsub>L\<^esub> x"
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1112
        by (meson A FA L.inf_closed L.inf_lower L.le_trans L.weak_sup_post_fixed_point assms(2) assms(4) assms(5) fA subsetCE)
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1113
      hence "f (\<Sqinter>\<^bsub>L\<^esub>A) \<sqsubseteq>\<^bsub>fpl L f\<^esub> \<Sqinter>\<^bsub>fpl L f\<^esub>A"
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1114
        by (rule_tac L'.inf_greatest, simp_all add: fA assms(3,5))
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1115
      thus ?thesis
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1116
        by (simp)
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1117
    qed
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1118
    show "\<Sqinter>\<^bsub>fpl L f\<^esub>A \<sqsubseteq>\<^bsub>L\<^esub> f (\<Sqinter>\<^bsub>L\<^esub>A)"
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1119
    proof -
68684
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
  1120
      have *: "\<Sqinter>\<^bsub>fpl L f\<^esub>A \<in> carrier L"
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
  1121
        using FA infA by blast
66187
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1122
      have "\<And>x. x \<in> A \<Longrightarrow> \<Sqinter>\<^bsub>fpl L f\<^esub>A \<sqsubseteq>\<^bsub>fpl L f\<^esub> x"
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1123
        by (rule L'.inf_lower, simp_all add: assms)
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1124
      hence "\<Sqinter>\<^bsub>fpl L f\<^esub>A \<sqsubseteq>\<^bsub>L\<^esub> (\<Sqinter>\<^bsub>L\<^esub>A)"
68684
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
  1125
        by (rule_tac L.inf_greatest, simp_all add: A *)
66187
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1126
      hence 1:"f(\<Sqinter>\<^bsub>fpl L f\<^esub>A) \<sqsubseteq>\<^bsub>L\<^esub> f(\<Sqinter>\<^bsub>L\<^esub>A)"
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1127
        by (metis (no_types, lifting) A FA L.inf_closed assms(2) infA subsetCE use_iso1)
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1128
      have 2:"\<Sqinter>\<^bsub>fpl L f\<^esub>A \<sqsubseteq>\<^bsub>L\<^esub> f (\<Sqinter>\<^bsub>fpl L f\<^esub>A)"
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1129
        by (metis (no_types, lifting) FA L.sym L.use_fps L.weak_complete_lattice_axioms PiE assms(4) infA subsetCE weak_complete_lattice_def weak_partial_order.weak_refl)
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1130
      show ?thesis  
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1131
        using FA fA infA by (auto intro!: L.le_trans[OF 2 1] ic fc, metis FA PiE assms(4) subsetCE)
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1132
    qed
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1133
  qed
85925d4ae93d Additional corollary Knaster_Tarski_idem_inf_eq.
ballarin
parents: 65099
diff changeset
  1134
qed  
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1135
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1136
subsection \<open>Examples\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1137
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1138
subsubsection \<open>The Powerset of a Set is a Complete Lattice\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1139
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1140
theorem powerset_is_complete_lattice:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67340
diff changeset
  1141
  "complete_lattice \<lparr>carrier = Pow A, eq = (=), le = (\<subseteq>)\<rparr>"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1142
  (is "complete_lattice ?L")
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1143
proof (rule partial_order.complete_latticeI)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1144
  show "partial_order ?L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1145
    by standard auto
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1146
next
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1147
  fix B
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1148
  assume "B \<subseteq> carrier ?L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1149
  then have "least ?L (\<Union> B) (Upper ?L B)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1150
    by (fastforce intro!: least_UpperI simp: Upper_def)
67091
1393c2340eec more symbols;
wenzelm
parents: 66580
diff changeset
  1151
  then show "\<exists>s. least ?L s (Upper ?L B)" ..
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1152
next
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1153
  fix B
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1154
  assume "B \<subseteq> carrier ?L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1155
  then have "greatest ?L (\<Inter> B \<inter> A) (Lower ?L B)"
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68684
diff changeset
  1156
    txt \<open>\<^term>\<open>\<Inter> B\<close> is not the infimum of \<^term>\<open>B\<close>:
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68684
diff changeset
  1157
      \<^term>\<open>\<Inter> {} = UNIV\<close> which is in general bigger than \<^term>\<open>A\<close>! \<close>
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1158
    by (fastforce intro!: greatest_LowerI simp: Lower_def)
67091
1393c2340eec more symbols;
wenzelm
parents: 66580
diff changeset
  1159
  then show "\<exists>i. greatest ?L i (Lower ?L B)" ..
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1160
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1161
66579
2db3fe23fdaf Revert 5a42eddc11c1.
ballarin
parents: 66501
diff changeset
  1162
text \<open>Another example, that of the lattice of subgroups of a group,
2db3fe23fdaf Revert 5a42eddc11c1.
ballarin
parents: 66501
diff changeset
  1163
  can be found in Group theory (Section~\ref{sec:subgroup-lattice}).\<close>
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1164
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1165
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1166
subsection \<open>Limit preserving functions\<close>
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1167
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1168
definition weak_sup_pres :: "('a, 'c) gorder_scheme \<Rightarrow> ('b, 'd) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool" where
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1169
"weak_sup_pres X Y f \<equiv> complete_lattice X \<and> complete_lattice Y \<and> (\<forall> A \<subseteq> carrier X. A \<noteq> {} \<longrightarrow> f (\<Squnion>\<^bsub>X\<^esub> A) = (\<Squnion>\<^bsub>Y\<^esub> (f ` A)))"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1170
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1171
definition sup_pres :: "('a, 'c) gorder_scheme \<Rightarrow> ('b, 'd) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool" where
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1172
"sup_pres X Y f \<equiv> complete_lattice X \<and> complete_lattice Y \<and> (\<forall> A \<subseteq> carrier X. f (\<Squnion>\<^bsub>X\<^esub> A) = (\<Squnion>\<^bsub>Y\<^esub> (f ` A)))"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1173
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1174
definition weak_inf_pres :: "('a, 'c) gorder_scheme \<Rightarrow> ('b, 'd) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool" where
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1175
"weak_inf_pres X Y f \<equiv> complete_lattice X \<and> complete_lattice Y \<and> (\<forall> A \<subseteq> carrier X. A \<noteq> {} \<longrightarrow> f (\<Sqinter>\<^bsub>X\<^esub> A) = (\<Sqinter>\<^bsub>Y\<^esub> (f ` A)))"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1176
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1177
definition inf_pres :: "('a, 'c) gorder_scheme \<Rightarrow> ('b, 'd) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool" where
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1178
"inf_pres X Y f \<equiv> complete_lattice X \<and> complete_lattice Y \<and> (\<forall> A \<subseteq> carrier X. f (\<Sqinter>\<^bsub>X\<^esub> A) = (\<Sqinter>\<^bsub>Y\<^esub> (f ` A)))"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1179
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1180
lemma weak_sup_pres:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1181
  "sup_pres X Y f \<Longrightarrow> weak_sup_pres X Y f"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1182
  by (simp add: sup_pres_def weak_sup_pres_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1183
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1184
lemma weak_inf_pres:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1185
  "inf_pres X Y f \<Longrightarrow> weak_inf_pres X Y f"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1186
  by (simp add: inf_pres_def weak_inf_pres_def)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1187
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1188
lemma sup_pres_is_join_pres:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1189
  assumes "weak_sup_pres X Y f"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1190
  shows "join_pres X Y f"
68684
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
  1191
  using assms by (auto simp: join_pres_def weak_sup_pres_def join_def)
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1192
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1193
lemma inf_pres_is_meet_pres:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1194
  assumes "weak_inf_pres X Y f"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1195
  shows "meet_pres X Y f"
68684
9a42b84f8838 de-applying
paulson <lp15@cam.ac.uk>
parents: 68488
diff changeset
  1196
  using assms by (auto simp: meet_pres_def weak_inf_pres_def meet_def)
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1197
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff changeset
  1198
end