| author | wenzelm | 
| Mon, 18 Nov 2024 15:05:31 +0100 | |
| changeset 81483 | 7d4df25af572 | 
| parent 81142 | 6ad2c917dd2e | 
| permissions | -rw-r--r-- | 
| 35849 | 1  | 
(* Title: HOL/Algebra/Lattice.thy  | 
2  | 
Author: Clemens Ballarin, started 7 November 2003  | 
|
3  | 
Copyright: Clemens Ballarin  | 
|
| 
27714
 
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
 
ballarin 
parents: 
27713 
diff
changeset
 | 
4  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
5  | 
Most congruence rules by Stephan Hohe.  | 
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
6  | 
With additional contributions from Alasdair Armstrong and Simon Foster.  | 
| 14551 | 7  | 
*)  | 
8  | 
||
| 35849 | 9  | 
theory Lattice  | 
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
10  | 
imports Order  | 
| 44472 | 11  | 
begin  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
12  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
13  | 
section \<open>Lattices\<close>  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
14  | 
|
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
15  | 
subsection \<open>Supremum and infimum\<close>  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
16  | 
|
| 35847 | 17  | 
definition  | 
| 81142 | 18  | 
sup :: "[_, 'a set] => 'a" (\<open>(\<open>open_block notation=\<open>prefix \<Squnion>\<close>\<close>\<Squnion>\<index>_)\<close> [90] 90)  | 
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35847 
diff
changeset
 | 
19  | 
where "\<Squnion>\<^bsub>L\<^esub>A = (SOME x. least L x (Upper L A))"  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
20  | 
|
| 35847 | 21  | 
definition  | 
| 81142 | 22  | 
inf :: "[_, 'a set] => 'a" (\<open>(\<open>open_block notation=\<open>prefix \<Sqinter>\<close>\<close>\<Sqinter>\<index>_)\<close> [90] 90)  | 
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35847 
diff
changeset
 | 
23  | 
where "\<Sqinter>\<^bsub>L\<^esub>A = (SOME x. greatest L x (Lower L A))"  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
24  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
25  | 
definition supr ::  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
26  | 
  "('a, 'b) gorder_scheme \<Rightarrow> 'c set \<Rightarrow> ('c \<Rightarrow> 'a) \<Rightarrow> 'a "
 | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
27  | 
where "supr L A f = \<Squnion>\<^bsub>L\<^esub>(f ` A)"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
28  | 
|
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
29  | 
definition infi ::  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
30  | 
  "('a, 'b) gorder_scheme \<Rightarrow> 'c set \<Rightarrow> ('c \<Rightarrow> 'a) \<Rightarrow> 'a "
 | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
31  | 
where "infi L A f = \<Sqinter>\<^bsub>L\<^esub>(f ` A)"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
32  | 
|
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
33  | 
syntax  | 
| 81142 | 34  | 
  "_inf1"     :: "('a, 'b) gorder_scheme \<Rightarrow> pttrns \<Rightarrow> 'a \<Rightarrow> 'a"
 | 
35  | 
(\<open>(\<open>indent=3 notation=\<open>binder IINF\<close>\<close>IINF\<index> _./ _)\<close> [0, 10] 10)  | 
|
36  | 
  "_inf"      :: "('a, 'b) gorder_scheme \<Rightarrow> pttrn \<Rightarrow> 'c set \<Rightarrow> 'a \<Rightarrow> 'a"
 | 
|
37  | 
(\<open>(\<open>indent=3 notation=\<open>binder IINF\<close>\<close>IINF\<index> _:_./ _)\<close> [0, 0, 10] 10)  | 
|
38  | 
  "_sup1"     :: "('a, 'b) gorder_scheme \<Rightarrow> pttrns \<Rightarrow> 'a \<Rightarrow> 'a"
 | 
|
39  | 
(\<open>(\<open>indent=3 notation=\<open>binder SSUP\<close>\<close>SSUP\<index> _./ _)\<close> [0, 10] 10)  | 
|
40  | 
  "_sup"      :: "('a, 'b) gorder_scheme \<Rightarrow> pttrn \<Rightarrow> 'c set \<Rightarrow> 'a \<Rightarrow> 'a"
 | 
|
41  | 
(\<open>(\<open>indent=3 notation=\<open>binder SSUP\<close>\<close>SSUP\<index> _:_./ _)\<close> [0, 0, 10] 10)  | 
|
| 80768 | 42  | 
syntax_consts  | 
43  | 
"_inf1" "_inf" == infi and  | 
|
44  | 
"_sup1" "_sup" == supr  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
45  | 
translations  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
46  | 
"IINF\<^bsub>L\<^esub> x. B" == "CONST infi L CONST UNIV (%x. B)"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
47  | 
"IINF\<^bsub>L\<^esub> x:A. B" == "CONST infi L A (%x. B)"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
48  | 
"SSUP\<^bsub>L\<^esub> x. B" == "CONST supr L CONST UNIV (%x. B)"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
49  | 
"SSUP\<^bsub>L\<^esub> x:A. B" == "CONST supr L A (%x. B)"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
50  | 
|
| 35847 | 51  | 
definition  | 
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
80768 
diff
changeset
 | 
52  | 
join :: "[_, 'a, 'a] => 'a" (infixl \<open>\<squnion>\<index>\<close> 65)  | 
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35847 
diff
changeset
 | 
53  | 
  where "x \<squnion>\<^bsub>L\<^esub> y = \<Squnion>\<^bsub>L\<^esub>{x, y}"
 | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
54  | 
|
| 35847 | 55  | 
definition  | 
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
80768 
diff
changeset
 | 
56  | 
meet :: "[_, 'a, 'a] => 'a" (infixl \<open>\<sqinter>\<index>\<close> 70)  | 
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35847 
diff
changeset
 | 
57  | 
  where "x \<sqinter>\<^bsub>L\<^esub> y = \<Sqinter>\<^bsub>L\<^esub>{x, y}"
 | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
58  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
59  | 
definition  | 
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
80768 
diff
changeset
 | 
60  | 
  LEAST_FP :: "('a, 'b) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a" (\<open>LFP\<index>\<close>) where
 | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67226 
diff
changeset
 | 
61  | 
  "LEAST_FP L f = \<Sqinter>\<^bsub>L\<^esub> {u \<in> carrier L. f u \<sqsubseteq>\<^bsub>L\<^esub> u}"    \<comment> \<open>least fixed point\<close>
 | 
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
62  | 
|
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
63  | 
definition  | 
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
80768 
diff
changeset
 | 
64  | 
  GREATEST_FP:: "('a, 'b) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a" (\<open>GFP\<index>\<close>) where
 | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67226 
diff
changeset
 | 
