|
1 (* Title: HOL/Import/HOL_Light_Maps.thy |
|
2 Author: Cezary Kaliszyk, University of Innsbruck |
|
3 Author: Alexander Krauss, QAware GmbH |
|
4 |
|
5 Based on earlier code by Steven Obua and Sebastian Skalberg |
|
6 *) |
|
7 |
|
8 header {* Type and constant mappings of HOL Light importer *} |
|
9 |
|
10 theory HOL_Light_Maps |
|
11 imports Import_Setup |
|
12 begin |
|
13 |
|
14 lemma [import_const T]: |
|
15 "True = ((\<lambda>p :: bool. p) = (\<lambda>p. p))" |
|
16 by simp |
|
17 |
|
18 lemma [import_const "/\\"]: |
|
19 "(op \<and>) = (\<lambda>p q. (\<lambda>f. f p q \<Colon> bool) = (\<lambda>f. f True True))" |
|
20 by metis |
|
21 |
|
22 lemma [import_const "==>"]: |
|
23 "(op \<longrightarrow>) = (\<lambda>(p\<Colon>bool) q\<Colon>bool. (p \<and> q) = p)" |
|
24 by auto |
|
25 |
|
26 lemma [import_const "!"]: |
|
27 "All = (\<lambda>P\<Colon>'A \<Rightarrow> bool. P = (\<lambda>x\<Colon>'A. True))" |
|
28 by auto |
|
29 |
|
30 lemma [import_const "?"]: |
|
31 "Ex = (\<lambda>P\<Colon>'A \<Rightarrow> bool. All (\<lambda>q\<Colon>bool. (All (\<lambda>x\<Colon>'A\<Colon>type. (P x) \<longrightarrow> q)) \<longrightarrow> q))" |
|
32 by auto |
|
33 |
|
34 lemma [import_const "\\/"]: |
|
35 "(op \<or>) = (\<lambda>p q. \<forall>r. (p \<longrightarrow> r) \<longrightarrow> (q \<longrightarrow> r) \<longrightarrow> r)" |
|
36 by auto |
|
37 |
|
38 lemma [import_const F]: |
|
39 "False = (\<forall>p. p)" |
|
40 by auto |
|
41 |
|
42 lemma [import_const "~"]: |
|
43 "Not = (\<lambda>p. p \<longrightarrow> False)" |
|
44 by simp |
|
45 |
|
46 lemma [import_const "?!"]: |
|
47 "Ex1 = (\<lambda>P\<Colon>'A \<Rightarrow> bool. Ex P \<and> (\<forall>x y. P x \<and> P y \<longrightarrow> x = y))" |
|
48 by auto |
|
49 |
|
50 lemma [import_const "_FALSITY_"]: "False = False" |
|
51 by simp |
|
52 |
|
53 lemma hl_ax1: "\<forall>t\<Colon>'A \<Rightarrow> 'B. (\<lambda>x. t x) = t" |
|
54 by metis |
|
55 |
|
56 lemma hl_ax2: "\<forall>P x\<Colon>'A. P x \<longrightarrow> P (Eps P)" |
|
57 by (auto intro: someI) |
|
58 |
|
59 lemma [import_const COND]: |
|
60 "If = (\<lambda>t (t1 :: 'A) t2. SOME x. (t = True \<longrightarrow> x = t1) \<and> (t = False \<longrightarrow> x = t2))" |
|
61 unfolding fun_eq_iff by auto |
|
62 |
|
63 lemma [import_const o]: |
|
64 "(op \<circ>) = (\<lambda>(f\<Colon>'B \<Rightarrow> 'C) g x\<Colon>'A. f (g x))" |
|
65 unfolding fun_eq_iff by simp |
|
66 |
|
67 lemma [import_const I]: "id = (\<lambda>x\<Colon>'A. x)" |
|
68 by auto |
|
69 |
|
70 lemma [import_type 1 one_ABS one_REP]: |
|
71 "type_definition Rep_unit Abs_unit (Collect (\<lambda>b. b))" |
|
72 by (metis (full_types) Collect_cong singleton_conv2 type_definition_unit) |
|
73 |
|
74 lemma [import_const one]: "() = (SOME x\<Colon>unit. True)" |
|
75 by auto |
|
76 |
|
77 lemma [import_const mk_pair]: |
|
78 "Pair_Rep = (\<lambda>(x\<Colon>'A) (y\<Colon>'B) (a\<Colon>'A) b\<Colon>'B. a = x \<and> b = y)" |
|
79 by (simp add: Pair_Rep_def fun_eq_iff) |
|
80 |
|
81 lemma [import_type prod ABS_prod REP_prod]: |
|
82 "type_definition Rep_prod Abs_prod (Collect (\<lambda>x\<Colon>'A \<Rightarrow> 'B \<Rightarrow> bool. \<exists>a b. x = Pair_Rep a b))" |
|
83 using type_definition_prod[unfolded Product_Type.prod_def] by simp |
|
84 |
|
85 lemma [import_const ","]: "Pair = (\<lambda>(x\<Colon>'A) y\<Colon>'B. Abs_prod (Pair_Rep x y))" |
|
86 by (metis Pair_def) |
|
87 |
|
88 lemma [import_const FST]: |
|
89 "fst = (\<lambda>p\<Colon>'A \<times> 'B. SOME x\<Colon>'A. \<exists>y\<Colon>'B. p = (x, y))" |
|
90 by auto |
|
91 |
|
92 lemma [import_const SND]: |
|
93 "snd = (\<lambda>p\<Colon>'A \<times> 'B. SOME y\<Colon>'B. \<exists>x\<Colon>'A. p = (x, y))" |
|
94 by auto |
|
95 |
|
96 (*lemma [import_const CURRY]: |
|
97 "curry = (\<lambda>(f\<Colon>'A \<times> 'B \<Rightarrow> 'C) x y. f (x, y))" |
|
98 using curry_def .*) |
|
99 |
|
100 lemma [import_const ONE_ONE : Fun.inj]: |
|
101 "inj = (\<lambda>(f\<Colon>'A \<Rightarrow> 'B). \<forall>x1 x2. f x1 = f x2 \<longrightarrow> x1 = x2)" |
|
102 by (auto simp add: fun_eq_iff inj_on_def) |
|
103 |
|
104 lemma [import_const ONTO : Fun.surj]: |
|
105 "surj = (\<lambda>(f\<Colon>'A \<Rightarrow> 'B). \<forall>y. \<exists>x. y = f x)" |
|
106 by (auto simp add: fun_eq_iff) |
|
107 |
|
108 lemma hl_ax3: "\<exists>f\<Colon>ind \<Rightarrow> ind. inj f \<and> \<not> surj f" |
|
109 by (rule_tac x="Suc_Rep" in exI) |
|
110 (metis Suc_Rep_inject' injI Suc_Rep_not_Zero_Rep surjD) |
|
111 |
|
112 import_type_map num : Nat.nat |
|
113 import_const_map "_0" : Groups.zero_class.zero |
|
114 import_const_map SUC : Nat.Suc |
|
115 |
|
116 lemma NOT_SUC: "\<forall>n. Suc n \<noteq> 0" |
|
117 by simp |
|
118 |
|
119 lemma SUC_INJ: "\<forall>m n. (Suc m = Suc n) = (m = n)" |
|
120 by simp |
|
121 |
|
122 lemma num_INDUCTION: |
|
123 "\<forall>P. P 0 \<and> (\<forall>n. P n \<longrightarrow> P (Suc n)) \<longrightarrow> (\<forall>n. P n)" |
|
124 by (auto intro: nat.induct) |
|
125 |
|
126 lemma [import_const NUMERAL]: "id = (\<lambda>x :: nat. x)" |
|
127 by auto |
|
128 |
|
129 definition [simp]: "bit0 n = 2 * n" |
|
130 |
|
131 lemma [import_const BIT0]: |
|
132 "bit0 = (SOME fn. fn (id 0) = id 0 \<and> (\<forall>n. fn (Suc n) = Suc (Suc (fn n))))" |
|
133 apply (auto intro!: some_equality[symmetric]) |
|
134 apply (auto simp add: fun_eq_iff) |
|
135 apply (induct_tac x) |
|
136 apply auto |
|
137 done |
|
138 |
|
139 definition [import_const BIT1, simp]: |
|
140 "bit1 = (\<lambda>x. Suc (bit0 x))" |
|
141 |
|
142 definition [simp]: "pred n = n - 1" |
|
143 |
|
144 lemma PRE[import_const PRE : HOL_Light_Maps.pred]: |
|
145 "pred (id (0\<Colon>nat)) = id (0\<Colon>nat) \<and> (\<forall>n\<Colon>nat. pred (Suc n) = n)" |
|
146 by simp |
|
147 |
|
148 lemma ADD[import_const "+" : Groups.plus_class.plus]: |
|
149 "(\<forall>n :: nat. (id 0) + n = n) \<and> (\<forall>m n. (Suc m) + n = Suc (m + n))" |
|
150 by simp |
|
151 |
|
152 lemma MULT[import_const "*" : Groups.times_class.times]: |
|
153 "(\<forall>n :: nat. (id 0) * n = id 0) \<and> (\<forall>m n. (Suc m) * n = (m * n) + n)" |
|
154 by simp |
|
155 |
|
156 lemma EXP[import_const EXP : Power.power_class.power]: |
|
157 "(\<forall>m. m ^ (id 0) = id (bit1 0)) \<and> (\<forall>(m :: nat) n. m ^ (Suc n) = m * (m ^ n))" |
|
158 by simp |
|
159 |
|
160 lemma LE[import_const "<=" : Orderings.ord_class.less_eq]: |
|
161 "(\<forall>m :: nat. m \<le> (id 0) = (m = id 0)) \<and> (\<forall>m n. m \<le> (Suc n) = (m = Suc n \<or> m \<le> n))" |
|
162 by auto |
|
163 |
|
164 lemma LT[import_const "<" : Orderings.ord_class.less]: |
|
165 "(\<forall>m :: nat. m < (id 0) = False) \<and> (\<forall>m n. m < (Suc n) = (m = n \<or> m < n))" |
|
166 by auto |
|
167 |
|
168 lemma DEF_GE[import_const ">=" : "Orderings.ord_class.greater_eq"]: |
|
169 "(op \<ge>) = (\<lambda>x y :: nat. y \<le> x)" |
|
170 by simp |
|
171 |
|
172 lemma DEF_GT[import_const ">" : "Orderings.ord_class.greater"]: |
|
173 "(op >) = (\<lambda>x y :: nat. y < x)" |
|
174 by simp |
|
175 |
|
176 lemma DEF_MAX[import_const "MAX"]: |
|
177 "max = (\<lambda>x y :: nat. if x \<le> y then y else x)" |
|
178 by (auto simp add: min_max.le_iff_sup fun_eq_iff) |
|
179 |
|
180 lemma DEF_MIN[import_const "MIN"]: |
|
181 "min = (\<lambda>x y :: nat. if x \<le> y then x else y)" |
|
182 by (auto simp add: min_max.le_iff_inf fun_eq_iff) |
|
183 |
|
184 lemma EVEN[import_const "EVEN" : "Parity.even_odd_class.even"]: |
|
185 "even (id 0\<Colon>nat) = True \<and> (\<forall>n. even (Suc n) = (\<not> even n))" |
|
186 by simp |
|
187 |
|
188 lemma SUB[import_const "-" : "Groups.minus_class.minus"]: |
|
189 "(\<forall>m\<Colon>nat. m - (id 0) = m) \<and> (\<forall>m n. m - (Suc n) = pred (m - n))" |
|
190 by simp |
|
191 |
|
192 lemma FACT[import_const "FACT" : "Fact.fact_class.fact"]: |
|
193 "fact (id 0) = id (bit1 0) \<and> (\<forall>n. fact (Suc n) = Suc n * fact n)" |
|
194 by simp |
|
195 |
|
196 import_const_map MOD : Divides.div_class.mod |
|
197 import_const_map DIV : Divides.div_class.div |
|
198 |
|
199 lemma DIVISION_0: |
|
200 "\<forall>m n\<Colon>nat. if n = id 0 then m div n = id 0 \<and> m mod n = m else m = m div n * n + m mod n \<and> m mod n < n" |
|
201 by simp |
|
202 |
|
203 lemmas [import_type sum "_dest_sum" "_mk_sum"] = type_definition_sum |
|
204 import_const_map INL : Sum_Type.Inl |
|
205 import_const_map INR : Sum_Type.Inr |
|
206 |
|
207 lemma sum_INDUCT: |
|
208 "\<forall>P. (\<forall>a. P (Inl a)) \<and> (\<forall>a. P (Inr a)) \<longrightarrow> (\<forall>x. P x)" |
|
209 by (auto intro: sum.induct) |
|
210 |
|
211 lemma sum_RECURSION: |
|
212 "\<forall>Inl' Inr'. \<exists>fn. (\<forall>a. fn (Inl a) = Inl' a) \<and> (\<forall>a. fn (Inr a) = Inr' a)" |
|
213 by (intro allI, rule_tac x="sum_case Inl' Inr'" in exI) auto |
|
214 |
|
215 lemma OUTL[import_const "OUTL" : "Sum_Type.Projl"]: |
|
216 "Sum_Type.Projl (Inl x) = x" |
|
217 by simp |
|
218 |
|
219 lemma OUTR[import_const "OUTR" : "Sum_Type.Projr"]: |
|
220 "Sum_Type.Projr (Inr y) = y" |
|
221 by simp |
|
222 |
|
223 import_type_map list : List.list |
|
224 import_const_map NIL : List.list.Nil |
|
225 import_const_map CONS : List.list.Cons |
|
226 |
|
227 lemma list_INDUCT: |
|
228 "\<forall>P\<Colon>'A list \<Rightarrow> bool. P [] \<and> (\<forall>a0 a1. P a1 \<longrightarrow> P (a0 # a1)) \<longrightarrow> (\<forall>x. P x)" |
|
229 using list.induct by auto |
|
230 |
|
231 lemma list_RECURSION: |
|
232 "\<forall>nil' cons'. \<exists>fn\<Colon>'A list \<Rightarrow> 'Z. fn [] = nil' \<and> (\<forall>(a0\<Colon>'A) a1\<Colon>'A list. fn (a0 # a1) = cons' a0 a1 (fn a1))" |
|
233 by (intro allI, rule_tac x="list_rec nil' cons'" in exI) auto |
|
234 |
|
235 lemma HD[import_const HD : List.hd]: |
|
236 "hd ((h\<Colon>'A) # t) = h" |
|
237 by simp |
|
238 |
|
239 lemma TL[import_const TL : List.tl]: |
|
240 "tl ((h\<Colon>'A) # t) = t" |
|
241 by simp |
|
242 |
|
243 lemma APPEND[import_const APPEND : List.append]: |
|
244 "(\<forall>l\<Colon>'A list. [] @ l = l) \<and> (\<forall>(h\<Colon>'A) t l. (h # t) @ l = h # t @ l)" |
|
245 by simp |
|
246 |
|
247 lemma REVERSE[import_const REVERSE : List.rev]: |
|
248 "rev [] = ([] :: 'A list) \<and> rev ((x\<Colon>'A) # l) = rev l @ [x]" |
|
249 by simp |
|
250 |
|
251 lemma LENGTH[import_const LENGTH : List.length]: |
|
252 "length [] = id 0 \<and> (\<forall>(h\<Colon>'A) t. length (h # t) = Suc (length t))" |
|
253 by simp |
|
254 |
|
255 lemma MAP[import_const MAP : List.map]: |
|
256 "(\<forall>f\<Colon>'A \<Rightarrow> 'B. map f [] = []) \<and> |
|
257 (\<forall>(f\<Colon>'A \<Rightarrow> 'B) h t. map f (h # t) = f h # map f t)" |
|
258 by simp |
|
259 |
|
260 lemma LAST[import_const LAST : List.last]: |
|
261 "last ((h\<Colon>'A) # t) = (if t = [] then h else last t)" |
|
262 by simp |
|
263 |
|
264 lemma BUTLAST[import_const BUTLAST : List.butlast]: |
|
265 "butlast [] = ([] :: 't18337 list) \<and> |
|
266 butlast ((h :: 't18337) # t) = (if t = [] then [] else h # butlast t)" |
|
267 by simp |
|
268 |
|
269 lemma REPLICATE[import_const REPLICATE : List.replicate]: |
|
270 "replicate (id (0\<Colon>nat)) (x\<Colon>'t18358) = [] \<and> |
|
271 replicate (Suc n) x = x # replicate n x" |
|
272 by simp |
|
273 |
|
274 lemma NULL[import_const NULL : List.null]: |
|
275 "List.null ([]\<Colon>'t18373 list) = True \<and> List.null ((h\<Colon>'t18373) # t) = False" |
|
276 unfolding null_def by simp |
|
277 |
|
278 lemma ALL[import_const ALL : List.list_all]: |
|
279 "list_all (P\<Colon>'t18393 \<Rightarrow> bool) [] = True \<and> |
|
280 list_all P (h # t) = (P h \<and> list_all P t)" |
|
281 by simp |
|
282 |
|
283 lemma EX[import_const EX : List.list_ex]: |
|
284 "list_ex (P\<Colon>'t18414 \<Rightarrow> bool) [] = False \<and> |
|
285 list_ex P (h # t) = (P h \<or> list_ex P t)" |
|
286 by simp |
|
287 |
|
288 lemma ITLIST[import_const ITLIST : List.foldr]: |
|
289 "foldr (f\<Colon>'t18437 \<Rightarrow> 't18436 \<Rightarrow> 't18436) [] b = b \<and> |
|
290 foldr f (h # t) b = f h (foldr f t b)" |
|
291 by simp |
|
292 |
|
293 lemma ALL2_DEF[import_const ALL2 : List.list_all2]: |
|
294 "list_all2 (P\<Colon>'t18495 \<Rightarrow> 't18502 \<Rightarrow> bool) [] (l2\<Colon>'t18502 list) = (l2 = []) \<and> |
|
295 list_all2 P ((h1\<Colon>'t18495) # (t1\<Colon>'t18495 list)) l2 = |
|
296 (if l2 = [] then False else P h1 (hd l2) \<and> list_all2 P t1 (tl l2))" |
|
297 by simp (induct_tac l2, simp_all) |
|
298 |
|
299 lemma FILTER[import_const FILTER : List.filter]: |
|
300 "filter (P\<Colon>'t18680 \<Rightarrow> bool) [] = [] \<and> |
|
301 filter P ((h\<Colon>'t18680) # t) = (if P h then h # filter P t else filter P t)" |
|
302 by simp |
|
303 |
|
304 lemma ZIP[import_const ZIP : List.zip]: |
|
305 "zip [] [] = ([] :: ('t18824 \<times> 't18825) list) \<and> |
|
306 zip ((h1\<Colon>'t18849) # t1) ((h2\<Colon>'t18850) # t2) = (h1, h2) # zip t1 t2" |
|
307 by simp |
|
308 |
|
309 lemma WF[import_const WF : Wellfounded.wfP]: |
|
310 "wfP u \<longleftrightarrow> (\<forall>P. (\<exists>x :: 'A. P x) \<longrightarrow> (\<exists>x. P x \<and> (\<forall>y. u y x \<longrightarrow> \<not> P y)))" |
|
311 proof (intro allI iffI impI wfI_min[to_pred], elim exE wfE_min[to_pred]) |
|
312 fix x :: "'a \<Rightarrow> 'a \<Rightarrow> bool" and xa :: "'a" and Q |
|
313 assume a: "xa \<in> Q" |
|
314 assume "\<forall>P. Ex P \<longrightarrow> (\<exists>xa. P xa \<and> (\<forall>y. x y xa \<longrightarrow> \<not> P y))" |
|
315 then have "Ex (\<lambda>x. x \<in> Q) \<longrightarrow> (\<exists>xa. (\<lambda>x. x \<in> Q) xa \<and> (\<forall>y. x y xa \<longrightarrow> \<not> (\<lambda>x. x \<in> Q) y))" by auto |
|
316 then show "\<exists>z\<in>Q. \<forall>y. x y z \<longrightarrow> y \<notin> Q" using a by auto |
|
317 next |
|
318 fix x P and xa :: 'A and z |
|
319 assume "P xa" "z \<in> {a. P a}" "\<And>y. x y z \<Longrightarrow> y \<notin> {a. P a}" |
|
320 then show "\<exists>xa. P xa \<and> (\<forall>y. x y xa \<longrightarrow> \<not> P y)" by auto |
|
321 qed auto |
|
322 |
|
323 end |
|
324 |