1 theory Test |
1 theory Test |
2 imports "Parser" |
2 imports "Parser" "../Attic/Prove" |
3 begin |
3 begin |
4 |
4 |
5 text {* weirdo example from Peter Sewell's bestiary *} |
5 text {* weirdo example from Peter Sewell's bestiary *} |
6 |
6 |
7 nominal_datatype weird = |
7 (*nominal_datatype weird = |
8 WBind x::"name" y::"name" p1::"weird" p2::"weird" p3::"weird" |
8 WBind x::"name" y::"name" p1::"weird" p2::"weird" p3::"weird" |
9 bind x in p1, bind x in p2, bind y in p2, bind y in p3 |
9 bind x in p1, bind x in p2, bind y in p2, bind y in p3 |
10 | WV "name" |
10 | WV "name" |
11 | WP "weird" "weird" |
11 | WP "weird" "weird" |
12 |
12 |
13 thm permute_weird_raw.simps[no_vars] |
13 thm permute_weird_raw.simps[no_vars] |
14 thm alpha_weird_raw.intros[no_vars] |
14 thm alpha_weird_raw.intros[no_vars] |
15 thm fv_weird_raw.simps[no_vars] |
15 thm fv_weird_raw.simps[no_vars] |
|
16 |
|
17 thm eqvts |
|
18 |
|
19 local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding weird_inj}, []), (build_alpha_inj @{thms alpha_weird_raw.intros} @{thms weird_raw.distinct weird_raw.inject} @{thms alpha_weird_raw.cases} ctxt)) ctxt)) *} |
|
20 thm weird_inj |
|
21 |
|
22 local_setup {* |
|
23 (fn ctxt => snd (Local_Theory.note ((@{binding alpha_eqvt}, []), |
|
24 build_alpha_eqvts [@{term alpha_weird_raw}] [@{term "permute :: perm \<Rightarrow> weird_raw \<Rightarrow> weird_raw"}] @{thms permute_weird_raw.simps weird_inj} @{thm alpha_weird_raw.induct} ctxt) ctxt)) *} |
|
25 |
|
26 (*prove {* (snd o snd) (build_alpha_refl_gl [@{term alpha_weird_raw}] ("x","y","z")) *} |
|
27 |
|
28 apply (tactic {* transp_tac @{context} @{thm alpha_weird_raw.induct} @{thms weird_inj} @{thms weird_raw.inject} @{thms weird_raw.distinct} @{thms alpha_weird_raw.cases} @{thms alpha_eqvt} 1 *}) |
|
29 *) |
|
30 lemma "alpha_weird_raw x y \<longrightarrow> (\<forall>z. alpha_weird_raw y z \<longrightarrow> alpha_weird_raw x z)" |
|
31 apply (rule impI) |
|
32 apply (erule alpha_weird_raw.induct) |
|
33 apply (simp_all add: weird_inj) |
|
34 defer |
|
35 apply (rule allI) |
|
36 apply (rule impI) |
|
37 apply (erule alpha_weird_raw.cases) |
|
38 apply (simp_all add: weird_inj) |
|
39 apply (rule allI) |
|
40 apply (rule impI) |
|
41 apply (erule alpha_weird_raw.cases) |
|
42 apply (simp_all add: weird_inj) |
|
43 apply (erule conjE)+ |
|
44 apply (erule exE)+ |
|
45 apply (erule conjE)+ |
|
46 apply (erule exE)+ |
|
47 apply (rule conjI) |
|
48 apply (rule_tac x="pica + pic" in exI) |
|
49 apply (erule alpha_gen_compose_trans) |
|
50 apply assumption |
|
51 apply (simp add: alpha_eqvt) |
|
52 apply (rule_tac x="pia + pib" in exI) |
|
53 apply (rule_tac x="piaa + piba" in exI) |
|
54 apply (rule conjI) |
|
55 apply (erule alpha_gen_compose_trans) |
|
56 apply assumption |
|
57 apply (simp add: alpha_eqvt) |
|
58 apply (rule conjI) |
|
59 defer |
|
60 apply (rule_tac x="pid + pi" in exI) |
|
61 apply (erule alpha_gen_compose_trans) |
|
62 apply assumption |
|
63 apply (simp add: alpha_eqvt) |
|
64 sorry |
|
65 |
|
66 lemma "alpha_weird_raw x y \<Longrightarrow> alpha_weird_raw y x" |
|
67 apply (erule alpha_weird_raw.induct) |
|
68 apply (simp_all add: weird_inj) |
|
69 apply (erule conjE)+ |
|
70 apply (erule exE)+ |
|
71 apply (erule conjE)+ |
|
72 apply (erule exE)+ |
|
73 apply (rule conjI) |
|
74 apply (rule_tac x="- pic" in exI) |
|
75 apply (erule alpha_gen_compose_sym) |
|
76 apply (simp_all add: alpha_eqvt) |
|
77 apply (rule_tac x="- pia" in exI) |
|
78 apply (rule_tac x="- pib" in exI) |
|
79 apply (simp add: minus_add[symmetric]) |
|
80 apply (rule conjI) |
|
81 apply (erule alpha_gen_compose_sym) |
|
82 apply (simp_all add: alpha_eqvt) |
|
83 apply (rule conjI) |
|
84 apply (simp add: supp_minus_perm Int_commute) |
|
85 apply (rule_tac x="- pi" in exI) |
|
86 apply (erule alpha_gen_compose_sym) |
|
87 apply (simp_all add: alpha_eqvt) |
|
88 done |
16 |
89 |
17 |
90 |
18 abbreviation "WBind \<equiv> WBind_raw" |
91 abbreviation "WBind \<equiv> WBind_raw" |
19 abbreviation "WP \<equiv> WP_raw" |
92 abbreviation "WP \<equiv> WP_raw" |
20 abbreviation "WV \<equiv> WV_raw" |
93 abbreviation "WV \<equiv> WV_raw" |
212 "bv5 Lnil = {}" |
294 "bv5 Lnil = {}" |
213 | "bv5 (Lcons n t ltl) = {atom n} \<union> (bv5 ltl)" |
295 | "bv5 (Lcons n t ltl) = {atom n} \<union> (bv5 ltl)" |
214 |
296 |
215 (* example 6 from Terms.thy *) |
297 (* example 6 from Terms.thy *) |
216 |
298 |
|
299 (* BV is not respectful, needs to fail |
217 nominal_datatype trm6 = |
300 nominal_datatype trm6 = |
218 Vr6 "name" |
301 Vr6 "name" |
219 | Lm6 x::"name" t::"trm6" bind x in t |
302 | Lm6 x::"name" t::"trm6" bind x in t |
220 | Lt6 left::"trm6" right::"trm6" bind "bv6 left" in right |
303 | Lt6 left::"trm6" right::"trm6" bind "bv6 left" in right |
221 binder |
304 binder |
222 bv6 |
305 bv6 |
223 where |
306 where |
224 "bv6 (Vr6 n) = {}" |
307 "bv6 (Vr6 n) = {}" |
225 | "bv6 (Lm6 n t) = {atom n} \<union> bv6 t" |
308 | "bv6 (Lm6 n t) = {atom n} \<union> bv6 t" |
226 | "bv6 (Lt6 l r) = bv6 l \<union> bv6 r" |
309 | "bv6 (Lt6 l r) = bv6 l \<union> bv6 r" *) |
227 |
310 |
228 (* example 7 from Terms.thy *) |
311 (* example 7 from Terms.thy *) |
229 |
312 |
|
313 (* BV is not respectful, needs to fail |
230 nominal_datatype trm7 = |
314 nominal_datatype trm7 = |
231 Vr7 "name" |
315 Vr7 "name" |
232 | Lm7 l::"name" r::"trm7" bind l in r |
316 | Lm7 l::"name" r::"trm7" bind l in r |
233 | Lt7 l::"trm7" r::"trm7" bind "bv7 l" in r |
317 | Lt7 l::"trm7" r::"trm7" bind "bv7 l" in r |
234 binder |
318 binder |
235 bv7 |
319 bv7 |
236 where |
320 where |
237 "bv7 (Vr7 n) = {atom n}" |
321 "bv7 (Vr7 n) = {atom n}" |
238 | "bv7 (Lm7 n t) = bv7 t - {atom n}" |
322 | "bv7 (Lm7 n t) = bv7 t - {atom n}" |
239 | "bv7 (Lt7 l r) = bv7 l \<union> bv7 r" |
323 | "bv7 (Lt7 l r) = bv7 l \<union> bv7 r" *) |
240 |
324 |
241 (* example 8 from Terms.thy *) |
325 (* example 8 from Terms.thy *) |
242 |
326 |
243 nominal_datatype foo8 = |
327 nominal_datatype foo8 = |
244 Foo0 "name" |
328 Foo0 "name" |
252 "bv8 (Bar0 x) = {}" |
336 "bv8 (Bar0 x) = {}" |
253 | "bv8 (Bar1 v x b) = {atom v}" |
337 | "bv8 (Bar1 v x b) = {atom v}" |
254 |
338 |
255 (* example 9 from Terms.thy *) |
339 (* example 9 from Terms.thy *) |
256 |
340 |
|
341 (* BV is not respectful, needs to fail |
257 nominal_datatype lam9 = |
342 nominal_datatype lam9 = |
258 Var9 "name" |
343 Var9 "name" |
259 | Lam9 n::"name" l::"lam9" bind n in l |
344 | Lam9 n::"name" l::"lam9" bind n in l |
260 and bla9 = |
345 and bla9 = |
261 Bla9 f::"lam9" s::"lam9" bind "bv9 f" in s |
346 Bla9 f::"lam9" s::"lam9" bind "bv9 f" in s |
262 binder |
347 binder |
263 bv9 |
348 bv9 |
264 where |
349 where |
265 "bv9 (Var9 x) = {}" |
350 "bv9 (Var9 x) = {}" |
266 | "bv9 (Lam9 x b) = {atom x}" |
351 | "bv9 (Lam9 x b) = {atom x}" *) |
267 |
352 |
268 (* example from my PHD *) |
353 (* example from my PHD *) |
269 |
354 |
270 atom_decl coname |
355 atom_decl coname |
271 |
356 |
272 nominal_datatype phd = |
357 (*nominal_datatype phd = |
273 Ax "name" "coname" |
358 Ax "name" "coname" |
274 | Cut n::"coname" t1::"phd" c::"coname" t2::"phd" bind n in t1, bind c in t2 |
359 | Cut n::"coname" t1::"phd" c::"coname" t2::"phd" bind n in t1, bind c in t2 |
275 | AndR c1::"coname" t1::"phd" c2::"coname" t2::"phd" "coname" bind c1 in t1, bind c2 in t2 |
360 | AndR c1::"coname" t1::"phd" c2::"coname" t2::"phd" "coname" bind c1 in t1, bind c2 in t2 |
276 | AndL1 n::"name" t::"phd" "name" bind n in t |
361 | AndL1 n::"name" t::"phd" "name" bind n in t |
277 | AndL2 n::"name" t::"phd" "name" bind n in t |
362 | AndL2 n::"name" t::"phd" "name" bind n in t |
278 | ImpL c::"coname" t1::"phd" n::"name" t2::"phd" "name" bind c in t1, bind n in t2 |
363 | ImpL c::"coname" t1::"phd" n::"name" t2::"phd" "name" bind c in t1, bind n in t2 |
279 | ImpR c::"coname" n::"name" t::"phd" "coname" bind n in t, bind c in t |
364 | ImpR c::"coname" n::"name" t::"phd" "coname" bind n in t, bind c in t |
280 |
365 |
281 (* PROBLEM?: why does it create for the Cut AndR ImpL cases |
|
282 two permutations, but only one is used *) |
|
283 thm alpha_phd_raw.intros[no_vars] |
366 thm alpha_phd_raw.intros[no_vars] |
284 thm fv_phd_raw.simps[no_vars] |
367 thm fv_phd_raw.simps[no_vars] |
|
368 *) |
285 |
369 |
286 (* example form Leroy 96 about modules; OTT *) |
370 (* example form Leroy 96 about modules; OTT *) |
287 |
371 |
288 nominal_datatype mexp = |
372 nominal_datatype mexp = |
289 Acc "path" |
373 Acc "path" |
318 Sref1 "name" |
402 Sref1 "name" |
319 | Sref2 "path" "name" |
403 | Sref2 "path" "name" |
320 and trmtrm = |
404 and trmtrm = |
321 Tref1 "name" |
405 Tref1 "name" |
322 | Tref2 "path" "name" |
406 | Tref2 "path" "name" |
323 | Lam v::"name" "tyty" M::"trmtrm" bind v in M |
407 | Lam' v::"name" "tyty" M::"trmtrm" bind v in M |
324 | App "trmtrm" "trmtrm" |
408 | App' "trmtrm" "trmtrm" |
325 | Let "body" "trmtrm" |
409 | Let' "body" "trmtrm" |
326 binder |
410 binder |
327 cbinders :: "defn \<Rightarrow> atom set" |
411 cbinders :: "defn \<Rightarrow> atom set" |
328 and Cbinders :: "spec \<Rightarrow> atom set" |
412 and Cbinders :: "spec \<Rightarrow> atom set" |
329 where |
413 where |
330 "cbinders (Type t T) = {atom t}" |
414 "cbinders (Type t T) = {atom t}" |
412 where |
497 where |
413 "bp (PVar x) = {atom x}" |
498 "bp (PVar x) = {atom x}" |
414 | "bp (PUnit) = {}" |
499 | "bp (PUnit) = {}" |
415 | "bp (PPair p1 p2) = bp p1 \<union> bp p2" |
500 | "bp (PPair p1 p2) = bp p1 \<union> bp p2" |
416 |
501 |
|
502 thm quot_respect |
417 (* example 6 from Peter Sewell's bestiary *) |
503 (* example 6 from Peter Sewell's bestiary *) |
418 nominal_datatype exp6 = |
504 (*nominal_datatype exp6 = |
419 EVar name |
505 EVar name |
420 | EPair exp6 exp6 |
506 | EPair exp6 exp6 |
421 | ELetRec x::name p::pat6 e1::exp6 e2::exp6 bind x in e1, bind x in e2, bind "bp6 p" in e1 |
507 | ELetRec x::name p::pat6 e1::exp6 e2::exp6 bind x in e1, bind x in e2, bind "bp6 p" in e1 |
422 and pat6 = |
508 and pat6 = |
423 PVar name |
509 PVar' name |
424 | PUnit |
510 | PUnit' |
425 | PPair pat6 pat6 |
511 | PPair' pat6 pat6 |
426 binder |
512 binder |
427 bp6 :: "pat6 \<Rightarrow> atom set" |
513 bp6 :: "pat6 \<Rightarrow> atom set" |
428 where |
514 where |
429 "bp6 (PVar x) = {atom x}" |
515 "bp6 (PVar' x) = {atom x}" |
430 | "bp6 (PUnit) = {}" |
516 | "bp6 (PUnit') = {}" |
431 | "bp6 (PPair p1 p2) = bp6 p1 \<union> bp6 p2" |
517 | "bp6 (PPair' p1 p2) = bp6 p1 \<union> bp6 p2"*) |
432 |
518 |
433 (* example 7 from Peter Sewell's bestiary *) |
519 (* example 7 from Peter Sewell's bestiary *) |
434 nominal_datatype exp7 = |
520 (*nominal_datatype exp7 = |
435 EVar name |
521 EVar name |
436 | EUnit |
522 | EUnit |
437 | EPair exp7 exp7 |
523 | EPair exp7 exp7 |
438 | ELetRec l::lrbs e::exp7 bind "b7s l" in e, bind "b7s l" in l |
524 | ELetRec l::lrbs e::exp7 bind "b7s l" in e, bind "b7s l" in l |
439 and lrb = |
525 and lrb = |
445 b7 :: "lrb \<Rightarrow> atom set" and |
531 b7 :: "lrb \<Rightarrow> atom set" and |
446 b7s :: "lrbs \<Rightarrow> atom set" |
532 b7s :: "lrbs \<Rightarrow> atom set" |
447 where |
533 where |
448 "b7 (Assign x e) = {atom x}" |
534 "b7 (Assign x e) = {atom x}" |
449 | "b7s (Single a) = b7 a" |
535 | "b7s (Single a) = b7 a" |
450 | "b7s (More a as) = (b7 a) \<union> (b7s as)" |
536 | "b7s (More a as) = (b7 a) \<union> (b7s as)"*) |
451 |
537 |
452 (* example 8 from Peter Sewell's bestiary *) |
538 (* example 8 from Peter Sewell's bestiary *) |
453 nominal_datatype exp8 = |
539 (*nominal_datatype exp8 = |
454 EVar name |
540 EVar name |
455 | EUnit |
541 | EUnit |
456 | EPair exp8 exp8 |
542 | EPair exp8 exp8 |
457 | ELetRec l::lrbs8 e::exp8 bind "b_lrbs8 l" in e, bind "b_lrbs8 l" in l |
543 | ELetRec l::lrbs8 e::exp8 bind "b_lrbs8 l" in e, bind "b_lrbs8 l" in l |
458 and fnclause = |
544 and fnclause = |
482 | "b_pat (PUnit) = {}" |
568 | "b_pat (PUnit) = {}" |
483 | "b_pat (PPair p1 p2) = b_pat p1 \<union> b_pat p2" |
569 | "b_pat (PPair p1 p2) = b_pat p1 \<union> b_pat p2" |
484 | "b_fnclauses (S fc) = (b_fnclause fc)" |
570 | "b_fnclauses (S fc) = (b_fnclause fc)" |
485 | "b_fnclauses (ORs fc fcs) = (b_fnclause fc) \<union> (b_fnclauses fcs)" |
571 | "b_fnclauses (ORs fc fcs) = (b_fnclause fc) \<union> (b_fnclauses fcs)" |
486 | "b_lrb8 (Clause fcs) = (b_fnclauses fcs)" |
572 | "b_lrb8 (Clause fcs) = (b_fnclauses fcs)" |
487 | "b_fnclause (K x pat exp8) = {atom x}" |
573 | "b_fnclause (K x pat exp8) = {atom x}"*) |
488 |
574 |
489 |
575 |
490 |
576 |
491 |
577 |
492 (* example 9 from Peter Sewell's bestiary *) |
578 (* example 9 from Peter Sewell's bestiary *) |