65  | 
  "GREATEST_FP L f = \<Squnion>\<^bsub>L\<^esub> {u \<in> carrier L. u \<sqsubseteq>\<^bsub>L\<^esub> f u}"    \<comment> \<open>greatest fixed point\<close>
 | 
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
66  | 
|
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
67  | 
|
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
68  | 
subsection \<open>Dual operators\<close>  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
69  | 
|
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
70  | 
lemma sup_dual [simp]:  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
71  | 
"\<Squnion>\<^bsub>inv_gorder L\<^esub>A = \<Sqinter>\<^bsub>L\<^esub>A"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
72  | 
by (simp add: sup_def inf_def)  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
73  | 
|
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
74  | 
lemma inf_dual [simp]:  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
75  | 
"\<Sqinter>\<^bsub>inv_gorder L\<^esub>A = \<Squnion>\<^bsub>L\<^esub>A"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
76  | 
by (simp add: sup_def inf_def)  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
77  | 
|
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
78  | 
lemma join_dual [simp]:  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
79  | 
"p \<squnion>\<^bsub>inv_gorder L\<^esub> q = p \<sqinter>\<^bsub>L\<^esub> q"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
80  | 
by (simp add:join_def meet_def)  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
81  | 
|
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
82  | 
lemma meet_dual [simp]:  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
83  | 
"p \<sqinter>\<^bsub>inv_gorder L\<^esub> q = p \<squnion>\<^bsub>L\<^esub> q"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
84  | 
by (simp add:join_def meet_def)  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
85  | 
|
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
86  | 
lemma top_dual [simp]:  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
87  | 
"\<top>\<^bsub>inv_gorder L\<^esub> = \<bottom>\<^bsub>L\<^esub>"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
88  | 
by (simp add: top_def bottom_def)  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
89  | 
|
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
90  | 
lemma bottom_dual [simp]:  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
91  | 
"\<bottom>\<^bsub>inv_gorder L\<^esub> = \<top>\<^bsub>L\<^esub>"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
92  | 
by (simp add: top_def bottom_def)  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
93  | 
|
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
94  | 
lemma LFP_dual [simp]:  | 
| 
66580
 
e5b1d4d55bf6
Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
 
ballarin 
parents: 
65099 
diff
changeset
 | 
95  | 
"LEAST_FP (inv_gorder L) f = GREATEST_FP L f"  | 
| 
 
e5b1d4d55bf6
Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
 
ballarin 
parents: 
65099 
diff
changeset
 | 
96  | 
by (simp add:LEAST_FP_def GREATEST_FP_def)  | 
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
97  | 
|
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
98  | 
lemma GFP_dual [simp]:  | 
| 
66580
 
e5b1d4d55bf6
Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
 
ballarin 
parents: 
65099 
diff
changeset
 | 
99  | 
"GREATEST_FP (inv_gorder L) f = LEAST_FP L f"  | 
| 
 
e5b1d4d55bf6
Avoid \mu and \nu as constant syntax, use LFP and GFP instead.
 
ballarin 
parents: 
65099 
diff
changeset
 | 
100  | 
by (simp add:LEAST_FP_def GREATEST_FP_def)  | 
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
101  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
102  | 
|
| 61382 | 103  | 
subsection \<open>Lattices\<close>  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
104  | 
|
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
105  | 
locale weak_upper_semilattice = weak_partial_order +  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
106  | 
assumes sup_of_two_exists:  | 
| 67091 | 107  | 
    "[| x \<in> carrier L; y \<in> carrier L |] ==> \<exists>s. least L s (Upper L {x, y})"
 | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
108  | 
|
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
109  | 
locale weak_lower_semilattice = weak_partial_order +  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
110  | 
assumes inf_of_two_exists:  | 
| 67091 | 111  | 
    "[| x \<in> carrier L; y \<in> carrier L |] ==> \<exists>s. greatest L s (Lower L {x, y})"
 | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
112  | 
|
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
113  | 
locale weak_lattice = weak_upper_semilattice + weak_lower_semilattice  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
114  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
115  | 
lemma (in weak_lattice) dual_weak_lattice:  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
116  | 
"weak_lattice (inv_gorder L)"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
117  | 
proof -  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
118  | 
interpret dual: weak_partial_order "inv_gorder L"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
119  | 
by (metis dual_weak_order)  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
120  | 
show ?thesis  | 
| 77362 | 121  | 
proof qed (simp_all add: inf_of_two_exists sup_of_two_exists)  | 
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
122  | 
qed  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
123  | 
|
| 14666 | 124  | 
|
| 61382 | 125  | 
subsubsection \<open>Supremum\<close>  | 
| 14551 | 126  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
127  | 
lemma (in weak_upper_semilattice) joinI:  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
128  | 
  "[| !!l. least L l (Upper L {x, y}) ==> P l; x \<in> carrier L; y \<in> carrier L |]
 | 
| 14551 | 129  | 
==> P (x \<squnion> y)"  | 
130  | 
proof (unfold join_def sup_def)  | 
|
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
131  | 
assume L: "x \<in> carrier L" "y \<in> carrier L"  | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
132  | 
    and P: "!!l. least L l (Upper L {x, y}) ==> P l"
 | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
133  | 
  with sup_of_two_exists obtain s where "least L s (Upper L {x, y})" by fast
 | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
134  | 
  with L show "P (SOME l. least L l (Upper L {x, y}))"
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
135  | 
by (fast intro: someI2 P)  | 
| 14551 | 136  | 
qed  | 
137  | 
||
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
138  | 
lemma (in weak_upper_semilattice) join_closed [simp]:  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
139  | 
"[| x \<in> carrier L; y \<in> carrier L |] ==> x \<squnion> y \<in> carrier L"  | 
| 27700 | 140  | 
by (rule joinI) (rule least_closed)  | 
| 14551 | 141  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
142  | 
lemma (in weak_upper_semilattice) join_cong_l:  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
143  | 
assumes carr: "x \<in> carrier L" "x' \<in> carrier L" "y \<in> carrier L"  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
144  | 
and xx': "x .= x'"  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
145  | 
shows "x \<squnion> y .= x' \<squnion> y"  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
146  | 
proof (rule joinI, rule joinI)  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
147  | 
fix a b  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
148  | 
from xx' carr  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
149  | 
      have seq: "{x, y} {.=} {x', y}" by (rule set_eq_pairI)
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
150  | 
|
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
151  | 
  assume leasta: "least L a (Upper L {x, y})"
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
152  | 
  assume "least L b (Upper L {x', y})"
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
153  | 
with carr  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
154  | 
      have leastb: "least L b (Upper L {x, y})"
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
155  | 
by (simp add: least_Upper_cong_r[OF _ _ seq])  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
156  | 
|
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
157  | 
from leasta leastb  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
158  | 
show "a .= b" by (rule weak_least_unique)  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
159  | 
qed (rule carr)+  | 
| 14551 | 160  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
161  | 
lemma (in weak_upper_semilattice) join_cong_r:  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
162  | 
assumes carr: "x \<in> carrier L" "y \<in> carrier L" "y' \<in> carrier L"  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
163  | 
and yy': "y .= y'"  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
164  | 
shows "x \<squnion> y .= x \<squnion> y'"  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
165  | 
proof (rule joinI, rule joinI)  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
166  | 
fix a b  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
167  | 
  have "{x, y} = {y, x}" by fast
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
168  | 
also from carr yy'  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
169  | 
      have "{y, x} {.=} {y', x}" by (intro set_eq_pairI)
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
170  | 
  also have "{y', x} = {x, y'}" by fast
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
171  | 
finally  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
172  | 
      have seq: "{x, y} {.=} {x, y'}" .
 | 
