2 imports "Parser" "../Attic/Prove" |
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 |
21 |
21 |
22 local_setup {* |
22 local_setup {* |
23 (fn ctxt => snd (Local_Theory.note ((@{binding alpha_eqvt}, []), |
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)) *} |
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 |
25 |
26 ML {* |
26 (*prove {* (snd o snd) (build_alpha_refl_gl [@{term alpha_weird_raw}] ("x","y","z")) *} |
27 fun is_ex (Const ("Ex", _) $ Abs _) = true |
27 |
28 | is_ex _ = false; |
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 *} |
29 *) |
30 |
|
31 ML {* |
|
32 fun eetac rule = Subgoal.FOCUS_PARAMS |
|
33 (fn (focus) => |
|
34 let |
|
35 val concl = #concl focus |
|
36 val prems = Logic.strip_imp_prems (term_of concl) |
|
37 val exs = filter (fn x => is_ex (HOLogic.dest_Trueprop x)) prems |
|
38 val cexs = map (SOME o (cterm_of (ProofContext.theory_of (#context focus)))) exs |
|
39 val thins = map (fn cex => Drule.instantiate' [] [cex] Drule.thin_rl) cexs |
|
40 in |
|
41 (etac rule THEN' RANGE[ |
|
42 atac, |
|
43 eresolve_tac thins |
|
44 ]) 1 |
|
45 end |
|
46 ) |
|
47 *} |
|
48 |
|
49 ML {* |
|
50 fun transp_tac ctxt induct alpha_inj term_inj distinct cases eqvt = |
|
51 ind_tac induct THEN_ALL_NEW |
|
52 (TRY o rtac allI THEN' imp_elim_tac cases ctxt) THEN_ALL_NEW |
|
53 asm_full_simp_tac ((mk_minimal_ss ctxt) addsimps alpha_inj) THEN_ALL_NEW |
|
54 split_conjs THEN_ALL_NEW REPEAT o (eetac @{thm exi_sum} ctxt) |
|
55 THEN_ALL_NEW split_conjs |
|
56 *} |
|
57 (*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 *})*) |
|
58 |
|
59 lemma "alpha_weird_raw x y \<longrightarrow> (\<forall>z. alpha_weird_raw y z \<longrightarrow> alpha_weird_raw x z)" |
30 lemma "alpha_weird_raw x y \<longrightarrow> (\<forall>z. alpha_weird_raw y z \<longrightarrow> alpha_weird_raw x z)" |
60 apply (rule impI) |
31 apply (rule impI) |
61 apply (erule alpha_weird_raw.induct) |
32 apply (erule alpha_weird_raw.induct) |
62 apply (simp_all add: weird_inj) |
33 apply (simp_all add: weird_inj) |
63 defer |
34 defer |
88 defer |
59 defer |
89 apply (rule_tac x="pid + pi" in exI) |
60 apply (rule_tac x="pid + pi" in exI) |
90 apply (erule alpha_gen_compose_trans) |
61 apply (erule alpha_gen_compose_trans) |
91 apply assumption |
62 apply assumption |
92 apply (simp add: alpha_eqvt) |
63 apply (simp add: alpha_eqvt) |
93 done |
64 sorry |
94 |
65 |
95 lemma "alpha_weird_raw x y \<Longrightarrow> alpha_weird_raw y x" |
66 lemma "alpha_weird_raw x y \<Longrightarrow> alpha_weird_raw y x" |
96 apply (erule alpha_weird_raw.induct) |
67 apply (erule alpha_weird_raw.induct) |
97 apply (simp_all add: weird_inj) |
68 apply (simp_all add: weird_inj) |
98 apply (erule conjE)+ |
69 apply (erule conjE)+ |
99 apply (erule exE)+ |
70 apply (erule exE)+ |
|
71 apply (erule conjE)+ |
|
72 apply (erule exE)+ |
100 apply (rule conjI) |
73 apply (rule conjI) |
101 defer (* simple *) |
74 apply (rule_tac x="- pic" in exI) |
102 apply (rule conjI) |
75 apply (erule alpha_gen_compose_sym) |
|
76 apply (simp_all add: alpha_eqvt) |
103 apply (rule_tac x="- pia" in exI) |
77 apply (rule_tac x="- pia" in exI) |
104 apply (rule_tac x="- pib" in exI) |
78 apply (rule_tac x="- pib" in exI) |
105 apply (simp add: minus_add[symmetric]) |
79 apply (simp add: minus_add[symmetric]) |
|
80 apply (rule conjI) |
106 apply (erule alpha_gen_compose_sym) |
81 apply (erule alpha_gen_compose_sym) |
107 apply (simp_all add: alpha_eqvt) |
82 apply (simp_all add: alpha_eqvt) |
|
83 apply (rule conjI) |
|
84 apply (simp add: supp_minus_perm Int_commute) |
108 apply (rule_tac x="- pi" in exI) |
85 apply (rule_tac x="- pi" in exI) |
109 apply (erule alpha_gen_compose_sym) |
|
110 apply (simp_all add: alpha_eqvt) |
|
111 apply (rule_tac x="- pic" in exI) |
|
112 apply (erule alpha_gen_compose_sym) |
86 apply (erule alpha_gen_compose_sym) |
113 apply (simp_all add: alpha_eqvt) |
87 apply (simp_all add: alpha_eqvt) |
114 done |
88 done |
115 |
89 |
116 |
90 |
311 "bv5 Lnil = {}" |
294 "bv5 Lnil = {}" |
312 | "bv5 (Lcons n t ltl) = {atom n} \<union> (bv5 ltl)" |
295 | "bv5 (Lcons n t ltl) = {atom n} \<union> (bv5 ltl)" |
313 |
296 |
314 (* example 6 from Terms.thy *) |
297 (* example 6 from Terms.thy *) |
315 |
298 |
|
299 (* BV is not respectful, needs to fail |
316 nominal_datatype trm6 = |
300 nominal_datatype trm6 = |
317 Vr6 "name" |
301 Vr6 "name" |
318 | Lm6 x::"name" t::"trm6" bind x in t |
302 | Lm6 x::"name" t::"trm6" bind x in t |
319 | Lt6 left::"trm6" right::"trm6" bind "bv6 left" in right |
303 | Lt6 left::"trm6" right::"trm6" bind "bv6 left" in right |
320 binder |
304 binder |
321 bv6 |
305 bv6 |
322 where |
306 where |
323 "bv6 (Vr6 n) = {}" |
307 "bv6 (Vr6 n) = {}" |
324 | "bv6 (Lm6 n t) = {atom n} \<union> bv6 t" |
308 | "bv6 (Lm6 n t) = {atom n} \<union> bv6 t" |
325 | "bv6 (Lt6 l r) = bv6 l \<union> bv6 r" |
309 | "bv6 (Lt6 l r) = bv6 l \<union> bv6 r" *) |
326 |
310 |
327 (* example 7 from Terms.thy *) |
311 (* example 7 from Terms.thy *) |
328 |
312 |
|
313 (* BV is not respectful, needs to fail |
329 nominal_datatype trm7 = |
314 nominal_datatype trm7 = |
330 Vr7 "name" |
315 Vr7 "name" |
331 | Lm7 l::"name" r::"trm7" bind l in r |
316 | Lm7 l::"name" r::"trm7" bind l in r |
332 | Lt7 l::"trm7" r::"trm7" bind "bv7 l" in r |
317 | Lt7 l::"trm7" r::"trm7" bind "bv7 l" in r |
333 binder |
318 binder |
334 bv7 |
319 bv7 |
335 where |
320 where |
336 "bv7 (Vr7 n) = {atom n}" |
321 "bv7 (Vr7 n) = {atom n}" |
337 | "bv7 (Lm7 n t) = bv7 t - {atom n}" |
322 | "bv7 (Lm7 n t) = bv7 t - {atom n}" |
338 | "bv7 (Lt7 l r) = bv7 l \<union> bv7 r" |
323 | "bv7 (Lt7 l r) = bv7 l \<union> bv7 r" *) |
339 |
324 |
340 (* example 8 from Terms.thy *) |
325 (* example 8 from Terms.thy *) |
341 |
326 |
342 nominal_datatype foo8 = |
327 nominal_datatype foo8 = |
343 Foo0 "name" |
328 Foo0 "name" |
351 "bv8 (Bar0 x) = {}" |
336 "bv8 (Bar0 x) = {}" |
352 | "bv8 (Bar1 v x b) = {atom v}" |
337 | "bv8 (Bar1 v x b) = {atom v}" |
353 |
338 |
354 (* example 9 from Terms.thy *) |
339 (* example 9 from Terms.thy *) |
355 |
340 |
|
341 (* BV is not respectful, needs to fail |
356 nominal_datatype lam9 = |
342 nominal_datatype lam9 = |
357 Var9 "name" |
343 Var9 "name" |
358 | Lam9 n::"name" l::"lam9" bind n in l |
344 | Lam9 n::"name" l::"lam9" bind n in l |
359 and bla9 = |
345 and bla9 = |
360 Bla9 f::"lam9" s::"lam9" bind "bv9 f" in s |
346 Bla9 f::"lam9" s::"lam9" bind "bv9 f" in s |
361 binder |
347 binder |
362 bv9 |
348 bv9 |
363 where |
349 where |
364 "bv9 (Var9 x) = {}" |
350 "bv9 (Var9 x) = {}" |
365 | "bv9 (Lam9 x b) = {atom x}" |
351 | "bv9 (Lam9 x b) = {atom x}" *) |
366 |
352 |
367 (* example from my PHD *) |
353 (* example from my PHD *) |
368 |
354 |
369 atom_decl coname |
355 atom_decl coname |
370 |
356 |
371 nominal_datatype phd = |
357 (*nominal_datatype phd = |
372 Ax "name" "coname" |
358 Ax "name" "coname" |
373 | 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 |
374 | 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 |
375 | AndL1 n::"name" t::"phd" "name" bind n in t |
361 | AndL1 n::"name" t::"phd" "name" bind n in t |
376 | AndL2 n::"name" t::"phd" "name" bind n in t |
362 | AndL2 n::"name" t::"phd" "name" bind n in t |
377 | 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 |
378 | 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 |
379 |
365 |
380 (* PROBLEM?: why does it create for the Cut AndR ImpL cases |
|
381 two permutations, but only one is used *) |
|
382 thm alpha_phd_raw.intros[no_vars] |
366 thm alpha_phd_raw.intros[no_vars] |
383 thm fv_phd_raw.simps[no_vars] |
367 thm fv_phd_raw.simps[no_vars] |
|
368 *) |
384 |
369 |
385 (* example form Leroy 96 about modules; OTT *) |
370 (* example form Leroy 96 about modules; OTT *) |
386 |
371 |
387 nominal_datatype mexp = |
372 nominal_datatype mexp = |
388 Acc "path" |
373 Acc "path" |
417 Sref1 "name" |
402 Sref1 "name" |
418 | Sref2 "path" "name" |
403 | Sref2 "path" "name" |
419 and trmtrm = |
404 and trmtrm = |
420 Tref1 "name" |
405 Tref1 "name" |
421 | Tref2 "path" "name" |
406 | Tref2 "path" "name" |
422 | Lam v::"name" "tyty" M::"trmtrm" bind v in M |
407 | Lam' v::"name" "tyty" M::"trmtrm" bind v in M |
423 | App "trmtrm" "trmtrm" |
408 | App' "trmtrm" "trmtrm" |
424 | Let "body" "trmtrm" |
409 | Let' "body" "trmtrm" |
425 binder |
410 binder |
426 cbinders :: "defn \<Rightarrow> atom set" |
411 cbinders :: "defn \<Rightarrow> atom set" |
427 and Cbinders :: "spec \<Rightarrow> atom set" |
412 and Cbinders :: "spec \<Rightarrow> atom set" |
428 where |
413 where |
429 "cbinders (Type t T) = {atom t}" |
414 "cbinders (Type t T) = {atom t}" |
508 where |
494 where |
509 "bp (PVar x) = {atom x}" |
495 "bp (PVar x) = {atom x}" |
510 | "bp (PUnit) = {}" |
496 | "bp (PUnit) = {}" |
511 | "bp (PPair p1 p2) = bp p1 \<union> bp p2" |
497 | "bp (PPair p1 p2) = bp p1 \<union> bp p2" |
512 |
498 |
|
499 thm quot_respect |
513 (* example 6 from Peter Sewell's bestiary *) |
500 (* example 6 from Peter Sewell's bestiary *) |
514 nominal_datatype exp6 = |
501 (*nominal_datatype exp6 = |
515 EVar name |
502 EVar name |
516 | EPair exp6 exp6 |
503 | EPair exp6 exp6 |
517 | ELetRec x::name p::pat6 e1::exp6 e2::exp6 bind x in e1, bind x in e2, bind "bp6 p" in e1 |
504 | ELetRec x::name p::pat6 e1::exp6 e2::exp6 bind x in e1, bind x in e2, bind "bp6 p" in e1 |
518 and pat6 = |
505 and pat6 = |
519 PVar name |
506 PVar' name |
520 | PUnit |
507 | PUnit' |
521 | PPair pat6 pat6 |
508 | PPair' pat6 pat6 |
522 binder |
509 binder |
523 bp6 :: "pat6 \<Rightarrow> atom set" |
510 bp6 :: "pat6 \<Rightarrow> atom set" |
524 where |
511 where |
525 "bp6 (PVar x) = {atom x}" |
512 "bp6 (PVar' x) = {atom x}" |
526 | "bp6 (PUnit) = {}" |
513 | "bp6 (PUnit') = {}" |
527 | "bp6 (PPair p1 p2) = bp6 p1 \<union> bp6 p2" |
514 | "bp6 (PPair' p1 p2) = bp6 p1 \<union> bp6 p2"*) |
528 |
515 |
529 (* example 7 from Peter Sewell's bestiary *) |
516 (* example 7 from Peter Sewell's bestiary *) |
530 nominal_datatype exp7 = |
517 (*nominal_datatype exp7 = |
531 EVar name |
518 EVar name |
532 | EUnit |
519 | EUnit |
533 | EPair exp7 exp7 |
520 | EPair exp7 exp7 |
534 | ELetRec l::lrbs e::exp7 bind "b7s l" in e, bind "b7s l" in l |
521 | ELetRec l::lrbs e::exp7 bind "b7s l" in e, bind "b7s l" in l |
535 and lrb = |
522 and lrb = |
541 b7 :: "lrb \<Rightarrow> atom set" and |
528 b7 :: "lrb \<Rightarrow> atom set" and |
542 b7s :: "lrbs \<Rightarrow> atom set" |
529 b7s :: "lrbs \<Rightarrow> atom set" |
543 where |
530 where |
544 "b7 (Assign x e) = {atom x}" |
531 "b7 (Assign x e) = {atom x}" |
545 | "b7s (Single a) = b7 a" |
532 | "b7s (Single a) = b7 a" |
546 | "b7s (More a as) = (b7 a) \<union> (b7s as)" |
533 | "b7s (More a as) = (b7 a) \<union> (b7s as)"*) |
547 |
534 |
548 (* example 8 from Peter Sewell's bestiary *) |
535 (* example 8 from Peter Sewell's bestiary *) |
549 nominal_datatype exp8 = |
536 (*nominal_datatype exp8 = |
550 EVar name |
537 EVar name |
551 | EUnit |
538 | EUnit |
552 | EPair exp8 exp8 |
539 | EPair exp8 exp8 |
553 | ELetRec l::lrbs8 e::exp8 bind "b_lrbs8 l" in e, bind "b_lrbs8 l" in l |
540 | ELetRec l::lrbs8 e::exp8 bind "b_lrbs8 l" in e, bind "b_lrbs8 l" in l |
554 and fnclause = |
541 and fnclause = |
578 | "b_pat (PUnit) = {}" |
565 | "b_pat (PUnit) = {}" |
579 | "b_pat (PPair p1 p2) = b_pat p1 \<union> b_pat p2" |
566 | "b_pat (PPair p1 p2) = b_pat p1 \<union> b_pat p2" |
580 | "b_fnclauses (S fc) = (b_fnclause fc)" |
567 | "b_fnclauses (S fc) = (b_fnclause fc)" |
581 | "b_fnclauses (ORs fc fcs) = (b_fnclause fc) \<union> (b_fnclauses fcs)" |
568 | "b_fnclauses (ORs fc fcs) = (b_fnclause fc) \<union> (b_fnclauses fcs)" |
582 | "b_lrb8 (Clause fcs) = (b_fnclauses fcs)" |
569 | "b_lrb8 (Clause fcs) = (b_fnclauses fcs)" |
583 | "b_fnclause (K x pat exp8) = {atom x}" |
570 | "b_fnclause (K x pat exp8) = {atom x}"*) |
584 |
571 |
585 |
572 |
586 |
573 |
587 |
574 |
588 (* example 9 from Peter Sewell's bestiary *) |
575 (* example 9 from Peter Sewell's bestiary *) |