author | nipkow |
Mon, 15 Jan 2024 22:50:13 +0100 | |
changeset 79490 | b287510a4202 |
parent 71469 | d7ef73df3d15 |
child 79564 | 33b10cd883ae |
permissions | -rw-r--r-- |
62695 | 1 |
(* Title: HOL/Datatype_Examples/Lift_BNF.thy |
2 |
Author: Dmitriy Traytel, ETH Zürich |
|
3 |
Copyright 2015 |
|
4 |
||
5 |
Demonstration of the "lift_bnf" command. |
|
6 |
*) |
|
7 |
||
63167 | 8 |
section \<open>Demonstration of the \textbf{lift_bnf} Command\<close> |
62695 | 9 |
|
60918 | 10 |
theory Lift_BNF |
11 |
imports Main |
|
12 |
begin |
|
13 |
||
71262 | 14 |
subsection \<open>Lifting \textbf{typedef}s\<close> |
15 |
||
60918 | 16 |
typedef 'a nonempty_list = "{xs :: 'a list. xs \<noteq> []}" |
17 |
by blast |
|
71262 | 18 |
setup_lifting type_definition_nonempty_list |
60918 | 19 |
|
20 |
lift_bnf (no_warn_wits) (neset: 'a) nonempty_list |
|
21 |
for map: nemap rel: nerel |
|
71262 | 22 |
by auto |
60918 | 23 |
|
24 |
typedef ('a :: finite, 'b) fin_nonempty_list = "{(xs :: 'a set, ys :: 'b list). ys \<noteq> []}" |
|
25 |
by blast |
|
71262 | 26 |
setup_lifting type_definition_fin_nonempty_list |
60918 | 27 |
|
71262 | 28 |
lift_bnf (no_warn_wits) (dead 'a :: finite, 'b) fin_nonempty_list |
29 |
[wits: "\<lambda>b. ({} :: 'a :: finite set, [b :: 'b])"] |
|
60918 | 30 |
by auto |
31 |
||
32 |
datatype 'a tree = Leaf | Node 'a "'a tree nonempty_list" |
|
33 |
||
34 |
record 'a point = |
|
35 |
xCoord :: 'a |
|
36 |
yCoord :: 'a |
|
37 |
||
71262 | 38 |
|
39 |
copy_bnf (no_warn_transfer) ('a, 's) point_ext |
|
60918 | 40 |
|
41 |
typedef 'a it = "UNIV :: 'a set" |
|
42 |
by blast |
|
43 |
||
71262 | 44 |
setup_lifting type_definition_it |
45 |
||
60918 | 46 |
copy_bnf (plugins del: size) 'a it |
47 |
||
48 |
typedef ('a, 'b) T_prod = "UNIV :: ('a \<times> 'b) set" |
|
49 |
by blast |
|
50 |
||
71262 | 51 |
setup_lifting type_definition_T_prod |
52 |
||
60918 | 53 |
copy_bnf ('a, 'b) T_prod |
54 |
||
55 |
typedef ('a, 'b, 'c) T_func = "UNIV :: ('a \<Rightarrow> 'b * 'c) set" |
|
56 |
by blast |
|
57 |
||
71262 | 58 |
setup_lifting type_definition_T_func |
59 |
||
60918 | 60 |
copy_bnf ('a, 'b, 'c) T_func |
61 |
||
62 |
typedef ('a, 'b) sum_copy = "UNIV :: ('a + 'b) set" |
|
63 |
by blast |
|
64 |
||
71262 | 65 |
setup_lifting type_definition_sum_copy |
66 |
||
60918 | 67 |
copy_bnf ('a, 'b) sum_copy |
68 |
||
69 |
typedef ('a, 'b) T_sum = "{Inl x | x. True} :: ('a + 'b) set" |
|
70 |
by blast |
|
71 |
||
71262 | 72 |
lift_bnf (no_warn_wits, no_warn_transfer) ('a, 'b) T_sum [wits: "Inl :: 'a \<Rightarrow> 'a + 'b"] |
73 |
by (force simp: map_sum_def sum_set_defs split: sum.splits)+ |
|
60918 | 74 |
|
75 |
typedef ('key, 'value) alist = "{xs :: ('key \<times> 'value) list. (distinct \<circ> map fst) xs}" |
|
76 |
morphisms impl_of Alist |
|
77 |
proof |
|
78 |
show "[] \<in> {xs. (distinct o map fst) xs}" |
|
79 |
by simp |
|
80 |
qed |
|
81 |
||
71262 | 82 |
setup_lifting type_definition_alist |
83 |
||
60918 | 84 |
lift_bnf (dead 'k, 'v) alist [wits: "Nil :: ('k \<times> 'v) list"] |
71262 | 85 |
by auto |
60918 | 86 |
|
87 |
typedef 'a myopt = "{X :: 'a set. finite X \<and> card X \<le> 1}" by (rule exI[of _ "{}"]) auto |
|
88 |
lemma myopt_type_def: "type_definition |
|
89 |
(\<lambda>X. if card (Rep_myopt X) = 1 then Some (the_elem (Rep_myopt X)) else None) |
|
90 |
(\<lambda>x. Abs_myopt (case x of Some x \<Rightarrow> {x} | _ \<Rightarrow> {})) |
|
91 |
(UNIV :: 'a option set)" |
|
92 |
apply unfold_locales |
|
93 |
apply (auto simp: Abs_myopt_inverse dest!: card_eq_SucD split: option.splits) |
|
94 |
apply (metis Rep_myopt_inverse) |
|
95 |
apply (metis One_nat_def Rep_myopt Rep_myopt_inverse Suc_le_mono card_0_eq le0 le_antisym mem_Collect_eq nat.exhaust) |
|
96 |
done |
|
97 |
||
71262 | 98 |
copy_bnf (no_warn_transfer) 'a myopt via myopt_type_def |
60918 | 99 |
|
62137 | 100 |
typedef ('k, 'v) fmap = "{M :: ('k \<rightharpoonup> 'v). finite (dom M)}" |
101 |
by (rule exI[of _ Map.empty]) simp_all |
|
71262 | 102 |
setup_lifting type_definition_fmap |
62137 | 103 |
|
104 |
lift_bnf (dead 'k, 'v) fmap [wits: "Map.empty :: 'k \<Rightarrow> 'v option"] |
|
105 |
by auto |
|
106 |
||
71262 | 107 |
typedef ('a, 'b) tuple_dead = "UNIV :: ('a \<times> 'b) set" .. |
108 |
typedef ('a, 'b) tuple_dead1 = "UNIV :: ('a \<times> 'b) set" .. |
|
109 |
typedef ('a, 'b) tuple_dead2 = "UNIV :: ('a \<times> 'b) set" .. |
|
110 |
typedef ('a, 'b, 'c) triple_dead = "UNIV :: ('a \<times> 'b \<times> 'c) set" .. |
|
111 |
typedef ('a, 'b, 'c) triple_dead1 = "UNIV :: ('a \<times> 'b \<times> 'c) set" .. |
|
112 |
typedef ('a, 'b, 'c) triple_dead2 = "UNIV :: ('a \<times> 'b \<times> 'c) set" .. |
|
113 |
typedef ('a, 'b, 'c) triple_dead3 = "UNIV :: ('a \<times> 'b \<times> 'c) set" .. |
|
114 |
typedef ('a, 'b, 'c) triple_dead12 = "UNIV :: ('a \<times> 'b \<times> 'c) set" .. |
|
115 |
typedef ('a, 'b, 'c) triple_dead13 = "UNIV :: ('a \<times> 'b \<times> 'c) set" .. |
|
116 |
typedef ('a, 'b, 'c) triple_dead23 = "UNIV :: ('a \<times> 'b \<times> 'c) set" .. |
|
117 |
typedef ('a, 'b, 'c, 'd) quadruple_dead = "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" .. |
|
118 |
typedef ('a, 'b, 'c, 'd) quadruple_dead1 = "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" .. |
|
119 |
typedef ('a, 'b, 'c, 'd) quadruple_dead2 = "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" .. |
|
120 |
typedef ('a, 'b, 'c, 'd) quadruple_dead3 = "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" .. |
|
121 |
typedef ('a, 'b, 'c, 'd) quadruple_dead4 = "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" .. |
|
122 |
typedef ('a, 'b, 'c, 'd) quadruple_dead12 = "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" .. |
|
123 |
typedef ('a, 'b, 'c, 'd) quadruple_dead13 = "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" .. |
|
124 |
typedef ('a, 'b, 'c, 'd) quadruple_dead14 = "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" .. |
|
125 |
typedef ('a, 'b, 'c, 'd) quadruple_dead23 = "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" .. |
|
126 |
typedef ('a, 'b, 'c, 'd) quadruple_dead24 = "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" .. |
|
127 |
typedef ('a, 'b, 'c, 'd) quadruple_dead34 = "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" .. |
|
128 |
typedef ('a, 'b, 'c, 'd) quadruple_dead123 = "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" .. |
|
129 |
typedef ('a, 'b, 'c, 'd) quadruple_dead124 = "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" .. |
|
130 |
typedef ('a, 'b, 'c, 'd) quadruple_dead134 = "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" .. |
|
131 |
typedef ('a, 'b, 'c, 'd) quadruple_dead234 = "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" .. |
|
132 |
setup_lifting type_definition_tuple_dead |
|
133 |
setup_lifting type_definition_tuple_dead1 |
|
134 |
setup_lifting type_definition_tuple_dead2 |
|
135 |
setup_lifting type_definition_triple_dead |
|
136 |
setup_lifting type_definition_triple_dead1 |
|
137 |
setup_lifting type_definition_triple_dead2 |
|
138 |
setup_lifting type_definition_triple_dead3 |
|
139 |
setup_lifting type_definition_triple_dead12 |
|
140 |
setup_lifting type_definition_triple_dead13 |
|
141 |
setup_lifting type_definition_triple_dead23 |
|
142 |
setup_lifting type_definition_quadruple_dead |
|
143 |
setup_lifting type_definition_quadruple_dead1 |
|
144 |
setup_lifting type_definition_quadruple_dead2 |
|
145 |
setup_lifting type_definition_quadruple_dead3 |
|
146 |
setup_lifting type_definition_quadruple_dead4 |
|
147 |
setup_lifting type_definition_quadruple_dead12 |
|
148 |
setup_lifting type_definition_quadruple_dead13 |
|
149 |
setup_lifting type_definition_quadruple_dead14 |
|
150 |
setup_lifting type_definition_quadruple_dead23 |
|
151 |
setup_lifting type_definition_quadruple_dead24 |
|
152 |
setup_lifting type_definition_quadruple_dead34 |
|
153 |
setup_lifting type_definition_quadruple_dead123 |
|
154 |
setup_lifting type_definition_quadruple_dead124 |
|
155 |
setup_lifting type_definition_quadruple_dead134 |
|
156 |
setup_lifting type_definition_quadruple_dead234 |
|
157 |
||
158 |
lift_bnf (no_warn_wits) ( 'a, 'b) tuple_dead by auto |
|
159 |
lift_bnf (no_warn_wits) (dead 'a, 'b) tuple_dead1 by auto |
|
160 |
lift_bnf (no_warn_wits) ( 'a, dead 'b) tuple_dead2 by auto |
|
161 |
lift_bnf (no_warn_wits) ( 'a, 'b, 'c) triple_dead by auto |
|
162 |
lift_bnf (no_warn_wits) (dead 'a, 'b, 'c) triple_dead1 by auto |
|
163 |
lift_bnf (no_warn_wits) ( 'a, dead 'b, 'c) triple_dead2 by auto |
|
164 |
lift_bnf (no_warn_wits) ( 'a, 'b, dead 'c) triple_dead3 by auto |
|
165 |
lift_bnf (no_warn_wits) (dead 'a, dead 'b, 'c) triple_dead12 by auto |
|
166 |
lift_bnf (no_warn_wits) (dead 'a, 'b, dead 'c) triple_dead13 by auto |
|
167 |
lift_bnf (no_warn_wits) ( 'a, dead 'b, dead 'c) triple_dead23 by auto |
|
168 |
lift_bnf (no_warn_wits) ( 'a, 'b, 'c, 'd) quadruple_dead by auto |
|
169 |
lift_bnf (no_warn_wits) (dead 'a, 'b, 'c, 'd) quadruple_dead1 by auto |
|
170 |
lift_bnf (no_warn_wits) ( 'a, dead 'b, 'c, 'd) quadruple_dead2 by auto |
|
171 |
lift_bnf (no_warn_wits) ( 'a, 'b, dead 'c, 'd) quadruple_dead3 by auto |
|
172 |
lift_bnf (no_warn_wits) ( 'a, 'b, 'c, dead 'd) quadruple_dead4 by auto |
|
173 |
lift_bnf (no_warn_wits) (dead 'a, dead 'b, 'c, 'd) quadruple_dead12 by auto |
|
174 |
lift_bnf (no_warn_wits) (dead 'a, 'b, dead 'c, 'd) quadruple_dead13 by auto |
|
175 |
lift_bnf (no_warn_wits) (dead 'a, 'b, 'c, dead 'd) quadruple_dead14 by auto |
|
176 |
lift_bnf (no_warn_wits) ( 'a, dead 'b, dead 'c, 'd) quadruple_dead23 by auto |
|
177 |
lift_bnf (no_warn_wits) ( 'a, dead 'b, 'c, dead 'd) quadruple_dead24 by auto |
|
178 |
lift_bnf (no_warn_wits) ( 'a, 'b, dead 'c, dead 'd) quadruple_dead34 by auto |
|
179 |
lift_bnf (no_warn_wits) (dead 'a, dead 'b, dead 'c, 'd) quadruple_dead123 by auto |
|
180 |
lift_bnf (no_warn_wits) (dead 'a, dead 'b, 'c, dead 'd) quadruple_dead124 by auto |
|
181 |
lift_bnf (no_warn_wits) (dead 'a, 'b, dead 'c, dead 'd) quadruple_dead134 by auto |
|
182 |
lift_bnf (no_warn_wits) ( 'a, dead 'b, dead 'c, dead 'd) quadruple_dead234 by auto |
|
183 |
||
184 |
subsection \<open>Lifting \textbf{quotient_type}s\<close> |
|
185 |
||
186 |
quotient_type 'a cpy_list = "'a list" / "(=)" |
|
187 |
by (rule identity_equivp) |
|
188 |
||
71469
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
traytel
parents:
71262
diff
changeset
|
189 |
lift_bnf 'a cpy_list |
71262 | 190 |
by (auto intro: list_all2_trans) |
191 |
||
192 |
quotient_type ('a, 'b) funk = "('a \<Rightarrow> 'b)" / "\<lambda>f g. \<forall>a. f a = g a" |
|
193 |
unfolding equivp_def by metis |
|
194 |
||
195 |
lemma funk_closure: "{(x, x'). \<forall>a. x a = x' a} `` {x. range x \<subseteq> A} = {x. range x \<subseteq> A}" |
|
196 |
by auto |
|
197 |
||
71469
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
traytel
parents:
71262
diff
changeset
|
198 |
lift_bnf (dead 'a, 'b) funk |
71262 | 199 |
unfolding funk_closure rel_fun_def by (auto 0 10 split: if_splits) |
200 |
||
201 |
lemma assumes "equivp REL" |
|
202 |
shows "REL OO REL OO R = REL OO R" |
|
203 |
using equivp_transp[OF assms] equivp_reflp[OF assms] |
|
204 |
by blast |
|
205 |
||
206 |
inductive ignore_Inl :: "'a + 'a \<Rightarrow> 'a + 'a \<Rightarrow> bool" where |
|
207 |
"ignore_Inl (Inl x) (Inl y)" |
|
208 |
| "ignore_Inl (Inr x) (Inr x)" |
|
209 |
||
210 |
inductive_simps [simp]: |
|
211 |
"ignore_Inl (Inl x) y" |
|
212 |
"ignore_Inl (Inr x') y" |
|
213 |
"ignore_Inl y (Inl x)" |
|
214 |
"ignore_Inl y (Inr x')" |
|
215 |
||
216 |
quotient_type 'a opt = "'a + 'a" / ignore_Inl |
|
217 |
apply(auto intro!: equivpI simp add: reflp_def symp_def transp_def elim!: ignore_Inl.cases) |
|
218 |
subgoal for x by(cases x) simp_all. |
|
219 |
||
220 |
lemma ignore_Inl_map_respect: |
|
221 |
"(rel_fun (=) (rel_fun ignore_Inl ignore_Inl)) (\<lambda>f. map_sum f f) (\<lambda>f. map_sum f f)" |
|
222 |
by(auto simp add: rel_fun_def elim: ignore_Inl.cases) |
|
223 |
||
224 |
lemma ignore_Inl_pos_distr: |
|
225 |
"A OO B \<noteq> bot \<Longrightarrow> ignore_Inl OO rel_sum A A OO ignore_Inl OO rel_sum B B OO ignore_Inl \<le> |
|
226 |
ignore_Inl OO rel_sum (A OO B) (A OO B) OO ignore_Inl" |
|
227 |
by(fastforce elim!: ignore_Inl.cases simp add: relcompp_apply fun_eq_iff intro: exI[where x="Inl _"]) |
|
228 |
||
229 |
lemma ignore_Inl_Inter: |
|
230 |
"\<lbrakk> S \<noteq> {}; \<Inter>S \<noteq> {} \<rbrakk> \<Longrightarrow> (\<Inter>A\<in>S. {(x, y). ignore_Inl x y} `` {x. Basic_BNFs.setl x \<union> Basic_BNFs.setr x \<subseteq> A}) \<subseteq> {(x, y). ignore_Inl x y} `` (\<Inter>A\<in>S. {x. Basic_BNFs.setl x \<union> Basic_BNFs.setr x \<subseteq> A})" |
|
231 |
apply(clarsimp; safe; clarsimp) |
|
232 |
subgoal for x A x' |
|
233 |
apply(cases x) |
|
234 |
apply(drule bspec[where x="Inl x'"]) |
|
235 |
apply clarsimp |
|
236 |
apply simp |
|
237 |
apply clarsimp |
|
238 |
apply(drule bspec[where x="Inr _"]) |
|
239 |
apply simp |
|
240 |
apply simp |
|
241 |
done |
|
242 |
done |
|
243 |
||
244 |
lift_bnf 'a opt [wits: "(Inl undefined :: 'a + 'a)"] |
|
245 |
apply(auto simp add: rel_fun_def elim: ignore_Inl.cases)[] |
|
246 |
apply(fastforce simp add: rel_sum.simps) |
|
247 |
subgoal for Ss |
|
248 |
using ignore_Inl_Inter[unfolded Plus_def, of Ss] |
|
249 |
by blast |
|
250 |
apply (auto simp: Ball_def image_iff sum_set_defs elim: ignore_Inl.cases split: sum.splits) [] |
|
251 |
done |
|
252 |
||
60918 | 253 |
end |