| 14551 | 173  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
174  | 
  assume leasta: "least L a (Upper L {x, y})"
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
175  | 
  assume "least L b (Upper L {x, y'})"
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
176  | 
with carr  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
177  | 
      have leastb: "least L b (Upper L {x, y})"
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
178  | 
by (simp add: least_Upper_cong_r[OF _ _ seq])  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
179  | 
|
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
180  | 
from leasta leastb  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
181  | 
show "a .= b" by (rule weak_least_unique)  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
182  | 
qed (rule carr)+  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
183  | 
|
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
184  | 
lemma (in weak_partial_order) sup_of_singletonI: (* only reflexivity needed ? *)  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
185  | 
  "x \<in> carrier L ==> least L x (Upper L {x})"
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
186  | 
by (rule least_UpperI) auto  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
187  | 
|
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
188  | 
lemma (in weak_partial_order) weak_sup_of_singleton [simp]:  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
189  | 
  "x \<in> carrier L ==> \<Squnion>{x} .= x"
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
190  | 
unfolding sup_def  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
191  | 
by (rule someI2) (auto intro: weak_least_unique sup_of_singletonI)  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
192  | 
|
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
193  | 
lemma (in weak_partial_order) sup_of_singleton_closed [simp]:  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
194  | 
  "x \<in> carrier L \<Longrightarrow> \<Squnion>{x} \<in> carrier L"
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
195  | 
unfolding sup_def  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
196  | 
by (rule someI2) (auto intro: sup_of_singletonI)  | 
| 14666 | 197  | 
|
| 63167 | 198  | 
text \<open>Condition on \<open>A\<close>: supremum exists.\<close>  | 
| 14551 | 199  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
200  | 
lemma (in weak_upper_semilattice) sup_insertI:  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
201  | 
"[| !!s. least L s (Upper L (insert x A)) ==> P s;  | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
202  | 
least L a (Upper L A); x \<in> carrier L; A \<subseteq> carrier L |]  | 
| 14693 | 203  | 
==> P (\<Squnion>(insert x A))"  | 
| 14551 | 204  | 
proof (unfold sup_def)  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
205  | 
assume L: "x \<in> carrier L" "A \<subseteq> carrier L"  | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
206  | 
and P: "!!l. least L l (Upper L (insert x A)) ==> P l"  | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
207  | 
and least_a: "least L a (Upper L A)"  | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
208  | 
from L least_a have La: "a \<in> carrier L" by simp  | 
| 14551 | 209  | 
from L sup_of_two_exists least_a  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
210  | 
  obtain s where least_s: "least L s (Upper L {a, x})" by blast
 | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
211  | 
show "P (SOME l. least L l (Upper L (insert x A)))"  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
212  | 
proof (rule someI2)  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
213  | 
show "least L s (Upper L (insert x A))"  | 
| 14551 | 214  | 
proof (rule least_UpperI)  | 
215  | 
fix z  | 
|
| 14693 | 216  | 
assume "z \<in> insert x A"  | 
217  | 
then show "z \<sqsubseteq> s"  | 
|
218  | 
proof  | 
|
219  | 
assume "z = x" then show ?thesis  | 
|
220  | 
by (simp add: least_Upper_above [OF least_s] L La)  | 
|
221  | 
next  | 
|
222  | 
assume "z \<in> A"  | 
|
223  | 
with L least_s least_a show ?thesis  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
224  | 
by (rule_tac le_trans [where y = a]) (auto dest: least_Upper_above)  | 
| 14693 | 225  | 
qed  | 
226  | 
next  | 
|
227  | 
fix y  | 
|
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
228  | 
assume y: "y \<in> Upper L (insert x A)"  | 
| 14693 | 229  | 
show "s \<sqsubseteq> y"  | 
230  | 
proof (rule least_le [OF least_s], rule Upper_memI)  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30363 
diff
changeset
 | 
231  | 
fix z  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30363 
diff
changeset
 | 
232  | 
        assume z: "z \<in> {a, x}"
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30363 
diff
changeset
 | 
233  | 
then show "z \<sqsubseteq> y"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30363 
diff
changeset
 | 
234  | 
proof  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
235  | 
have y': "y \<in> Upper L A"  | 
| 77362 | 236  | 
by (meson Upper_antimono in_mono subset_insertI y)  | 
| 14693 | 237  | 
assume "z = a"  | 
238  | 
with y' least_a show ?thesis by (fast dest: least_le)  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30363 
diff
changeset
 | 
239  | 
next  | 
| 77362 | 240  | 
          assume "z \<in> {x}"
 | 
| 14693 | 241  | 
with y L show ?thesis by blast  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30363 
diff
changeset
 | 
242  | 
qed  | 
| 23350 | 243  | 
qed (rule Upper_closed [THEN subsetD, OF y])  | 
| 14693 | 244  | 
next  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
245  | 
from L show "insert x A \<subseteq> carrier L" by simp  | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
246  | 
from least_s show "s \<in> carrier L" by simp  | 
| 14551 | 247  | 
qed  | 
| 23350 | 248  | 
qed (rule P)  | 
| 14551 | 249  | 
qed  | 
250  | 
||
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
251  | 
lemma (in weak_upper_semilattice) finite_sup_least:  | 
| 67091 | 252  | 
  "[| finite A; A \<subseteq> carrier L; A \<noteq> {} |] ==> least L (\<Squnion>A) (Upper L A)"
 | 
| 22265 | 253  | 
proof (induct set: finite)  | 
| 14693 | 254  | 
case empty  | 
255  | 
then show ?case by simp  | 
|
| 14551 | 256  | 
next  | 
| 15328 | 257  | 
case (insert x A)  | 
| 14551 | 258  | 
show ?case  | 
259  | 
  proof (cases "A = {}")
 | 
|
260  | 
case True  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
261  | 
with insert show ?thesis  | 
| 44472 | 262  | 
by simp (simp add: least_cong [OF weak_sup_of_singleton] sup_of_singletonI)  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30363 
diff
changeset
 | 
263  | 
(* The above step is hairy; least_cong can make simp loop.  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30363 
diff
changeset
 | 
264  | 
Would want special version of simp to apply least_cong. *)  | 
| 14551 | 265  | 
next  | 
266  | 
case False  | 
|
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
267  | 
with insert have "least L (\<Squnion>A) (Upper L A)" by simp  | 
| 14693 | 268  | 
with _ show ?thesis  | 
269  | 
by (rule sup_insertI) (simp_all add: insert [simplified])  | 
|
| 14551 | 270  | 
qed  | 
271  | 
qed  | 
|
272  | 
||
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
273  | 
lemma (in weak_upper_semilattice) finite_sup_insertI:  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
274  | 
assumes P: "!!l. least L l (Upper L (insert x A)) ==> P l"  | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
275  | 
and xA: "finite A" "x \<in> carrier L" "A \<subseteq> carrier L"  | 
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
276  | 
shows "P (\<Squnion> (insert x A))"  | 
| 14551 | 277  | 
proof (cases "A = {}")
 | 
278  | 
case True with P and xA show ?thesis  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
279  | 
by (simp add: finite_sup_least)  | 
| 14551 | 280  | 
next  | 
281  | 
case False with P and xA show ?thesis  | 
|
282  | 
by (simp add: sup_insertI finite_sup_least)  | 
|
283  | 
qed  | 
|
284  | 
||
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
285  | 
lemma (in weak_upper_semilattice) finite_sup_closed [simp]:  | 
| 67091 | 286  | 
  "[| finite A; A \<subseteq> carrier L; A \<noteq> {} |] ==> \<Squnion>A \<in> carrier L"
 | 
