author | nipkow |
Wed, 10 Jan 2018 15:25:09 +0100 | |
changeset 67399 | eab6ce8368fa |
parent 62390 | 842917225d56 |
child 77811 | ae9e6218443d |
permissions | -rw-r--r-- |
57998
8b7508f848ef
rename Quickcheck_Types to Lattice_Constructions and remove quickcheck setup
Andreas Lochbihler
parents:
57997
diff
changeset
|
1 |
(* Title: HOL/Library/Lattice_Constructions.thy |
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
2 |
Author: Lukas Bulwahn |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
3 |
Copyright 2010 TU Muenchen |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
4 |
*) |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
5 |
|
57998
8b7508f848ef
rename Quickcheck_Types to Lattice_Constructions and remove quickcheck setup
Andreas Lochbihler
parents:
57997
diff
changeset
|
6 |
theory Lattice_Constructions |
57645 | 7 |
imports Main |
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
8 |
begin |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
9 |
|
60500 | 10 |
subsection \<open>Values extended by a bottom element\<close> |
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
11 |
|
58310 | 12 |
datatype 'a bot = Value 'a | Bot |
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
13 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
14 |
instantiation bot :: (preorder) preorder |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
15 |
begin |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
16 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
17 |
definition less_eq_bot where |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
18 |
"x \<le> y \<longleftrightarrow> (case x of Bot \<Rightarrow> True | Value x \<Rightarrow> (case y of Bot \<Rightarrow> False | Value y \<Rightarrow> x \<le> y))" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
19 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
20 |
definition less_bot where |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
21 |
"x < y \<longleftrightarrow> (case y of Bot \<Rightarrow> False | Value y \<Rightarrow> (case x of Bot \<Rightarrow> True | Value x \<Rightarrow> x < y))" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
22 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
23 |
lemma less_eq_bot_Bot [simp]: "Bot \<le> x" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
24 |
by (simp add: less_eq_bot_def) |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
25 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
26 |
lemma less_eq_bot_Bot_code [code]: "Bot \<le> x \<longleftrightarrow> True" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
27 |
by simp |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
28 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
29 |
lemma less_eq_bot_Bot_is_Bot: "x \<le> Bot \<Longrightarrow> x = Bot" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
30 |
by (cases x) (simp_all add: less_eq_bot_def) |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
31 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
32 |
lemma less_eq_bot_Value_Bot [simp, code]: "Value x \<le> Bot \<longleftrightarrow> False" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
33 |
by (simp add: less_eq_bot_def) |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
34 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
35 |
lemma less_eq_bot_Value [simp, code]: "Value x \<le> Value y \<longleftrightarrow> x \<le> y" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
36 |
by (simp add: less_eq_bot_def) |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
37 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
38 |
lemma less_bot_Bot [simp, code]: "x < Bot \<longleftrightarrow> False" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
39 |
by (simp add: less_bot_def) |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
40 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
41 |
lemma less_bot_Bot_is_Value: "Bot < x \<Longrightarrow> \<exists>z. x = Value z" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
42 |
by (cases x) (simp_all add: less_bot_def) |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
43 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
44 |
lemma less_bot_Bot_Value [simp]: "Bot < Value x" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
45 |
by (simp add: less_bot_def) |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
46 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
47 |
lemma less_bot_Bot_Value_code [code]: "Bot < Value x \<longleftrightarrow> True" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
48 |
by simp |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
49 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
50 |
lemma less_bot_Value [simp, code]: "Value x < Value y \<longleftrightarrow> x < y" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
51 |
by (simp add: less_bot_def) |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
52 |
|
60508 | 53 |
instance |
60679 | 54 |
by standard |
60508 | 55 |
(auto simp add: less_eq_bot_def less_bot_def less_le_not_le elim: order_trans split: bot.splits) |
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
56 |
|
60508 | 57 |
end |
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
58 |
|
60508 | 59 |
instance bot :: (order) order |
60679 | 60 |
by standard (auto simp add: less_eq_bot_def less_bot_def split: bot.splits) |
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
61 |
|
60508 | 62 |
instance bot :: (linorder) linorder |
60679 | 63 |
by standard (auto simp add: less_eq_bot_def less_bot_def split: bot.splits) |
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
64 |
|
43815
4f6e2965d821
adjusted to tightened specification of classes bot and top
haftmann
parents:
40654
diff
changeset
|
65 |
instantiation bot :: (order) bot |
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
66 |
begin |
60508 | 67 |
definition "bot = Bot" |
68 |
instance .. |
|
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
69 |
end |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
70 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
71 |
instantiation bot :: (top) top |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
72 |
begin |
60508 | 73 |
definition "top = Value top" |
74 |
instance .. |
|
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
75 |
end |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
76 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
77 |
instantiation bot :: (semilattice_inf) semilattice_inf |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
78 |
begin |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
79 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
80 |
definition inf_bot |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
81 |
where |
60508 | 82 |
"inf x y = |
83 |
(case x of |
|
84 |
Bot \<Rightarrow> Bot |
|
85 |
| Value v \<Rightarrow> |
|
86 |
(case y of |
|
87 |
Bot \<Rightarrow> Bot |
|
88 |
| Value v' \<Rightarrow> Value (inf v v')))" |
|
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
89 |
|
60508 | 90 |
instance |
60679 | 91 |
by standard (auto simp add: inf_bot_def less_eq_bot_def split: bot.splits) |
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
92 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
93 |
end |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
94 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
95 |
instantiation bot :: (semilattice_sup) semilattice_sup |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
96 |
begin |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
97 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
98 |
definition sup_bot |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
99 |
where |
60508 | 100 |
"sup x y = |
101 |
(case x of |
|
102 |
Bot \<Rightarrow> y |
|
103 |
| Value v \<Rightarrow> |
|
104 |
(case y of |
|
105 |
Bot \<Rightarrow> x |
|
106 |
| Value v' \<Rightarrow> Value (sup v v')))" |
|
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
107 |
|
60508 | 108 |
instance |
60679 | 109 |
by standard (auto simp add: sup_bot_def less_eq_bot_def split: bot.splits) |
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
110 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
111 |
end |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
112 |
|
57543
36041934e429
adapt and reactivate Quickcheck_Types and add two test cases
Andreas Lochbihler
parents:
44890
diff
changeset
|
113 |
instance bot :: (lattice) bounded_lattice_bot |
60508 | 114 |
by intro_classes (simp add: bot_bot_def) |
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
115 |
|
60508 | 116 |
|
117 |
subsection \<open>Values extended by a top element\<close> |
|
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
118 |
|
58310 | 119 |
datatype 'a top = Value 'a | Top |
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
120 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
121 |
instantiation top :: (preorder) preorder |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
122 |
begin |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
123 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
124 |
definition less_eq_top where |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
125 |
"x \<le> y \<longleftrightarrow> (case y of Top \<Rightarrow> True | Value y \<Rightarrow> (case x of Top \<Rightarrow> False | Value x \<Rightarrow> x \<le> y))" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
126 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
127 |
definition less_top where |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
128 |
"x < y \<longleftrightarrow> (case x of Top \<Rightarrow> False | Value x \<Rightarrow> (case y of Top \<Rightarrow> True | Value y \<Rightarrow> x < y))" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
129 |
|
60508 | 130 |
lemma less_eq_top_Top [simp]: "x \<le> Top" |
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
131 |
by (simp add: less_eq_top_def) |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
132 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
133 |
lemma less_eq_top_Top_code [code]: "x \<le> Top \<longleftrightarrow> True" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
134 |
by simp |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
135 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
136 |
lemma less_eq_top_is_Top: "Top \<le> x \<Longrightarrow> x = Top" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
137 |
by (cases x) (simp_all add: less_eq_top_def) |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
138 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
139 |
lemma less_eq_top_Top_Value [simp, code]: "Top \<le> Value x \<longleftrightarrow> False" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
140 |
by (simp add: less_eq_top_def) |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
141 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
142 |
lemma less_eq_top_Value_Value [simp, code]: "Value x \<le> Value y \<longleftrightarrow> x \<le> y" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
143 |
by (simp add: less_eq_top_def) |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
144 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
145 |
lemma less_top_Top [simp, code]: "Top < x \<longleftrightarrow> False" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
146 |
by (simp add: less_top_def) |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
147 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
148 |
lemma less_top_Top_is_Value: "x < Top \<Longrightarrow> \<exists>z. x = Value z" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
149 |
by (cases x) (simp_all add: less_top_def) |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
150 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
151 |
lemma less_top_Value_Top [simp]: "Value x < Top" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
152 |
by (simp add: less_top_def) |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
153 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
154 |
lemma less_top_Value_Top_code [code]: "Value x < Top \<longleftrightarrow> True" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
155 |
by simp |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
156 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
157 |
lemma less_top_Value [simp, code]: "Value x < Value y \<longleftrightarrow> x < y" |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
158 |
by (simp add: less_top_def) |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
159 |
|
60508 | 160 |
instance |
60679 | 161 |
by standard |
60508 | 162 |
(auto simp add: less_eq_top_def less_top_def less_le_not_le elim: order_trans split: top.splits) |
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
163 |
|
60508 | 164 |
end |
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
165 |
|
60508 | 166 |
instance top :: (order) order |
60679 | 167 |
by standard (auto simp add: less_eq_top_def less_top_def split: top.splits) |
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
168 |
|
60508 | 169 |
instance top :: (linorder) linorder |
60679 | 170 |
by standard (auto simp add: less_eq_top_def less_top_def split: top.splits) |
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
171 |
|
43815
4f6e2965d821
adjusted to tightened specification of classes bot and top
haftmann
parents:
40654
diff
changeset
|
172 |
instantiation top :: (order) top |
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
173 |
begin |
60508 | 174 |
definition "top = Top" |
175 |
instance .. |
|
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
176 |
end |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
177 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
178 |
instantiation top :: (bot) bot |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
179 |
begin |
60508 | 180 |
definition "bot = Value bot" |
181 |
instance .. |
|
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
182 |
end |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
183 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
184 |
instantiation top :: (semilattice_inf) semilattice_inf |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
185 |
begin |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
186 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
187 |
definition inf_top |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
188 |
where |
60508 | 189 |
"inf x y = |
190 |
(case x of |
|
191 |
Top \<Rightarrow> y |
|
192 |
| Value v \<Rightarrow> |
|
193 |
(case y of |
|
194 |
Top \<Rightarrow> x |
|
195 |
| Value v' \<Rightarrow> Value (inf v v')))" |
|
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
196 |
|
60508 | 197 |
instance |
60679 | 198 |
by standard (auto simp add: inf_top_def less_eq_top_def split: top.splits) |
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
199 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
200 |
end |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
201 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
202 |
instantiation top :: (semilattice_sup) semilattice_sup |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
203 |
begin |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
204 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
205 |
definition sup_top |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
206 |
where |
60508 | 207 |
"sup x y = |
208 |
(case x of |
|
209 |
Top \<Rightarrow> Top |
|
210 |
| Value v \<Rightarrow> |
|
211 |
(case y of |
|
212 |
Top \<Rightarrow> Top |
|
213 |
| Value v' \<Rightarrow> Value (sup v v')))" |
|
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
214 |
|
60508 | 215 |
instance |
60679 | 216 |
by standard (auto simp add: sup_top_def less_eq_top_def split: top.splits) |
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
217 |
|
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
218 |
end |
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
219 |
|
57543
36041934e429
adapt and reactivate Quickcheck_Types and add two test cases
Andreas Lochbihler
parents:
44890
diff
changeset
|
220 |
instance top :: (lattice) bounded_lattice_top |
60679 | 221 |
by standard (simp add: top_top_def) |
60508 | 222 |
|
37915
e709e764f20c
adding Library theory for other quickcheck default types
bulwahn
parents:
diff
changeset
|
223 |
|
60500 | 224 |
subsection \<open>Values extended by a top and a bottom element\<close> |
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
225 |
|
58310 | 226 |
datatype 'a flat_complete_lattice = Value 'a | Bot | Top |
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
227 |
|
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
228 |
instantiation flat_complete_lattice :: (type) order |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
229 |
begin |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
230 |
|
60508 | 231 |
definition less_eq_flat_complete_lattice |
232 |
where |
|
233 |
"x \<le> y \<equiv> |
|
234 |
(case x of |
|
235 |
Bot \<Rightarrow> True |
|
236 |
| Value v1 \<Rightarrow> |
|
237 |
(case y of |
|
238 |
Bot \<Rightarrow> False |
|
239 |
| Value v2 \<Rightarrow> v1 = v2 |
|
240 |
| Top \<Rightarrow> True) |
|
241 |
| Top \<Rightarrow> y = Top)" |
|
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
242 |
|
60508 | 243 |
definition less_flat_complete_lattice |
244 |
where |
|
245 |
"x < y = |
|
246 |
(case x of |
|
247 |
Bot \<Rightarrow> y \<noteq> Bot |
|
248 |
| Value v1 \<Rightarrow> y = Top |
|
249 |
| Top \<Rightarrow> False)" |
|
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
250 |
|
60508 | 251 |
lemma [simp]: "Bot \<le> y" |
252 |
unfolding less_eq_flat_complete_lattice_def by auto |
|
253 |
||
254 |
lemma [simp]: "y \<le> Top" |
|
255 |
unfolding less_eq_flat_complete_lattice_def by (auto split: flat_complete_lattice.splits) |
|
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
256 |
|
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
257 |
lemma greater_than_two_values: |
60509 | 258 |
assumes "a \<noteq> b" "Value a \<le> z" "Value b \<le> z" |
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
259 |
shows "z = Top" |
60508 | 260 |
using assms |
261 |
by (cases z) (auto simp add: less_eq_flat_complete_lattice_def) |
|
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
262 |
|
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
263 |
lemma lesser_than_two_values: |
60509 | 264 |
assumes "a \<noteq> b" "z \<le> Value a" "z \<le> Value b" |
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
265 |
shows "z = Bot" |
60508 | 266 |
using assms |
267 |
by (cases z) (auto simp add: less_eq_flat_complete_lattice_def) |
|
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
268 |
|
60508 | 269 |
instance |
60679 | 270 |
by standard |
60508 | 271 |
(auto simp add: less_eq_flat_complete_lattice_def less_flat_complete_lattice_def |
272 |
split: flat_complete_lattice.splits) |
|
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
273 |
|
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
274 |
end |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
275 |
|
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
276 |
instantiation flat_complete_lattice :: (type) bot |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
277 |
begin |
60508 | 278 |
definition "bot = Bot" |
279 |
instance .. |
|
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
280 |
end |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
281 |
|
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
282 |
instantiation flat_complete_lattice :: (type) top |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
283 |
begin |
60508 | 284 |
definition "top = Top" |
285 |
instance .. |
|
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
286 |
end |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
287 |
|
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
288 |
instantiation flat_complete_lattice :: (type) lattice |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
289 |
begin |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
290 |
|
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
291 |
definition inf_flat_complete_lattice |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
292 |
where |
60508 | 293 |
"inf x y = |
294 |
(case x of |
|
295 |
Bot \<Rightarrow> Bot |
|
296 |
| Value v1 \<Rightarrow> |
|
297 |
(case y of |
|
298 |
Bot \<Rightarrow> Bot |
|
299 |
| Value v2 \<Rightarrow> if v1 = v2 then x else Bot |
|
300 |
| Top \<Rightarrow> x) |
|
301 |
| Top \<Rightarrow> y)" |
|
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
302 |
|
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
303 |
definition sup_flat_complete_lattice |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
304 |
where |
60508 | 305 |
"sup x y = |
306 |
(case x of |
|
307 |
Bot \<Rightarrow> y |
|
308 |
| Value v1 \<Rightarrow> |
|
309 |
(case y of |
|
310 |
Bot \<Rightarrow> x |
|
311 |
| Value v2 \<Rightarrow> if v1 = v2 then x else Top |
|
312 |
| Top \<Rightarrow> Top) |
|
313 |
| Top \<Rightarrow> Top)" |
|
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
314 |
|
60508 | 315 |
instance |
60679 | 316 |
by standard |
60508 | 317 |
(auto simp add: inf_flat_complete_lattice_def sup_flat_complete_lattice_def |
318 |
less_eq_flat_complete_lattice_def split: flat_complete_lattice.splits) |
|
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
319 |
|
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
320 |
end |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
321 |
|
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
322 |
instantiation flat_complete_lattice :: (type) complete_lattice |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
323 |
begin |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
324 |
|
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
325 |
definition Sup_flat_complete_lattice |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
326 |
where |
60508 | 327 |
"Sup A = |
328 |
(if A = {} \<or> A = {Bot} then Bot |
|
329 |
else if \<exists>v. A - {Bot} = {Value v} then Value (THE v. A - {Bot} = {Value v}) |
|
330 |
else Top)" |
|
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
331 |
|
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
332 |
definition Inf_flat_complete_lattice |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
333 |
where |
60508 | 334 |
"Inf A = |
335 |
(if A = {} \<or> A = {Top} then Top |
|
336 |
else if \<exists>v. A - {Top} = {Value v} then Value (THE v. A - {Top} = {Value v}) |
|
337 |
else Bot)" |
|
338 |
||
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
339 |
instance |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
340 |
proof |
60508 | 341 |
fix x :: "'a flat_complete_lattice" |
342 |
fix A |
|
343 |
assume "x \<in> A" |
|
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
344 |
{ |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
345 |
fix v |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
346 |
assume "A - {Top} = {Value v}" |
60508 | 347 |
then have "(THE v. A - {Top} = {Value v}) = v" |
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
348 |
by (auto intro!: the1_equality) |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
349 |
moreover |
60508 | 350 |
from \<open>x \<in> A\<close> \<open>A - {Top} = {Value v}\<close> have "x = Top \<or> x = Value v" |
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
351 |
by auto |
60508 | 352 |
ultimately have "Value (THE v. A - {Top} = {Value v}) \<le> x" |
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
353 |
by auto |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
354 |
} |
60508 | 355 |
with \<open>x \<in> A\<close> show "Inf A \<le> x" |
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
356 |
unfolding Inf_flat_complete_lattice_def |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
43815
diff
changeset
|
357 |
by fastforce |
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
358 |
next |
60508 | 359 |
fix z :: "'a flat_complete_lattice" |
360 |
fix A |
|
60509 | 361 |
show "z \<le> Inf A" if z: "\<And>x. x \<in> A \<Longrightarrow> z \<le> x" |
362 |
proof - |
|
363 |
consider "A = {} \<or> A = {Top}" |
|
364 |
| "A \<noteq> {}" "A \<noteq> {Top}" "\<exists>v. A - {Top} = {Value v}" |
|
365 |
| "A \<noteq> {}" "A \<noteq> {Top}" "\<not> (\<exists>v. A - {Top} = {Value v})" |
|
366 |
by blast |
|
367 |
then show ?thesis |
|
368 |
proof cases |
|
369 |
case 1 |
|
370 |
then have "Inf A = Top" |
|
371 |
unfolding Inf_flat_complete_lattice_def by auto |
|
372 |
then show ?thesis by simp |
|
373 |
next |
|
374 |
case 2 |
|
375 |
then obtain v where v1: "A - {Top} = {Value v}" |
|
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
376 |
by auto |
60509 | 377 |
then have v2: "(THE v. A - {Top} = {Value v}) = v" |
378 |
by (auto intro!: the1_equality) |
|
379 |
from 2 v2 have Inf: "Inf A = Value v" |
|
380 |
unfolding Inf_flat_complete_lattice_def by simp |
|
381 |
from v1 have "Value v \<in> A" by blast |
|
382 |
then have "z \<le> Value v" by (rule z) |
|
383 |
with Inf show ?thesis by simp |
|
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
384 |
next |
60509 | 385 |
case 3 |
386 |
then have Inf: "Inf A = Bot" |
|
387 |
unfolding Inf_flat_complete_lattice_def by auto |
|
388 |
have "z \<le> Bot" |
|
389 |
proof (cases "A - {Top} = {Bot}") |
|
390 |
case True |
|
391 |
then have "Bot \<in> A" by blast |
|
392 |
then show ?thesis by (rule z) |
|
393 |
next |
|
394 |
case False |
|
395 |
from 3 obtain a1 where a1: "a1 \<in> A - {Top}" |
|
396 |
by auto |
|
397 |
from 3 False a1 obtain a2 where "a2 \<in> A - {Top} \<and> a1 \<noteq> a2" |
|
398 |
by (cases a1) auto |
|
399 |
with a1 z[of "a1"] z[of "a2"] show ?thesis |
|
400 |
apply (cases a1) |
|
401 |
apply auto |
|
402 |
apply (cases a2) |
|
403 |
apply auto |
|
404 |
apply (auto dest!: lesser_than_two_values) |
|
405 |
done |
|
406 |
qed |
|
407 |
with Inf show ?thesis by simp |
|
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
408 |
qed |
60509 | 409 |
qed |
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
410 |
next |
60508 | 411 |
fix x :: "'a flat_complete_lattice" |
412 |
fix A |
|
413 |
assume "x \<in> A" |
|
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
414 |
{ |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
415 |
fix v |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
416 |
assume "A - {Bot} = {Value v}" |
60508 | 417 |
then have "(THE v. A - {Bot} = {Value v}) = v" |
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
418 |
by (auto intro!: the1_equality) |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
419 |
moreover |
60508 | 420 |
from \<open>x \<in> A\<close> \<open>A - {Bot} = {Value v}\<close> have "x = Bot \<or> x = Value v" |
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
421 |
by auto |
60508 | 422 |
ultimately have "x \<le> Value (THE v. A - {Bot} = {Value v})" |
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
423 |
by auto |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
424 |
} |
60508 | 425 |
with \<open>x \<in> A\<close> show "x \<le> Sup A" |
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
426 |
unfolding Sup_flat_complete_lattice_def |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
43815
diff
changeset
|
427 |
by fastforce |
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
428 |
next |
60508 | 429 |
fix z :: "'a flat_complete_lattice" |
430 |
fix A |
|
60509 | 431 |
show "Sup A \<le> z" if z: "\<And>x. x \<in> A \<Longrightarrow> x \<le> z" |
432 |
proof - |
|
433 |
consider "A = {} \<or> A = {Bot}" |
|
434 |
| "A \<noteq> {}" "A \<noteq> {Bot}" "\<exists>v. A - {Bot} = {Value v}" |
|
435 |
| "A \<noteq> {}" "A \<noteq> {Bot}" "\<not> (\<exists>v. A - {Bot} = {Value v})" |
|
436 |
by blast |
|
437 |
then show ?thesis |
|
438 |
proof cases |
|
439 |
case 1 |
|
440 |
then have "Sup A = Bot" |
|
441 |
unfolding Sup_flat_complete_lattice_def by auto |
|
442 |
then show ?thesis by simp |
|
443 |
next |
|
444 |
case 2 |
|
445 |
then obtain v where v1: "A - {Bot} = {Value v}" |
|
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
446 |
by auto |
60509 | 447 |
then have v2: "(THE v. A - {Bot} = {Value v}) = v" |
448 |
by (auto intro!: the1_equality) |
|
449 |
from 2 v2 have Sup: "Sup A = Value v" |
|
450 |
unfolding Sup_flat_complete_lattice_def by simp |
|
451 |
from v1 have "Value v \<in> A" by blast |
|
452 |
then have "Value v \<le> z" by (rule z) |
|
453 |
with Sup show ?thesis by simp |
|
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
454 |
next |
60509 | 455 |
case 3 |
456 |
then have Sup: "Sup A = Top" |
|
457 |
unfolding Sup_flat_complete_lattice_def by auto |
|
458 |
have "Top \<le> z" |
|
459 |
proof (cases "A - {Bot} = {Top}") |
|
460 |
case True |
|
461 |
then have "Top \<in> A" by blast |
|
462 |
then show ?thesis by (rule z) |
|
463 |
next |
|
464 |
case False |
|
465 |
from 3 obtain a1 where a1: "a1 \<in> A - {Bot}" |
|
466 |
by auto |
|
467 |
from 3 False a1 obtain a2 where "a2 \<in> A - {Bot} \<and> a1 \<noteq> a2" |
|
468 |
by (cases a1) auto |
|
469 |
with a1 z[of "a1"] z[of "a2"] show ?thesis |
|
470 |
apply (cases a1) |
|
471 |
apply auto |
|
472 |
apply (cases a2) |
|
473 |
apply (auto dest!: greater_than_two_values) |
|
474 |
done |
|
475 |
qed |
|
476 |
with Sup show ?thesis by simp |
|
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
477 |
qed |
60509 | 478 |
qed |
57543
36041934e429
adapt and reactivate Quickcheck_Types and add two test cases
Andreas Lochbihler
parents:
44890
diff
changeset
|
479 |
next |
36041934e429
adapt and reactivate Quickcheck_Types and add two test cases
Andreas Lochbihler
parents:
44890
diff
changeset
|
480 |
show "Inf {} = (top :: 'a flat_complete_lattice)" |
60508 | 481 |
by (simp add: Inf_flat_complete_lattice_def top_flat_complete_lattice_def) |
57543
36041934e429
adapt and reactivate Quickcheck_Types and add two test cases
Andreas Lochbihler
parents:
44890
diff
changeset
|
482 |
show "Sup {} = (bot :: 'a flat_complete_lattice)" |
60508 | 483 |
by (simp add: Sup_flat_complete_lattice_def bot_flat_complete_lattice_def) |
37918
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
484 |
qed |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
485 |
|
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
486 |
end |
eda5faaca9e2
adding a type for flat complete lattice to Quickcheck_Types
bulwahn
parents:
37915
diff
changeset
|
487 |
|
62390 | 488 |
end |