| 22265 | 287  | 
proof (induct set: finite)  | 
| 14551 | 288  | 
case empty then show ?case by simp  | 
289  | 
next  | 
|
| 15328 | 290  | 
case insert then show ?case  | 
| 14693 | 291  | 
by - (rule finite_sup_insertI, simp_all)  | 
| 14551 | 292  | 
qed  | 
293  | 
||
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
294  | 
lemma (in weak_upper_semilattice) join_left:  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
295  | 
"[| x \<in> carrier L; y \<in> carrier L |] ==> x \<sqsubseteq> x \<squnion> y"  | 
| 14693 | 296  | 
by (rule joinI [folded join_def]) (blast dest: least_mem)  | 
| 14551 | 297  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
298  | 
lemma (in weak_upper_semilattice) join_right:  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
299  | 
"[| x \<in> carrier L; y \<in> carrier L |] ==> y \<sqsubseteq> x \<squnion> y"  | 
| 14693 | 300  | 
by (rule joinI [folded join_def]) (blast dest: least_mem)  | 
| 14551 | 301  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
302  | 
lemma (in weak_upper_semilattice) sup_of_two_least:  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
303  | 
  "[| x \<in> carrier L; y \<in> carrier L |] ==> least L (\<Squnion>{x, y}) (Upper L {x, y})"
 | 
| 14551 | 304  | 
proof (unfold sup_def)  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
305  | 
assume L: "x \<in> carrier L" "y \<in> carrier L"  | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
306  | 
  with sup_of_two_exists obtain s where "least L s (Upper L {x, y})" by fast
 | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
307  | 
  with L show "least L (SOME z. least L z (Upper L {x, y})) (Upper L {x, y})"
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
308  | 
by (fast intro: someI2 weak_least_unique) (* blast fails *)  | 
| 14551 | 309  | 
qed  | 
310  | 
||
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
311  | 
lemma (in weak_upper_semilattice) join_le:  | 
| 14693 | 312  | 
assumes sub: "x \<sqsubseteq> z" "y \<sqsubseteq> z"  | 
| 23350 | 313  | 
and x: "x \<in> carrier L" and y: "y \<in> carrier L" and z: "z \<in> carrier L"  | 
| 14551 | 314  | 
shows "x \<squnion> y \<sqsubseteq> z"  | 
| 23350 | 315  | 
proof (rule joinI [OF _ x y])  | 
| 14551 | 316  | 
fix s  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
317  | 
  assume "least L s (Upper L {x, y})"
 | 
| 23350 | 318  | 
with sub z show "s \<sqsubseteq> z" by (fast elim: least_le intro: Upper_memI)  | 
| 14551 | 319  | 
qed  | 
| 14693 | 320  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
321  | 
lemma (in weak_lattice) weak_le_iff_meet:  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
322  | 
assumes "x \<in> carrier L" "y \<in> carrier L"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
323  | 
shows "x \<sqsubseteq> y \<longleftrightarrow> (x \<squnion> y) .= y"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
324  | 
by (meson assms(1) assms(2) join_closed join_le join_left join_right le_cong_r local.le_refl weak_le_antisym)  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
325  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
326  | 
lemma (in weak_upper_semilattice) weak_join_assoc_lemma:  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
327  | 
assumes L: "x \<in> carrier L" "y \<in> carrier L" "z \<in> carrier L"  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
328  | 
  shows "x \<squnion> (y \<squnion> z) .= \<Squnion>{x, y, z}"
 | 
| 14551 | 329  | 
proof (rule finite_sup_insertI)  | 
| 63167 | 330  | 
\<comment> \<open>The textbook argument in Jacobson I, p 457\<close>  | 
| 14551 | 331  | 
fix s  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
332  | 
  assume sup: "least L s (Upper L {x, y, z})"
 | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
333  | 
show "x \<squnion> (y \<squnion> z) .= s"  | 
| 33657 | 334  | 
proof (rule weak_le_antisym)  | 
| 14551 | 335  | 
from sup L show "x \<squnion> (y \<squnion> z) \<sqsubseteq> s"  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44655 
diff
changeset
 | 
336  | 
by (fastforce intro!: join_le elim: least_Upper_above)  | 
| 14551 | 337  | 
next  | 
338  | 
from sup L show "s \<sqsubseteq> x \<squnion> (y \<squnion> z)"  | 
|
339  | 
by (erule_tac least_le)  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
340  | 
(blast intro!: Upper_memI intro: le_trans join_left join_right join_closed)  | 
| 27700 | 341  | 
qed (simp_all add: L least_closed [OF sup])  | 
| 14551 | 342  | 
qed (simp_all add: L)  | 
343  | 
||
| 63167 | 344  | 
text \<open>Commutativity holds for \<open>=\<close>.\<close>  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
345  | 
|
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
346  | 
lemma join_comm:  | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
347  | 
fixes L (structure)  | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
348  | 
shows "x \<squnion> y = y \<squnion> x"  | 
| 14551 | 349  | 
by (unfold join_def) (simp add: insert_commute)  | 
350  | 
||
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
351  | 
lemma (in weak_upper_semilattice) weak_join_assoc:  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
352  | 
assumes L: "x \<in> carrier L" "y \<in> carrier L" "z \<in> carrier L"  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
353  | 
shows "(x \<squnion> y) \<squnion> z .= x \<squnion> (y \<squnion> z)"  | 
| 14551 | 354  | 
proof -  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
355  | 
(* FIXME: could be simplified by improved simp: uniform use of .=,  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
356  | 
omit [symmetric] in last step. *)  | 
| 14551 | 357  | 
have "(x \<squnion> y) \<squnion> z = z \<squnion> (x \<squnion> y)" by (simp only: join_comm)  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
358  | 
  also from L have "... .= \<Squnion>{z, x, y}" by (simp add: weak_join_assoc_lemma)
 | 
| 14693 | 359  | 
  also from L have "... = \<Squnion>{x, y, z}" by (simp add: insert_commute)
 | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
360  | 
also from L have "... .= x \<squnion> (y \<squnion> z)" by (simp add: weak_join_assoc_lemma [symmetric])  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
361  | 
finally show ?thesis by (simp add: L)  | 
| 14551 | 362  | 
qed  | 
363  | 
||
| 14693 | 364  | 
|
| 61382 | 365  | 
subsubsection \<open>Infimum\<close>  | 
| 14551 | 366  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
367  | 
lemma (in weak_lower_semilattice) meetI:  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
368  | 
  "[| !!i. greatest L i (Lower L {x, y}) ==> P i;
 | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
369  | 
x \<in> carrier L; y \<in> carrier L |]  | 
| 14551 | 370  | 
==> P (x \<sqinter> y)"  | 
371  | 
proof (unfold meet_def inf_def)  | 
|
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
372  | 
assume L: "x \<in> carrier L" "y \<in> carrier L"  | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
373  | 
    and P: "!!g. greatest L g (Lower L {x, y}) ==> P g"
 | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
374  | 
  with inf_of_two_exists obtain i where "greatest L i (Lower L {x, y})" by fast
 | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
375  | 
  with L show "P (SOME g. greatest L g (Lower L {x, y}))"
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
376  | 
by (fast intro: someI2 weak_greatest_unique P)  | 
| 14551 | 377  | 
qed  | 
378  | 
||
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
379  | 
lemma (in weak_lower_semilattice) meet_closed [simp]:  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
380  | 
"[| x \<in> carrier L; y \<in> carrier L |] ==> x \<sqinter> y \<in> carrier L"  | 
| 27700 | 381  | 
by (rule meetI) (rule greatest_closed)  | 
| 14551 | 382  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
383  | 
lemma (in weak_lower_semilattice) meet_cong_l:  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
384  | 
assumes carr: "x \<in> carrier L" "x' \<in> carrier L" "y \<in> carrier L"  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
385  | 
and xx': "x .= x'"  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
386  | 
shows "x \<sqinter> y .= x' \<sqinter> y"  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
387  | 
proof (rule meetI, rule meetI)  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
388  | 
fix a b  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
389  | 
from xx' carr  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
390  | 
      have seq: "{x, y} {.=} {x', y}" by (rule set_eq_pairI)
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
391  | 
|
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
392  | 
  assume greatesta: "greatest L a (Lower L {x, y})"
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
393  | 
  assume "greatest L b (Lower L {x', y})"
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
394  | 
with carr  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
395  | 
      have greatestb: "greatest L b (Lower L {x, y})"
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
396  | 
by (simp add: greatest_Lower_cong_r[OF _ _ seq])  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
397  | 
|
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
398  | 
from greatesta greatestb  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
399  | 
show "a .= b" by (rule weak_greatest_unique)  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
400  | 
qed (rule carr)+  | 
| 14551 | 401  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
402  | 
lemma (in weak_lower_semilattice) meet_cong_r:  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
403  | 
assumes carr: "x \<in> carrier L" "y \<in> carrier L" "y' \<in> carrier L"  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
404  | 
and yy': "y .= y'"  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
405  | 
shows "x \<sqinter> y .= x \<sqinter> y'"  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
406  | 
proof (rule meetI, rule meetI)  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
407  | 
fix a b  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
408  | 
  have "{x, y} = {y, x}" by fast
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
409  | 
also from carr yy'  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
410  | 
      have "{y, x} {.=} {y', x}" by (intro set_eq_pairI)
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
411  | 
  also have "{y', x} = {x, y'}" by fast
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
412  | 
finally  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
413  | 
      have seq: "{x, y} {.=} {x, y'}" .
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
414  | 
|
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
415  | 
  assume greatesta: "greatest L a (Lower L {x, y})"
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
416  | 
  assume "greatest L b (Lower L {x, y'})"
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
417  | 
with carr  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
418  | 
      have greatestb: "greatest L b (Lower L {x, y})"
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
419  | 
by (simp add: greatest_Lower_cong_r[OF _ _ seq])  | 
| 14551 | 420  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
421  | 
from greatesta greatestb  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
422  | 
show "a .= b" by (rule weak_greatest_unique)  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
423  | 
qed (rule carr)+  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
424  | 
|
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
425  | 
lemma (in weak_partial_order) inf_of_singletonI: (* only reflexivity needed ? *)  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
426  | 
  "x \<in> carrier L ==> greatest L x (Lower L {x})"
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
427  | 
by (rule greatest_LowerI) auto  | 
| 14551 | 428  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
429  | 
lemma (in weak_partial_order) weak_inf_of_singleton [simp]:  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
430  | 
  "x \<in> carrier L ==> \<Sqinter>{x} .= x"
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
431  | 
unfolding inf_def  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
432  | 
by (rule someI2) (auto intro: weak_greatest_unique inf_of_singletonI)  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
433  | 
|
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
434  | 
lemma (in weak_partial_order) inf_of_singleton_closed:  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
435  | 
  "x \<in> carrier L ==> \<Sqinter>{x} \<in> carrier L"
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
436  | 
unfolding inf_def  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
437  | 
by (rule someI2) (auto intro: inf_of_singletonI)  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
438  | 
|
| 63167 | 439  | 
text \<open>Condition on \<open>A\<close>: infimum exists.\<close>  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
440  | 
|
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
441  | 
lemma (in weak_lower_semilattice) inf_insertI:  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
442  | 
"[| !!i. greatest L i (Lower L (insert x A)) ==> P i;  | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
443  | 
greatest L a (Lower L A); x \<in> carrier L; A \<subseteq> carrier L |]  | 
| 14693 | 444  | 
==> P (\<Sqinter>(insert x A))"  | 
| 14551 | 445  | 
proof (unfold inf_def)  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
446  | 
assume L: "x \<in> carrier L" "A \<subseteq> carrier L"  | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
447  | 
and P: "!!g. greatest L g (Lower L (insert x A)) ==> P g"  | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
448  | 
and greatest_a: "greatest L a (Lower L A)"  | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
449  | 
from L greatest_a have La: "a \<in> carrier L" by simp  | 
| 14551 | 450  | 
from L inf_of_two_exists greatest_a  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
451  | 
  obtain i where greatest_i: "greatest L i (Lower L {a, x})" by blast
 | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
452  | 
show "P (SOME g. greatest L g (Lower L (insert x A)))"  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
453  | 
proof (rule someI2)  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
454  | 
show "greatest L i (Lower L (insert x A))"  | 
| 14551 | 455  | 
proof (rule greatest_LowerI)  | 
456  | 
fix z  | 
|
| 14693 | 457  | 
assume "z \<in> insert x A"  | 
458  | 
then show "i \<sqsubseteq> z"  | 
|
459  | 
proof  | 
|
460  | 
assume "z = x" then show ?thesis  | 
|
| 27700 | 461  | 
by (simp add: greatest_Lower_below [OF greatest_i] L La)  | 
| 14693 | 462  | 
next  | 
463  | 
assume "z \<in> A"  | 
|
464  | 
with L greatest_i greatest_a show ?thesis  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
465  | 
by (rule_tac le_trans [where y = a]) (auto dest: greatest_Lower_below)  | 
| 14693 | 466  | 
qed  | 
467  | 
next  | 
|
468  | 
fix y  | 
|
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
469  | 
assume y: "y \<in> Lower L (insert x A)"  | 
| 14693 | 470  | 
show "y \<sqsubseteq> i"  | 
471  | 
proof (rule greatest_le [OF greatest_i], rule Lower_memI)  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30363 
diff
changeset
 | 
472  | 
fix z  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30363 
diff
changeset
 | 
473  | 
        assume z: "z \<in> {a, x}"
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30363 
diff
changeset
 | 
474  | 
then show "y \<sqsubseteq> z"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30363 
diff
changeset
 | 
475  | 
proof  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
476  | 
have y': "y \<in> Lower L A"  | 
| 77362 | 477  | 
by (meson Lower_antimono in_mono subset_insertI y)  | 
| 14693 | 478  | 
assume "z = a"  | 
479  | 
with y' greatest_a show ?thesis by (fast dest: greatest_le)  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30363 
diff
changeset
 | 
480  | 
next  | 
| 14693 | 481  | 
          assume "z \<in> {x}"
 | 
482  | 
with y L show ?thesis by blast  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30363 
diff
changeset
 | 
483  | 
qed  | 
| 23350 | 484  | 
qed (rule Lower_closed [THEN subsetD, OF y])  | 
| 14693 | 485  | 
next  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
486  | 
from L show "insert x A \<subseteq> carrier L" by simp  | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
487  | 
from greatest_i show "i \<in> carrier L" by simp  | 
| 14551 | 488  | 
qed  | 
| 23350 | 489  | 
qed (rule P)  | 
| 14551 | 490  | 
qed  | 
491  | 
||
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
492  | 
lemma (in weak_lower_semilattice) finite_inf_greatest:  | 
| 67091 | 493  | 
  "[| finite A; A \<subseteq> carrier L; A \<noteq> {} |] ==> greatest L (\<Sqinter>A) (Lower L A)"
 | 
| 22265 | 494  | 
proof (induct set: finite)  | 
| 14551 | 495  | 
case empty then show ?case by simp  | 
496  | 
next  | 
|
| 15328 | 497  | 
case (insert x A)  | 
| 14551 | 498  | 
show ?case  | 
499  | 
  proof (cases "A = {}")
 | 
|
500  | 
case True  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
501  | 
with insert show ?thesis  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
502  | 
by simp (simp add: greatest_cong [OF weak_inf_of_singleton]  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30363 
diff
changeset
 | 
503  | 
inf_of_singleton_closed inf_of_singletonI)  | 
| 14551 | 504  | 
next  | 
505  | 
case False  | 
|
506  | 
from insert show ?thesis  | 
|
507  | 
proof (rule_tac inf_insertI)  | 
|
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
508  | 
from False insert show "greatest L (\<Sqinter>A) (Lower L A)" by simp  | 
| 14551 | 509  | 
qed simp_all  | 
510  | 
qed  | 
|
511  | 
qed  | 
|
512  | 
||
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
513  | 
lemma (in weak_lower_semilattice) finite_inf_insertI:  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
514  | 
assumes P: "!!i. greatest L i (Lower L (insert x A)) ==> P i"  | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
515  | 
and xA: "finite A" "x \<in> carrier L" "A \<subseteq> carrier L"  | 
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
516  | 
shows "P (\<Sqinter> (insert x A))"  | 
| 14551 | 517  | 
proof (cases "A = {}")
 | 
518  | 
case True with P and xA show ?thesis  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
519  | 
by (simp add: finite_inf_greatest)  | 
| 14551 | 520  | 
next  | 
521  | 
case False with P and xA show ?thesis  | 
|
522  | 
by (simp add: inf_insertI finite_inf_greatest)  | 
|
523  | 
qed  | 
|
524  | 
||
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
525  | 
lemma (in weak_lower_semilattice) finite_inf_closed [simp]:  | 
| 67091 | 526  | 
  "[| finite A; A \<subseteq> carrier L; A \<noteq> {} |] ==> \<Sqinter>A \<in> carrier L"
 | 
| 22265 | 527  | 
proof (induct set: finite)  | 
| 14551 | 528  | 
case empty then show ?case by simp  | 
529  | 
next  | 
|
| 15328 | 530  | 
case insert then show ?case  | 
| 14551 | 531  | 
by (rule_tac finite_inf_insertI) (simp_all)  | 
532  | 
qed  | 
|
533  | 
||
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
534  | 
lemma (in weak_lower_semilattice) meet_left:  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
535  | 
"[| x \<in> carrier L; y \<in> carrier L |] ==> x \<sqinter> y \<sqsubseteq> x"  | 
| 14693 | 536  | 
by (rule meetI [folded meet_def]) (blast dest: greatest_mem)  | 
| 14551 | 537  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
538  | 
lemma (in weak_lower_semilattice) meet_right:  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
539  | 
"[| x \<in> carrier L; y \<in> carrier L |] ==> x \<sqinter> y \<sqsubseteq> y"  | 
| 14693 | 540  | 
by (rule meetI [folded meet_def]) (blast dest: greatest_mem)  | 
| 14551 | 541  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
542  | 
lemma (in weak_lower_semilattice) inf_of_two_greatest:  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
543  | 
"[| x \<in> carrier L; y \<in> carrier L |] ==>  | 
| 60585 | 544  | 
  greatest L (\<Sqinter>{x, y}) (Lower L {x, y})"
 | 
| 14551 | 545  | 
proof (unfold inf_def)  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
546  | 
assume L: "x \<in> carrier L" "y \<in> carrier L"  | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
547  | 
  with inf_of_two_exists obtain s where "greatest L s (Lower L {x, y})" by fast
 | 
| 14551 | 548  | 
with L  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
549  | 
  show "greatest L (SOME z. greatest L z (Lower L {x, y})) (Lower L {x, y})"
 | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
550  | 
by (fast intro: someI2 weak_greatest_unique) (* blast fails *)  | 
| 14551 | 551  | 
qed  | 
552  | 
||
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
553  | 
lemma (in weak_lower_semilattice) meet_le:  | 
| 14693 | 554  | 
assumes sub: "z \<sqsubseteq> x" "z \<sqsubseteq> y"  | 
| 23350 | 555  | 
and x: "x \<in> carrier L" and y: "y \<in> carrier L" and z: "z \<in> carrier L"  | 
| 14551 | 556  | 
shows "z \<sqsubseteq> x \<sqinter> y"  | 
| 23350 | 557  | 
proof (rule meetI [OF _ x y])  | 
| 14551 | 558  | 
fix i  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
559  | 
  assume "greatest L i (Lower L {x, y})"
 | 
| 23350 | 560  | 
with sub z show "z \<sqsubseteq> i" by (fast elim: greatest_le intro: Lower_memI)  | 
| 14551 | 561  | 
qed  | 
| 14693 | 562  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
563  | 
lemma (in weak_lattice) weak_le_iff_join:  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
564  | 
assumes "x \<in> carrier L" "y \<in> carrier L"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
565  | 
shows "x \<sqsubseteq> y \<longleftrightarrow> x .= (x \<sqinter> y)"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
566  | 
by (meson assms(1) assms(2) local.le_refl local.le_trans meet_closed meet_le meet_left meet_right weak_le_antisym weak_refl)  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
567  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
568  | 
lemma (in weak_lower_semilattice) weak_meet_assoc_lemma:  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
569  | 
assumes L: "x \<in> carrier L" "y \<in> carrier L" "z \<in> carrier L"  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
570  | 
  shows "x \<sqinter> (y \<sqinter> z) .= \<Sqinter>{x, y, z}"
 | 
| 14551 | 571  | 
proof (rule finite_inf_insertI)  | 
| 61382 | 572  | 
txt \<open>The textbook argument in Jacobson I, p 457\<close>  | 
| 14551 | 573  | 
fix i  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
574  | 
  assume inf: "greatest L i (Lower L {x, y, z})"
 | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
575  | 
show "x \<sqinter> (y \<sqinter> z) .= i"  | 
| 33657 | 576  | 
proof (rule weak_le_antisym)  | 
| 14551 | 577  | 
from inf L show "i \<sqsubseteq> x \<sqinter> (y \<sqinter> z)"  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44655 
diff
changeset
 | 
578  | 
by (fastforce intro!: meet_le elim: greatest_Lower_below)  | 
| 14551 | 579  | 
next  | 
580  | 
from inf L show "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> i"  | 
|
581  | 
by (erule_tac greatest_le)  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
582  | 
(blast intro!: Lower_memI intro: le_trans meet_left meet_right meet_closed)  | 
| 27700 | 583  | 
qed (simp_all add: L greatest_closed [OF inf])  | 
| 14551 | 584  | 
qed (simp_all add: L)  | 
585  | 
||
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
586  | 
lemma meet_comm:  | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
587  | 
fixes L (structure)  | 
| 
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
588  | 
shows "x \<sqinter> y = y \<sqinter> x"  | 
| 14551 | 589  | 
by (unfold meet_def) (simp add: insert_commute)  | 
590  | 
||
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
591  | 
lemma (in weak_lower_semilattice) weak_meet_assoc:  | 
| 
22063
 
717425609192
Reverted to structure representation with records.
 
ballarin 
parents: 
21896 
diff
changeset
 | 
592  | 
assumes L: "x \<in> carrier L" "y \<in> carrier L" "z \<in> carrier L"  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
593  | 
shows "(x \<sqinter> y) \<sqinter> z .= x \<sqinter> (y \<sqinter> z)"  | 
| 14551 | 594  | 
proof -  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
595  | 
(* FIXME: improved simp, see weak_join_assoc above *)  | 
| 14551 | 596  | 
have "(x \<sqinter> y) \<sqinter> z = z \<sqinter> (x \<sqinter> y)" by (simp only: meet_comm)  | 
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
597  | 
  also from L have "... .= \<Sqinter> {z, x, y}" by (simp add: weak_meet_assoc_lemma)
 | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
598  | 
  also from L have "... = \<Sqinter> {x, y, z}" by (simp add: insert_commute)
 | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
599  | 
also from L have "... .= x \<sqinter> (y \<sqinter> z)" by (simp add: weak_meet_assoc_lemma [symmetric])  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
600  | 
finally show ?thesis by (simp add: L)  | 
| 14551 | 601  | 
qed  | 
602  | 
||
| 61382 | 603  | 
text \<open>Total orders are lattices.\<close>  | 
| 
24087
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
604  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
605  | 
sublocale weak_total_order \<subseteq> weak?: weak_lattice  | 
| 28823 | 606  | 
proof  | 
| 
24087
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
607  | 
fix x y  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
608  | 
assume L: "x \<in> carrier L" "y \<in> carrier L"  | 
| 67091 | 609  | 
  show "\<exists>s. least L s (Upper L {x, y})"
 | 
| 
24087
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
610  | 
proof -  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
611  | 
note total L  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
612  | 
moreover  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
613  | 
    {
 | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
614  | 
assume "x \<sqsubseteq> y"  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
615  | 
      with L have "least L y (Upper L {x, y})"
 | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
616  | 
by (rule_tac least_UpperI) auto  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
617  | 
}  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
618  | 
moreover  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
619  | 
    {
 | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
620  | 
assume "y \<sqsubseteq> x"  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
621  | 
      with L have "least L x (Upper L {x, y})"
 | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
622  | 
by (rule_tac least_UpperI) auto  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
623  | 
}  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
624  | 
ultimately show ?thesis by blast  | 
| 14551 | 625  | 
qed  | 
| 
24087
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
626  | 
next  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
627  | 
fix x y  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
628  | 
assume L: "x \<in> carrier L" "y \<in> carrier L"  | 
| 67091 | 629  | 
  show "\<exists>i. greatest L i (Lower L {x, y})"
 | 
| 
24087
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
630  | 
proof -  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
631  | 
note total L  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
632  | 
moreover  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
633  | 
    {
 | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
634  | 
assume "y \<sqsubseteq> x"  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
635  | 
      with L have "greatest L y (Lower L {x, y})"
 | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
636  | 
by (rule_tac greatest_LowerI) auto  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
637  | 
}  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
638  | 
moreover  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
639  | 
    {
 | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
640  | 
assume "x \<sqsubseteq> y"  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
641  | 
      with L have "greatest L x (Lower L {x, y})"
 | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
642  | 
by (rule_tac greatest_LowerI) auto  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
643  | 
}  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
644  | 
ultimately show ?thesis by blast  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
645  | 
qed  | 
| 
 
eb025d149a34
Proper interpretation of total orders in lattices.
 
ballarin 
parents: 
23463 
diff
changeset
 | 
646  | 
qed  | 
| 14551 | 647  | 
|
| 14693 | 648  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
649  | 
subsection \<open>Weak Bounded Lattices\<close>  | 
| 14551 | 650  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
651  | 
locale weak_bounded_lattice =  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
652  | 
weak_lattice +  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
653  | 
weak_partial_order_bottom +  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
654  | 
weak_partial_order_top  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
655  | 
begin  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
656  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
657  | 
lemma bottom_meet: "x \<in> carrier L \<Longrightarrow> \<bottom> \<sqinter> x .= \<bottom>"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
658  | 
by (metis bottom_least least_def meet_closed meet_left weak_le_antisym)  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
659  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
660  | 
lemma bottom_join: "x \<in> carrier L \<Longrightarrow> \<bottom> \<squnion> x .= x"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
661  | 
by (metis bottom_least join_closed join_le join_right le_refl least_def weak_le_antisym)  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
662  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
663  | 
lemma bottom_weak_eq:  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
664  | 
"\<lbrakk> b \<in> carrier L; \<And> x. x \<in> carrier L \<Longrightarrow> b \<sqsubseteq> x \<rbrakk> \<Longrightarrow> b .= \<bottom>"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
665  | 
by (metis bottom_closed bottom_lower weak_le_antisym)  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
666  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
667  | 
lemma top_join: "x \<in> carrier L \<Longrightarrow> \<top> \<squnion> x .= \<top>"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
668  | 
by (metis join_closed join_left top_closed top_higher weak_le_antisym)  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
669  | 
|
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
670  | 
lemma top_meet: "x \<in> carrier L \<Longrightarrow> \<top> \<sqinter> x .= x"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
671  | 
by (metis le_refl meet_closed meet_le meet_right top_closed top_higher weak_le_antisym)  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
672  | 
|
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
673  | 
lemma top_weak_eq: "\<lbrakk> t \<in> carrier L; \<And> x. x \<in> carrier L \<Longrightarrow> x \<sqsubseteq> t \<rbrakk> \<Longrightarrow> t .= \<top>"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
674  | 
by (metis top_closed top_higher weak_le_antisym)  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
675  | 
|
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
676  | 
end  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
677  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
678  | 
sublocale weak_bounded_lattice \<subseteq> weak_partial_order ..  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
679  | 
|
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
680  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
681  | 
subsection \<open>Lattices where \<open>eq\<close> is the Equality\<close>  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
682  | 
|
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
683  | 
locale upper_semilattice = partial_order +  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
684  | 
assumes sup_of_two_exists:  | 
| 67091 | 685  | 
    "[| x \<in> carrier L; y \<in> carrier L |] ==> \<exists>s. least L s (Upper L {x, y})"
 | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
686  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
687  | 
sublocale upper_semilattice \<subseteq> weak?: weak_upper_semilattice  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
688  | 
by unfold_locales (rule sup_of_two_exists)  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
689  | 
|
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
690  | 
locale lower_semilattice = partial_order +  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
691  | 
assumes inf_of_two_exists:  | 
| 67091 | 692  | 
    "[| x \<in> carrier L; y \<in> carrier L |] ==> \<exists>s. greatest L s (Lower L {x, y})"
 | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
693  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
694  | 
sublocale lower_semilattice \<subseteq> weak?: weak_lower_semilattice  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
695  | 
by unfold_locales (rule inf_of_two_exists)  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
696  | 
|
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
697  | 
locale lattice = upper_semilattice + lower_semilattice  | 
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
698  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
699  | 
sublocale lattice \<subseteq> weak_lattice ..  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
700  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
701  | 
lemma (in lattice) dual_lattice:  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
702  | 
"lattice (inv_gorder L)"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
703  | 
proof -  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
704  | 
interpret dual: weak_lattice "inv_gorder L"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
705  | 
by (metis dual_weak_lattice)  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
706  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
707  | 
show ?thesis  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
708  | 
apply (unfold_locales)  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
709  | 
apply (simp_all add: inf_of_two_exists sup_of_two_exists)  | 
| 77362 | 710  | 
apply (rule eq_is_equal)  | 
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
711  | 
done  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
712  | 
qed  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
713  | 
|
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
714  | 
lemma (in lattice) le_iff_join:  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
715  | 
assumes "x \<in> carrier L" "y \<in> carrier L"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
716  | 
shows "x \<sqsubseteq> y \<longleftrightarrow> x = (x \<sqinter> y)"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
717  | 
by (simp add: assms(1) assms(2) eq_is_equal weak_le_iff_join)  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
718  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
719  | 
lemma (in lattice) le_iff_meet:  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
720  | 
assumes "x \<in> carrier L" "y \<in> carrier L"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
721  | 
shows "x \<sqsubseteq> y \<longleftrightarrow> (x \<squnion> y) = y"  | 
| 77362 | 722  | 
by (simp add: assms eq_is_equal weak_le_iff_meet)  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
723  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
724  | 
text \<open> Total orders are lattices. \<close>  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
725  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
726  | 
sublocale total_order \<subseteq> weak?: lattice  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
727  | 
by standard (auto intro: weak.weak.sup_of_two_exists weak.weak.inf_of_two_exists)  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
728  | 
|
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
729  | 
text \<open>Functions that preserve joins and meets\<close>  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
730  | 
|
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
731  | 
definition join_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: 
63167 
diff
changeset
 | 
732  | 
"join_pres X Y f \<equiv> lattice X \<and> lattice Y \<and> (\<forall> x \<in> carrier X. \<forall> y \<in> carrier X. f (x \<squnion>\<^bsub>X\<^esub> y) = f x \<squnion>\<^bsub>Y\<^esub> f y)"  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
733  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
734  | 
definition meet_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: 
63167 
diff
changeset
 | 
735  | 
"meet_pres X Y f \<equiv> lattice X \<and> lattice Y \<and> (\<forall> x \<in> carrier X. \<forall> y \<in> carrier X. f (x \<sqinter>\<^bsub>X\<^esub> y) = f x \<sqinter>\<^bsub>Y\<^esub> f y)"  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
736  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
737  | 
lemma join_pres_isotone:  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
738  | 
assumes "f \<in> carrier X \<rightarrow> carrier Y" "join_pres X Y f"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
739  | 
shows "isotone X Y f"  | 
| 77362 | 740  | 
proof (rule isotoneI)  | 
741  | 
show "weak_partial_order X" "weak_partial_order Y"  | 
|
742  | 
using assms unfolding join_pres_def lattice_def upper_semilattice_def lower_semilattice_def  | 
|
743  | 
by (meson partial_order.axioms(1))+  | 
|
744  | 
show "\<And>x y. \<lbrakk>x \<in> carrier X; y \<in> carrier X; x \<sqsubseteq>\<^bsub>X\<^esub> y\<rbrakk> \<Longrightarrow> f x \<sqsubseteq>\<^bsub>Y\<^esub> f y"  | 
|
745  | 
by (metis (no_types, lifting) PiE assms join_pres_def lattice.le_iff_meet)  | 
|
746  | 
qed  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
747  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
748  | 
lemma meet_pres_isotone:  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
749  | 
assumes "f \<in> carrier X \<rightarrow> carrier Y" "meet_pres X Y f"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
750  | 
shows "isotone X Y f"  | 
| 77362 | 751  | 
proof (rule isotoneI)  | 
752  | 
show "weak_partial_order X" "weak_partial_order Y"  | 
|
753  | 
using assms unfolding meet_pres_def lattice_def upper_semilattice_def lower_semilattice_def  | 
|
754  | 
by (meson partial_order.axioms(1))+  | 
|
755  | 
show "\<And>x y. \<lbrakk>x \<in> carrier X; y \<in> carrier X; x \<sqsubseteq>\<^bsub>X\<^esub> y\<rbrakk> \<Longrightarrow> f x \<sqsubseteq>\<^bsub>Y\<^esub> f y"  | 
|
756  | 
by (metis (no_types, lifting) PiE assms lattice.le_iff_join meet_pres_def)  | 
|
757  | 
qed  | 
|
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
758  | 
|
| 
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
759  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
760  | 
subsection \<open>Bounded Lattices\<close>  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
761  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
762  | 
locale bounded_lattice =  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
763  | 
lattice +  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
764  | 
weak_partial_order_bottom +  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
765  | 
weak_partial_order_top  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
766  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
767  | 
sublocale bounded_lattice \<subseteq> weak_bounded_lattice ..  | 
| 
27713
 
95b36bfe7fc4
New locales for orders and lattices where the equivalence relation is not restricted to equality.
 
ballarin 
parents: 
27700 
diff
changeset
 | 
768  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
769  | 
context bounded_lattice  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
770  | 
begin  | 
| 14551 | 771  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
772  | 
lemma bottom_eq:  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
773  | 
"\<lbrakk> b \<in> carrier L; \<And> x. x \<in> carrier L \<Longrightarrow> b \<sqsubseteq> x \<rbrakk> \<Longrightarrow> b = \<bottom>"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
774  | 
by (metis bottom_closed bottom_lower le_antisym)  | 
| 14551 | 775  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
776  | 
lemma top_eq: "\<lbrakk> t \<in> carrier L; \<And> x. x \<in> carrier L \<Longrightarrow> x \<sqsubseteq> t \<rbrakk> \<Longrightarrow> t = \<top>"  | 
| 
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
777  | 
by (metis le_antisym top_closed top_higher)  | 
| 14551 | 778  | 
|
| 14693 | 779  | 
end  | 
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
780  | 
|
| 
69700
 
7a92cbec7030
new material about summations and powers, along with some tweaks
 
paulson <lp15@cam.ac.uk> 
parents: 
67443 
diff
changeset
 | 
781  | 
hide_const (open) Lattice.inf  | 
| 
 
7a92cbec7030
new material about summations and powers, along with some tweaks
 
paulson <lp15@cam.ac.uk> 
parents: 
67443 
diff
changeset
 | 
782  | 
hide_const (open) Lattice.sup  | 
| 
 
7a92cbec7030
new material about summations and powers, along with some tweaks
 
paulson <lp15@cam.ac.uk> 
parents: 
67443 
diff
changeset
 | 
783  | 
|
| 
65099
 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 
ballarin 
parents: 
63167 
diff
changeset
 | 
784  | 
end  |