| author | Cezary Kaliszyk <kaliszyk@in.tum.de> | 
| Tue, 08 Nov 2011 22:31:31 +0000 | |
| changeset 3047 | 014edadaeb59 | 
| parent 3046 | 9b0324e1293b | 
| child 3065 | 51ef8a3cb6ef | 
| permissions | -rw-r--r-- | 
| 1797 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 Christian Urban <urbanc@in.tum.de> parents: 
1773diff
changeset | 1 | theory Lambda | 
| 2950 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 2 | imports | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 3 | "../Nominal2" | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 4 | "~~/src/HOL/Library/Monad_Syntax" | 
| 1594 | 5 | begin | 
| 6 | ||
| 2885 
1264f2a21ea9
some rudimentary infrastructure for storing data about nominal datatypes
 Christian Urban <urbanc@in.tum.de> parents: 
2868diff
changeset | 7 | |
| 1594 | 8 | atom_decl name | 
| 9 | ||
| 2436 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 10 | nominal_datatype lam = | 
| 1800 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 11 | Var "name" | 
| 2436 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 12 | | App "lam" "lam" | 
| 2950 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 13 | | Lam x::"name" l::"lam"  binds x in l ("Lam [_]. _" [100, 100] 100)
 | 
| 2431 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2425diff
changeset | 14 | |
| 3047 
014edadaeb59
Add equivariance for alpha_lam_raw and abs_lam.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
3046diff
changeset | 15 | lemma alpha_lam_raw_eqvt[eqvt]: "p \<bullet> (alpha_lam_raw x y) = alpha_lam_raw (p \<bullet> x) (p \<bullet> y)" | 
| 
014edadaeb59
Add equivariance for alpha_lam_raw and abs_lam.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
3046diff
changeset | 16 | unfolding alpha_lam_raw_def | 
| 
014edadaeb59
Add equivariance for alpha_lam_raw and abs_lam.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
3046diff
changeset | 17 | by perm_simp rule | 
| 
014edadaeb59
Add equivariance for alpha_lam_raw and abs_lam.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
3046diff
changeset | 18 | |
| 
014edadaeb59
Add equivariance for alpha_lam_raw and abs_lam.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
3046diff
changeset | 19 | lemma abs_lam_eqvt[eqvt]: "(p \<bullet> abs_lam t) = abs_lam (p \<bullet> t)" | 
| 
014edadaeb59
Add equivariance for alpha_lam_raw and abs_lam.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
3046diff
changeset | 20 | proof - | 
| 
014edadaeb59
Add equivariance for alpha_lam_raw and abs_lam.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
3046diff
changeset | 21 | have "alpha_lam_raw (rep_lam (abs_lam t)) t" | 
| 
014edadaeb59
Add equivariance for alpha_lam_raw and abs_lam.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
3046diff
changeset | 22 | using rep_abs_rsp_left Quotient_lam equivp_reflp lam_equivp by metis | 
| 
014edadaeb59
Add equivariance for alpha_lam_raw and abs_lam.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
3046diff
changeset | 23 | then have "alpha_lam_raw (p \<bullet> rep_lam (abs_lam t)) (p \<bullet> t)" | 
| 
014edadaeb59
Add equivariance for alpha_lam_raw and abs_lam.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
3046diff
changeset | 24 | unfolding alpha_lam_raw_eqvt[symmetric] permute_pure . | 
| 
014edadaeb59
Add equivariance for alpha_lam_raw and abs_lam.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
3046diff
changeset | 25 | then have "abs_lam (p \<bullet> rep_lam (abs_lam t)) = abs_lam (p \<bullet> t)" | 
| 
014edadaeb59
Add equivariance for alpha_lam_raw and abs_lam.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
3046diff
changeset | 26 | using Quotient_rel Quotient_lam by metis | 
| 
014edadaeb59
Add equivariance for alpha_lam_raw and abs_lam.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
3046diff
changeset | 27 | thus ?thesis using permute_lam_def id_apply map_fun_apply by metis | 
| 
014edadaeb59
Add equivariance for alpha_lam_raw and abs_lam.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
3046diff
changeset | 28 | qed | 
| 
014edadaeb59
Add equivariance for alpha_lam_raw and abs_lam.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
3046diff
changeset | 29 | |
| 
014edadaeb59
Add equivariance for alpha_lam_raw and abs_lam.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
3046diff
changeset | 30 | |
| 2951 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 31 | section {* Simple examples from Norrish 2004 *}
 | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 32 | |
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 33 | nominal_primrec | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 34 | is_app :: "lam \<Rightarrow> bool" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 35 | where | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 36 | "is_app (Var x) = False" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 37 | | "is_app (App t1 t2) = True" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 38 | | "is_app (Lam [x]. t) = False" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 39 | apply(simp add: eqvt_def is_app_graph_def) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 40 | apply (rule, perm_simp, rule) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 41 | apply(rule TrueI) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 42 | apply(rule_tac y="x" in lam.exhaust) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 43 | apply(auto)[3] | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 44 | apply(all_trivials) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 45 | done | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 46 | |
| 2973 
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
 Christian Urban <urbanc@in.tum.de> parents: 
2972diff
changeset | 47 | termination (eqvt) by lexicographic_order | 
| 
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
 Christian Urban <urbanc@in.tum.de> parents: 
2972diff
changeset | 48 | |
| 2974 
b95a2065aa10
generated the partial eqvt-theorem for functions
 Christian Urban <urbanc@in.tum.de> parents: 
2973diff
changeset | 49 | thm is_app_def | 
| 2975 
c62e26830420
preliminary version of automatically generation the eqvt-lemmas for functions defined with nominal_primrec
 Christian Urban <urbanc@in.tum.de> parents: 
2974diff
changeset | 50 | thm is_app.eqvt | 
| 2973 
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
 Christian Urban <urbanc@in.tum.de> parents: 
2972diff
changeset | 51 | |
| 2975 
c62e26830420
preliminary version of automatically generation the eqvt-lemmas for functions defined with nominal_primrec
 Christian Urban <urbanc@in.tum.de> parents: 
2974diff
changeset | 52 | thm eqvts | 
| 2951 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 53 | |
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 54 | nominal_primrec | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 55 | rator :: "lam \<Rightarrow> lam option" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 56 | where | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 57 | "rator (Var x) = None" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 58 | | "rator (App t1 t2) = Some t1" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 59 | | "rator (Lam [x]. t) = None" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 60 | apply(simp add: eqvt_def rator_graph_def) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 61 | apply (rule, perm_simp, rule) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 62 | apply(rule TrueI) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 63 | apply(rule_tac y="x" in lam.exhaust) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 64 | apply(auto)[3] | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 65 | apply(simp_all) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 66 | done | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 67 | |
| 2973 
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
 Christian Urban <urbanc@in.tum.de> parents: 
2972diff
changeset | 68 | termination (eqvt) by lexicographic_order | 
| 2951 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 69 | |
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 70 | nominal_primrec | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 71 | rand :: "lam \<Rightarrow> lam option" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 72 | where | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 73 | "rand (Var x) = None" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 74 | | "rand (App t1 t2) = Some t2" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 75 | | "rand (Lam [x]. t) = None" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 76 | apply(simp add: eqvt_def rand_graph_def) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 77 | apply (rule, perm_simp, rule) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 78 | apply(rule TrueI) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 79 | apply(rule_tac y="x" in lam.exhaust) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 80 | apply(auto)[3] | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 81 | apply(simp_all) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 82 | done | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 83 | |
| 2973 
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
 Christian Urban <urbanc@in.tum.de> parents: 
2972diff
changeset | 84 | termination (eqvt) by lexicographic_order | 
| 2951 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 85 | |
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 86 | nominal_primrec | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 87 | is_eta_nf :: "lam \<Rightarrow> bool" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 88 | where | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 89 | "is_eta_nf (Var x) = True" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 90 | | "is_eta_nf (App t1 t2) = (is_eta_nf t1 \<and> is_eta_nf t2)" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 91 | | "is_eta_nf (Lam [x]. t) = (is_eta_nf t \<and> | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 92 | ((is_app t \<and> rand t = Some (Var x)) \<longrightarrow> atom x \<in> supp (rator t)))" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 93 | apply(simp add: eqvt_def is_eta_nf_graph_def) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 94 | apply (rule, perm_simp, rule) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 95 | apply(rule TrueI) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 96 | apply(rule_tac y="x" in lam.exhaust) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 97 | apply(auto)[3] | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 98 | apply(simp_all) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 99 | apply(erule_tac c="()" in Abs_lst1_fcb2') | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 100 | apply(simp_all add: pure_fresh fresh_star_def)[3] | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 101 | apply(simp add: eqvt_at_def conj_eqvt) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 102 | apply(perm_simp) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 103 | apply(rule refl) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 104 | apply(simp add: eqvt_at_def conj_eqvt) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 105 | apply(perm_simp) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 106 | apply(rule refl) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 107 | done | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 108 | |
| 2973 
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
 Christian Urban <urbanc@in.tum.de> parents: 
2972diff
changeset | 109 | termination (eqvt) by lexicographic_order | 
| 2951 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 110 | |
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 111 | nominal_datatype path = Left | Right | In | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 112 | |
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 113 | section {* Paths to a free variables *} 
 | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 114 | |
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 115 | instance path :: pure | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 116 | apply(default) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 117 | apply(induct_tac "x::path" rule: path.induct) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 118 | apply(simp_all) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 119 | done | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 120 | |
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 121 | nominal_primrec | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 122 | var_pos :: "name \<Rightarrow> lam \<Rightarrow> (path list) set" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 123 | where | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 124 |   "var_pos y (Var x) = (if y = x then {[]} else {})"
 | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 125 | | "var_pos y (App t1 t2) = (Cons Left ` (var_pos y t1)) \<union> (Cons Right ` (var_pos y t2))" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 126 | | "atom x \<sharp> y \<Longrightarrow> var_pos y (Lam [x]. t) = (Cons In ` (var_pos y t))" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 127 | apply(simp add: eqvt_def var_pos_graph_def) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 128 | apply (rule, perm_simp, rule) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 129 | apply(rule TrueI) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 130 | apply(case_tac x) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 131 | apply(rule_tac y="b" and c="a" in lam.strong_exhaust) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 132 | apply(auto simp add: fresh_star_def)[3] | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 133 | apply(simp_all) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 134 | apply(erule conjE)+ | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 135 | apply(erule_tac Abs_lst1_fcb2) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 136 | apply(simp add: pure_fresh) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 137 | apply(simp add: fresh_star_def) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 138 | apply(simp add: eqvt_at_def image_eqvt perm_supp_eq) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 139 | apply(perm_simp) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 140 | apply(rule refl) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 141 | apply(simp add: eqvt_at_def image_eqvt perm_supp_eq) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 142 | apply(perm_simp) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 143 | apply(rule refl) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 144 | done | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 145 | |
| 2973 
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
 Christian Urban <urbanc@in.tum.de> parents: 
2972diff
changeset | 146 | termination (eqvt) by lexicographic_order | 
| 2951 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 147 | |
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 148 | lemma var_pos1: | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 149 | assumes "atom y \<notin> supp t" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 150 |   shows "var_pos y t = {}"
 | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 151 | using assms | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 152 | apply(induct t rule: var_pos.induct) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 153 | apply(simp_all add: lam.supp supp_at_base fresh_at_base) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 154 | done | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 155 | |
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 156 | lemma var_pos2: | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 157 |   shows "var_pos y (Lam [y].t) = {}"
 | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 158 | apply(rule var_pos1) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 159 | apply(simp add: lam.supp) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 160 | done | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 161 | |
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 162 | |
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 163 | text {* strange substitution operation *}
 | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 164 | |
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 165 | nominal_primrec | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 166 |   subst' :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam"  ("_ [_ ::== _]" [90, 90, 90] 90)
 | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 167 | where | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 168 | "(Var x)[y ::== s] = (if x = y then s else (Var x))" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 169 | | "(App t1 t2)[y ::== s] = App (t1[y ::== s]) (t2[y ::== s])" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 170 | | "atom x \<sharp> (y, s) \<Longrightarrow> (Lam [x]. t)[y ::== s] = Lam [x].(t[y ::== (App (Var y) s)])" | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 171 | apply(simp add: eqvt_def subst'_graph_def) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 172 | apply (rule, perm_simp, rule) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 173 | apply(rule TrueI) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 174 | apply(case_tac x) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 175 | apply(rule_tac y="a" and c="(b, c)" in lam.strong_exhaust) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 176 | apply(auto simp add: fresh_star_def)[3] | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 177 | apply(simp_all) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 178 | apply(erule conjE)+ | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 179 | apply (erule_tac c="(ya,sa)" in Abs_lst1_fcb2) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 180 | apply(simp_all add: Abs_fresh_iff) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 181 | apply(simp add: fresh_star_def fresh_Pair) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 182 | apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 183 | apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq) | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 184 | done | 
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 185 | |
| 2973 
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
 Christian Urban <urbanc@in.tum.de> parents: 
2972diff
changeset | 186 | termination (eqvt) by lexicographic_order | 
| 2951 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 187 | |
| 
d75b3d8529e7
added some relatively simple examples from paper by Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
2950diff
changeset | 188 | |
| 2868 
2b8e387d2dfc
got rid of the boolean flag in the raw_equivariance function
 Christian Urban <urbanc@in.tum.de> parents: 
2860diff
changeset | 189 | section {* free name function *}
 | 
| 2858 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 190 | |
| 2860 
25a7f421a3ba
added a test that every function must be of pt-sort
 Christian Urban <urbanc@in.tum.de> parents: 
2858diff
changeset | 191 | text {* first returns an atom list *}
 | 
| 2858 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 192 | |
| 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 193 | nominal_primrec | 
| 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 194 | frees_lst :: "lam \<Rightarrow> atom list" | 
| 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 195 | where | 
| 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 196 | "frees_lst (Var x) = [atom x]" | 
| 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 197 | | "frees_lst (App t1 t2) = frees_lst t1 @ frees_lst t2" | 
| 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 198 | | "frees_lst (Lam [x]. t) = removeAll (atom x) (frees_lst t)" | 
| 2868 
2b8e387d2dfc
got rid of the boolean flag in the raw_equivariance function
 Christian Urban <urbanc@in.tum.de> parents: 
2860diff
changeset | 199 | unfolding eqvt_def | 
| 
2b8e387d2dfc
got rid of the boolean flag in the raw_equivariance function
 Christian Urban <urbanc@in.tum.de> parents: 
2860diff
changeset | 200 | unfolding frees_lst_graph_def | 
| 2858 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 201 | apply (rule, perm_simp, rule) | 
| 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 202 | apply(rule TrueI) | 
| 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 203 | apply(rule_tac y="x" in lam.exhaust) | 
| 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 204 | apply(auto) | 
| 2902 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 205 | apply (erule_tac c="()" in Abs_lst1_fcb2) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 206 | apply(simp add: supp_removeAll fresh_def) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 207 | apply(simp add: fresh_star_def fresh_Unit) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 208 | apply(simp add: eqvt_at_def removeAll_eqvt atom_eqvt) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 209 | apply(simp add: eqvt_at_def removeAll_eqvt atom_eqvt) | 
| 2858 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 210 | done | 
| 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 211 | |
| 2973 
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
 Christian Urban <urbanc@in.tum.de> parents: 
2972diff
changeset | 212 | termination (eqvt) by lexicographic_order | 
| 2858 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 213 | |
| 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 214 | text {* a small test lemma *}
 | 
| 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 215 | lemma shows "supp t = set (frees_lst t)" | 
| 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 216 | by (induct t rule: frees_lst.induct) (simp_all add: lam.supp supp_at_base) | 
| 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 217 | |
| 2860 
25a7f421a3ba
added a test that every function must be of pt-sort
 Christian Urban <urbanc@in.tum.de> parents: 
2858diff
changeset | 218 | text {* second returns an atom set - therefore needs an invariant *}
 | 
| 2858 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 219 | |
| 2821 
c7d4bd9e89e0
fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
 Christian Urban <urbanc@in.tum.de> parents: 
2819diff
changeset | 220 | nominal_primrec (invariant "\<lambda>x (y::atom set). finite y") | 
| 
c7d4bd9e89e0
fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
 Christian Urban <urbanc@in.tum.de> parents: 
2819diff
changeset | 221 | frees_set :: "lam \<Rightarrow> atom set" | 
| 
c7d4bd9e89e0
fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
 Christian Urban <urbanc@in.tum.de> parents: 
2819diff
changeset | 222 | where | 
| 
c7d4bd9e89e0
fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
 Christian Urban <urbanc@in.tum.de> parents: 
2819diff
changeset | 223 |   "frees_set (Var x) = {atom x}"
 | 
| 
c7d4bd9e89e0
fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
 Christian Urban <urbanc@in.tum.de> parents: 
2819diff
changeset | 224 | | "frees_set (App t1 t2) = frees_set t1 \<union> frees_set t2" | 
| 
c7d4bd9e89e0
fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
 Christian Urban <urbanc@in.tum.de> parents: 
2819diff
changeset | 225 | | "frees_set (Lam [x]. t) = (frees_set t) - {atom x}"
 | 
| 2845 
a99f488a96bb
Optimized proofs and removed some garbage.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2843diff
changeset | 226 | apply(simp add: eqvt_def frees_set_graph_def) | 
| 
a99f488a96bb
Optimized proofs and removed some garbage.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2843diff
changeset | 227 | apply(rule, perm_simp, rule) | 
| 
a99f488a96bb
Optimized proofs and removed some garbage.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2843diff
changeset | 228 | apply(erule frees_set_graph.induct) | 
| 
a99f488a96bb
Optimized proofs and removed some garbage.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2843diff
changeset | 229 | apply(auto)[9] | 
| 
a99f488a96bb
Optimized proofs and removed some garbage.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2843diff
changeset | 230 | apply(rule_tac y="x" in lam.exhaust) | 
| 
a99f488a96bb
Optimized proofs and removed some garbage.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2843diff
changeset | 231 | apply(auto)[3] | 
| 
a99f488a96bb
Optimized proofs and removed some garbage.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2843diff
changeset | 232 | apply(simp) | 
| 2902 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 233 | apply(erule_tac c="()" in Abs_lst1_fcb2) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 234 | apply(simp add: fresh_minus_atom_set) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 235 | apply(simp add: fresh_star_def fresh_Unit) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 236 | apply(simp add: Diff_eqvt eqvt_at_def, perm_simp, rule refl) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 237 | apply(simp add: Diff_eqvt eqvt_at_def, perm_simp, rule refl) | 
| 2845 
a99f488a96bb
Optimized proofs and removed some garbage.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2843diff
changeset | 238 | done | 
| 2821 
c7d4bd9e89e0
fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
 Christian Urban <urbanc@in.tum.de> parents: 
2819diff
changeset | 239 | |
| 2973 
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
 Christian Urban <urbanc@in.tum.de> parents: 
2972diff
changeset | 240 | termination (eqvt) | 
| 2858 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 241 | by lexicographic_order | 
| 2821 
c7d4bd9e89e0
fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
 Christian Urban <urbanc@in.tum.de> parents: 
2819diff
changeset | 242 | |
| 2822 
23befefc6e73
cleaned ups a bit the examples with the invariant framework; exported nominal_function_config datatype into separate structure and file
 Christian Urban <urbanc@in.tum.de> parents: 
2821diff
changeset | 243 | lemma "frees_set t = supp t" | 
| 2845 
a99f488a96bb
Optimized proofs and removed some garbage.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2843diff
changeset | 244 | by (induct rule: frees_set.induct) (simp_all add: lam.supp supp_at_base) | 
| 2821 
c7d4bd9e89e0
fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
 Christian Urban <urbanc@in.tum.de> parents: 
2819diff
changeset | 245 | |
| 2868 
2b8e387d2dfc
got rid of the boolean flag in the raw_equivariance function
 Christian Urban <urbanc@in.tum.de> parents: 
2860diff
changeset | 246 | section {* height function *}
 | 
| 2431 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2425diff
changeset | 247 | |
| 2666 
324a5d1289a3
added a few examples of functions to Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2664diff
changeset | 248 | nominal_primrec | 
| 2678 
494b859bfc16
defined height as a function that returns an integer
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 249 | height :: "lam \<Rightarrow> int" | 
| 2666 
324a5d1289a3
added a few examples of functions to Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2664diff
changeset | 250 | where | 
| 2678 
494b859bfc16
defined height as a function that returns an integer
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 251 | "height (Var x) = 1" | 
| 2683 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 252 | | "height (App t1 t2) = max (height t1) (height t2) + 1" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 253 | | "height (Lam [x].t) = height t + 1" | 
| 2902 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 254 | apply(simp add: eqvt_def height_graph_def) | 
| 2791 
5d0875b7ed3e
Simple eqvt proofs with perm_simps for clarity
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2789diff
changeset | 255 | apply (rule, perm_simp, rule) | 
| 2902 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 256 | apply(rule TrueI) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 257 | apply(rule_tac y="x" in lam.exhaust) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 258 | apply(auto) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 259 | apply (erule_tac c="()" in Abs_lst1_fcb2) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 260 | apply(simp_all add: fresh_def pure_supp eqvt_at_def fresh_star_def) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 261 | done | 
| 2666 
324a5d1289a3
added a few examples of functions to Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2664diff
changeset | 262 | |
| 2973 
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
 Christian Urban <urbanc@in.tum.de> parents: 
2972diff
changeset | 263 | termination (eqvt) | 
| 2858 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 264 | by lexicographic_order | 
| 2666 
324a5d1289a3
added a few examples of functions to Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2664diff
changeset | 265 | |
| 2707 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2685diff
changeset | 266 | thm height.simps | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2685diff
changeset | 267 | |
| 2683 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 268 | |
| 2868 
2b8e387d2dfc
got rid of the boolean flag in the raw_equivariance function
 Christian Urban <urbanc@in.tum.de> parents: 
2860diff
changeset | 269 | section {* capture-avoiding substitution *}
 | 
| 2683 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 270 | |
| 2675 
68ccf847507d
defined properly substitution
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 271 | nominal_primrec | 
| 2683 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 272 |   subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam"  ("_ [_ ::= _]" [90, 90, 90] 90)
 | 
| 2675 
68ccf847507d
defined properly substitution
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 273 | where | 
| 2683 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 274 | "(Var x)[y ::= s] = (if x = y then s else (Var x))" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 275 | | "(App t1 t2)[y ::= s] = App (t1[y ::= s]) (t2[y ::= s])" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 276 | | "atom x \<sharp> (y, s) \<Longrightarrow> (Lam [x]. t)[y ::= s] = Lam [x].(t[y ::= s])" | 
| 2791 
5d0875b7ed3e
Simple eqvt proofs with perm_simps for clarity
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2789diff
changeset | 277 | unfolding eqvt_def subst_graph_def | 
| 
5d0875b7ed3e
Simple eqvt proofs with perm_simps for clarity
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2789diff
changeset | 278 | apply (rule, perm_simp, rule) | 
| 2902 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 279 | apply(rule TrueI) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 280 | apply(auto simp add: lam.distinct lam.eq_iff) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 281 | apply(rule_tac y="a" and c="(aa, b)" in lam.strong_exhaust) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 282 | apply(blast)+ | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 283 | apply(simp_all add: fresh_star_def fresh_Pair_elim) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 284 | apply (erule_tac c="(ya,sa)" in Abs_lst1_fcb2) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 285 | apply(simp_all add: Abs_fresh_iff) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 286 | apply(simp add: fresh_star_def fresh_Pair) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 287 | apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 288 | apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq) | 
| 2675 
68ccf847507d
defined properly substitution
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 289 | done | 
| 
68ccf847507d
defined properly substitution
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 290 | |
| 2973 
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
 Christian Urban <urbanc@in.tum.de> parents: 
2972diff
changeset | 291 | termination (eqvt) | 
| 2858 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 292 | by lexicographic_order | 
| 2683 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 293 | |
| 2975 
c62e26830420
preliminary version of automatically generation the eqvt-lemmas for functions defined with nominal_primrec
 Christian Urban <urbanc@in.tum.de> parents: 
2974diff
changeset | 294 | thm subst.eqvt | 
| 2683 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 295 | |
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 296 | lemma forget: | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 297 | shows "atom x \<sharp> t \<Longrightarrow> t[x ::= s] = t" | 
| 2845 
a99f488a96bb
Optimized proofs and removed some garbage.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2843diff
changeset | 298 | by (nominal_induct t avoiding: x s rule: lam.strong_induct) | 
| 
a99f488a96bb
Optimized proofs and removed some garbage.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2843diff
changeset | 299 | (auto simp add: lam.fresh fresh_at_base) | 
| 2683 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 300 | |
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 301 | text {* same lemma but with subst.induction *}
 | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 302 | lemma forget2: | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 303 | shows "atom x \<sharp> t \<Longrightarrow> t[x ::= s] = t" | 
| 2845 
a99f488a96bb
Optimized proofs and removed some garbage.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2843diff
changeset | 304 | by (induct t x s rule: subst.induct) | 
| 
a99f488a96bb
Optimized proofs and removed some garbage.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2843diff
changeset | 305 | (auto simp add: lam.fresh fresh_at_base fresh_Pair) | 
| 2683 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 306 | |
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 307 | lemma fresh_fact: | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 308 | fixes z::"name" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 309 | assumes a: "atom z \<sharp> s" | 
| 2845 
a99f488a96bb
Optimized proofs and removed some garbage.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2843diff
changeset | 310 | and b: "z = y \<or> atom z \<sharp> t" | 
| 2683 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 311 | shows "atom z \<sharp> t[y ::= s]" | 
| 2845 
a99f488a96bb
Optimized proofs and removed some garbage.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2843diff
changeset | 312 | using a b | 
| 
a99f488a96bb
Optimized proofs and removed some garbage.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2843diff
changeset | 313 | by (nominal_induct t avoiding: z y s rule: lam.strong_induct) | 
| 
a99f488a96bb
Optimized proofs and removed some garbage.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2843diff
changeset | 314 | (auto simp add: lam.fresh fresh_at_base) | 
| 2683 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 315 | |
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 316 | lemma substitution_lemma: | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 317 | assumes a: "x \<noteq> y" "atom x \<sharp> u" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 318 | shows "t[x ::= s][y ::= u] = t[y ::= u][x ::= s[y ::= u]]" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 319 | using a | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 320 | by (nominal_induct t avoiding: x y s u rule: lam.strong_induct) | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 321 | (auto simp add: fresh_fact forget) | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 322 | |
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 323 | lemma subst_rename: | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 324 | assumes a: "atom y \<sharp> t" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 325 | shows "t[x ::= s] = ((y \<leftrightarrow> x) \<bullet>t)[y ::= s]" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 326 | using a | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 327 | apply (nominal_induct t avoiding: x y s rule: lam.strong_induct) | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 328 | apply (auto simp add: lam.fresh fresh_at_base) | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 329 | done | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 330 | |
| 2767 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 331 | lemma height_ge_one: | 
| 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 332 | shows "1 \<le> (height e)" | 
| 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 333 | by (induct e rule: lam.induct) (simp_all) | 
| 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 334 | |
| 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 335 | theorem height_subst: | 
| 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 336 | shows "height (e[x::=e']) \<le> ((height e) - 1) + (height e')" | 
| 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 337 | proof (nominal_induct e avoiding: x e' rule: lam.strong_induct) | 
| 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 338 | case (Var y) | 
| 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 339 | have "1 \<le> height e'" by (rule height_ge_one) | 
| 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 340 | then show "height (Var y[x::=e']) \<le> height (Var y) - 1 + height e'" by simp | 
| 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 341 | next | 
| 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 342 | case (Lam y e1) | 
| 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 343 | hence ih: "height (e1[x::=e']) \<le> ((height e1) - 1) + (height e')" by simp | 
| 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 344 | moreover | 
| 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 345 | have vc: "atom y\<sharp>x" "atom y\<sharp>e'" by fact+ (* usual variable convention *) | 
| 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 346 | ultimately show "height ((Lam [y]. e1)[x::=e']) \<le> height (Lam [y]. e1) - 1 + height e'" by simp | 
| 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 347 | next | 
| 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 348 | case (App e1 e2) | 
| 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 349 | hence ih1: "height (e1[x::=e']) \<le> ((height e1) - 1) + (height e')" | 
| 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 350 | and ih2: "height (e2[x::=e']) \<le> ((height e2) - 1) + (height e')" by simp_all | 
| 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 351 | then show "height ((App e1 e2)[x::=e']) \<le> height (App e1 e2) - 1 + height e'" by simp | 
| 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 352 | qed | 
| 2683 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 353 | |
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 354 | subsection {* single-step beta-reduction *}
 | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 355 | |
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 356 | inductive | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 357 |   beta :: "lam \<Rightarrow> lam \<Rightarrow> bool" (" _ \<longrightarrow>b _" [80,80] 80)
 | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 358 | where | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 359 | b1[intro]: "t1 \<longrightarrow>b t2 \<Longrightarrow> App t1 s \<longrightarrow>b App t2 s" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 360 | | b2[intro]: "s1 \<longrightarrow>b s2 \<Longrightarrow> App t s1 \<longrightarrow>b App t s2" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 361 | | b3[intro]: "t1 \<longrightarrow>b t2 \<Longrightarrow> Lam [x]. t1 \<longrightarrow>b Lam [x]. t2" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 362 | | b4[intro]: "atom x \<sharp> s \<Longrightarrow> App (Lam [x]. t) s \<longrightarrow>b t[x ::= s]" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 363 | |
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 364 | equivariance beta | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 365 | |
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 366 | nominal_inductive beta | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 367 | avoids b4: "x" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 368 | by (simp_all add: fresh_star_def fresh_Pair lam.fresh fresh_fact) | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 369 | |
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 370 | text {* One-Reduction *}
 | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 371 | |
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 372 | inductive | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 373 |   One :: "lam \<Rightarrow> lam \<Rightarrow> bool" (" _ \<longrightarrow>1 _" [80,80] 80)
 | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 374 | where | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 375 | o1[intro]: "Var x \<longrightarrow>1 Var x" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 376 | | o2[intro]: "\<lbrakk>t1 \<longrightarrow>1 t2; s1 \<longrightarrow>1 s2\<rbrakk> \<Longrightarrow> App t1 s1 \<longrightarrow>1 App t2 s2" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 377 | | o3[intro]: "t1 \<longrightarrow>1 t2 \<Longrightarrow> Lam [x].t1 \<longrightarrow>1 Lam [x].t2" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 378 | | o4[intro]: "\<lbrakk>atom x \<sharp> (s1, s2); t1 \<longrightarrow>1 t2; s1 \<longrightarrow>1 s2\<rbrakk> \<Longrightarrow> App (Lam [x].t1) s1 \<longrightarrow>1 t2[x ::= s2]" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 379 | |
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 380 | equivariance One | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 381 | |
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 382 | nominal_inductive One | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 383 | avoids o3: "x" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 384 | | o4: "x" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 385 | by (simp_all add: fresh_star_def fresh_Pair lam.fresh fresh_fact) | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 386 | |
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 387 | lemma One_refl: | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 388 | shows "t \<longrightarrow>1 t" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 389 | by (nominal_induct t rule: lam.strong_induct) (auto) | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 390 | |
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 391 | lemma One_subst: | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 392 | assumes a: "t1 \<longrightarrow>1 t2" "s1 \<longrightarrow>1 s2" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 393 | shows "t1[x ::= s1] \<longrightarrow>1 t2[x ::= s2]" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 394 | using a | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 395 | apply(nominal_induct t1 t2 avoiding: s1 s2 x rule: One.strong_induct) | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 396 | apply(auto simp add: substitution_lemma fresh_at_base fresh_fact fresh_Pair) | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 397 | done | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 398 | |
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 399 | lemma better_o4_intro: | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 400 | assumes a: "t1 \<longrightarrow>1 t2" "s1 \<longrightarrow>1 s2" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 401 | shows "App (Lam [x]. t1) s1 \<longrightarrow>1 t2[ x ::= s2]" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 402 | proof - | 
| 2685 
1df873b63cb2
added obtain_fresh lemma; tuned Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2683diff
changeset | 403 | obtain y::"name" where fs: "atom y \<sharp> (x, t1, s1, t2, s2)" by (rule obtain_fresh) | 
| 2683 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 404 | have "App (Lam [x]. t1) s1 = App (Lam [y]. ((y \<leftrightarrow> x) \<bullet> t1)) s1" using fs | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 405 | by (auto simp add: lam.eq_iff Abs1_eq_iff' flip_def fresh_Pair fresh_at_base) | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 406 | also have "\<dots> \<longrightarrow>1 ((y \<leftrightarrow> x) \<bullet> t2)[y ::= s2]" using fs a by (auto simp add: One.eqvt) | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 407 | also have "\<dots> = t2[x ::= s2]" using fs by (simp add: subst_rename[symmetric]) | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 408 | finally show "App (Lam [x].t1) s1 \<longrightarrow>1 t2[x ::= s2]" by simp | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 409 | qed | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 410 | |
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2678diff
changeset | 411 | section {* Locally Nameless Terms *}
 | 
| 2678 
494b859bfc16
defined height as a function that returns an integer
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 412 | |
| 2669 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2667diff
changeset | 413 | nominal_datatype ln = | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2667diff
changeset | 414 | LNBnd nat | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2667diff
changeset | 415 | | LNVar name | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2667diff
changeset | 416 | | LNApp ln ln | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2667diff
changeset | 417 | | LNLam ln | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2667diff
changeset | 418 | |
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2667diff
changeset | 419 | fun | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2667diff
changeset | 420 | lookup :: "name list \<Rightarrow> nat \<Rightarrow> name \<Rightarrow> ln" | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2667diff
changeset | 421 | where | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2667diff
changeset | 422 | "lookup [] n x = LNVar x" | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2667diff
changeset | 423 | | "lookup (y # ys) n x = (if x = y then LNBnd n else (lookup ys (n + 1) x))" | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2667diff
changeset | 424 | |
| 2822 
23befefc6e73
cleaned ups a bit the examples with the invariant framework; exported nominal_function_config datatype into separate structure and file
 Christian Urban <urbanc@in.tum.de> parents: 
2821diff
changeset | 425 | lemma supp_lookup: | 
| 2824 
44d937e8ae78
Proof of trans with invariant
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2822diff
changeset | 426 |   shows "supp (lookup xs n x) \<subseteq> {atom x}"
 | 
| 2822 
23befefc6e73
cleaned ups a bit the examples with the invariant framework; exported nominal_function_config datatype into separate structure and file
 Christian Urban <urbanc@in.tum.de> parents: 
2821diff
changeset | 427 | apply(induct arbitrary: n rule: lookup.induct) | 
| 2824 
44d937e8ae78
Proof of trans with invariant
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2822diff
changeset | 428 | apply(simp add: ln.supp supp_at_base) | 
| 2822 
23befefc6e73
cleaned ups a bit the examples with the invariant framework; exported nominal_function_config datatype into separate structure and file
 Christian Urban <urbanc@in.tum.de> parents: 
2821diff
changeset | 429 | apply(simp add: ln.supp pure_supp) | 
| 
23befefc6e73
cleaned ups a bit the examples with the invariant framework; exported nominal_function_config datatype into separate structure and file
 Christian Urban <urbanc@in.tum.de> parents: 
2821diff
changeset | 430 | done | 
| 2824 
44d937e8ae78
Proof of trans with invariant
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2822diff
changeset | 431 | |
| 
44d937e8ae78
Proof of trans with invariant
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2822diff
changeset | 432 | lemma supp_lookup_in: | 
| 
44d937e8ae78
Proof of trans with invariant
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2822diff
changeset | 433 |   shows "x \<in> set xs \<Longrightarrow> supp (lookup xs n x) = {}"
 | 
| 
44d937e8ae78
Proof of trans with invariant
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2822diff
changeset | 434 | by (induct arbitrary: n rule: lookup.induct)(auto simp add: ln.supp pure_supp) | 
| 
44d937e8ae78
Proof of trans with invariant
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2822diff
changeset | 435 | |
| 
44d937e8ae78
Proof of trans with invariant
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2822diff
changeset | 436 | lemma supp_lookup_notin: | 
| 
44d937e8ae78
Proof of trans with invariant
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2822diff
changeset | 437 |   shows "x \<notin> set xs \<Longrightarrow> supp (lookup xs n x) = {atom x}"
 | 
| 
44d937e8ae78
Proof of trans with invariant
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2822diff
changeset | 438 | by (induct arbitrary: n rule: lookup.induct) (auto simp add: ln.supp pure_supp supp_at_base) | 
| 2822 
23befefc6e73
cleaned ups a bit the examples with the invariant framework; exported nominal_function_config datatype into separate structure and file
 Christian Urban <urbanc@in.tum.de> parents: 
2821diff
changeset | 439 | |
| 2829 
0acb0b8f4106
Simplify ln-trans proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2828diff
changeset | 440 | lemma supp_lookup_fresh: | 
| 
0acb0b8f4106
Simplify ln-trans proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2828diff
changeset | 441 | shows "atom ` set xs \<sharp>* lookup xs n x" | 
| 
0acb0b8f4106
Simplify ln-trans proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2828diff
changeset | 442 | by (case_tac "x \<in> set xs") (auto simp add: fresh_star_def fresh_def supp_lookup_in supp_lookup_notin) | 
| 
0acb0b8f4106
Simplify ln-trans proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2828diff
changeset | 443 | |
| 
0acb0b8f4106
Simplify ln-trans proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2828diff
changeset | 444 | lemma lookup_eqvt[eqvt]: | 
| 2669 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2667diff
changeset | 445 | shows "(p \<bullet> lookup xs n x) = lookup (p \<bullet> xs) (p \<bullet> n) (p \<bullet> x)" | 
| 2767 
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2765diff
changeset | 446 | by (induct xs arbitrary: n) (simp_all add: permute_pure) | 
| 2669 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2667diff
changeset | 447 | |
| 2858 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 448 | text {* Function that translates lambda-terms into locally nameless terms *}
 | 
| 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 449 | |
| 2826 | 450 | nominal_primrec (invariant "\<lambda>(_, xs) y. atom ` set xs \<sharp>* y") | 
| 2669 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2667diff
changeset | 451 | trans :: "lam \<Rightarrow> name list \<Rightarrow> ln" | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2667diff
changeset | 452 | where | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2667diff
changeset | 453 | "trans (Var x) xs = lookup xs 0 x" | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2667diff
changeset | 454 | | "trans (App t1 t2) xs = LNApp (trans t1 xs) (trans t2 xs)" | 
| 2685 
1df873b63cb2
added obtain_fresh lemma; tuned Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2683diff
changeset | 455 | | "atom x \<sharp> xs \<Longrightarrow> trans (Lam [x]. t) xs = LNLam (trans t (x # xs))" | 
| 2829 
0acb0b8f4106
Simplify ln-trans proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2828diff
changeset | 456 | apply (simp add: eqvt_def trans_graph_def) | 
| 
0acb0b8f4106
Simplify ln-trans proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2828diff
changeset | 457 | apply (rule, perm_simp, rule) | 
| 2824 
44d937e8ae78
Proof of trans with invariant
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2822diff
changeset | 458 | apply (erule trans_graph.induct) | 
| 2902 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 459 | apply (auto simp add: ln.fresh)[3] | 
| 2829 
0acb0b8f4106
Simplify ln-trans proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2828diff
changeset | 460 | apply (simp add: supp_lookup_fresh) | 
| 2824 
44d937e8ae78
Proof of trans with invariant
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2822diff
changeset | 461 | apply (simp add: fresh_star_def ln.fresh) | 
| 
44d937e8ae78
Proof of trans with invariant
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2822diff
changeset | 462 | apply (simp add: ln.fresh fresh_star_def) | 
| 2902 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 463 | apply(auto)[1] | 
| 2829 
0acb0b8f4106
Simplify ln-trans proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2828diff
changeset | 464 | apply (rule_tac y="a" and c="b" in lam.strong_exhaust) | 
| 
0acb0b8f4106
Simplify ln-trans proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2828diff
changeset | 465 | apply (auto simp add: fresh_star_def)[3] | 
| 2902 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 466 | apply(simp_all) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 467 | apply(erule conjE)+ | 
| 2948 
b0b2adafb6d2
made the tests go through again
 Christian Urban <urbanc@in.tum.de> parents: 
2945diff
changeset | 468 | apply (erule_tac c="xsa" in Abs_lst1_fcb2') | 
| 2902 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 469 | apply (simp add: fresh_star_def) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 470 | apply (simp add: fresh_star_def) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 471 | apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 472 | apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq) | 
| 2829 
0acb0b8f4106
Simplify ln-trans proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2828diff
changeset | 473 | done | 
| 2669 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2667diff
changeset | 474 | |
| 2973 
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
 Christian Urban <urbanc@in.tum.de> parents: 
2972diff
changeset | 475 | termination (eqvt) | 
| 2858 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 476 | by lexicographic_order | 
| 2827 
394664816e24
defined the "count-bound-variables-occurences" function which has an accumulator like trans
 Christian Urban <urbanc@in.tum.de> parents: 
2826diff
changeset | 477 | |
| 2950 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 478 | |
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 479 | text {* count the occurences of lambdas in a term *}
 | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 480 | |
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 481 | nominal_primrec | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 482 | cntlams :: "lam \<Rightarrow> nat" | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 483 | where | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 484 | "cntlams (Var x) = 0" | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 485 | | "cntlams (App t1 t2) = (cntlams t1) + (cntlams t2)" | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 486 | | "cntlams (Lam [x]. t) = Suc (cntlams t)" | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 487 | apply(simp add: eqvt_def cntlams_graph_def) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 488 | apply(rule, perm_simp, rule) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 489 | apply(rule TrueI) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 490 | apply(rule_tac y="x" in lam.exhaust) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 491 | apply(auto)[3] | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 492 | apply(all_trivials) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 493 | apply(simp) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 494 | apply(simp) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 495 | apply(erule_tac c="()" in Abs_lst1_fcb2') | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 496 | apply(simp add: pure_fresh) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 497 | apply(simp add: fresh_star_def pure_fresh) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 498 | apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 499 | apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 500 | done | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 501 | |
| 2973 
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
 Christian Urban <urbanc@in.tum.de> parents: 
2972diff
changeset | 502 | termination (eqvt) | 
| 2950 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 503 | by lexicographic_order | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 504 | |
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 505 | |
| 2860 
25a7f421a3ba
added a test that every function must be of pt-sort
 Christian Urban <urbanc@in.tum.de> parents: 
2858diff
changeset | 506 | text {* count the bound-variable occurences in a lambda-term *}
 | 
| 2827 
394664816e24
defined the "count-bound-variables-occurences" function which has an accumulator like trans
 Christian Urban <urbanc@in.tum.de> parents: 
2826diff
changeset | 507 | |
| 2828 
81276d5c7438
cbvs can be easily defined without an invariant
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2827diff
changeset | 508 | nominal_primrec | 
| 2950 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 509 | cntbvs :: "lam \<Rightarrow> name list \<Rightarrow> nat" | 
| 2827 
394664816e24
defined the "count-bound-variables-occurences" function which has an accumulator like trans
 Christian Urban <urbanc@in.tum.de> parents: 
2826diff
changeset | 510 | where | 
| 2950 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 511 | "cntbvs (Var x) xs = (if x \<in> set xs then 1 else 0)" | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 512 | | "cntbvs (App t1 t2) xs = (cntbvs t1 xs) + (cntbvs t2 xs)" | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 513 | | "atom x \<sharp> xs \<Longrightarrow> cntbvs (Lam [x]. t) xs = cntbvs t (x # xs)" | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 514 | apply(simp add: eqvt_def cntbvs_graph_def) | 
| 2828 
81276d5c7438
cbvs can be easily defined without an invariant
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2827diff
changeset | 515 | apply(rule, perm_simp, rule) | 
| 2950 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 516 | apply(rule TrueI) | 
| 2828 
81276d5c7438
cbvs can be easily defined without an invariant
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2827diff
changeset | 517 | apply(case_tac x) | 
| 
81276d5c7438
cbvs can be easily defined without an invariant
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2827diff
changeset | 518 | apply(rule_tac y="a" and c="b" in lam.strong_exhaust) | 
| 2902 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 519 | apply(auto simp add: fresh_star_def)[3] | 
| 2950 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 520 | apply(all_trivials) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 521 | apply(simp) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 522 | apply(simp) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 523 | apply(simp) | 
| 2902 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 524 | apply(erule conjE) | 
| 2912 
3c363a5070a5
copied all work to Lambda.thy; had to derive a special version of fcb1 for concrete atom
 Christian Urban <urbanc@in.tum.de> parents: 
2902diff
changeset | 525 | apply(erule Abs_lst1_fcb2') | 
| 2902 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 526 | apply(simp add: pure_fresh fresh_star_def) | 
| 2950 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 527 | apply(simp add: fresh_star_def) | 
| 2902 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 528 | apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 529 | apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq) | 
| 2828 
81276d5c7438
cbvs can be easily defined without an invariant
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2827diff
changeset | 530 | done | 
| 2827 
394664816e24
defined the "count-bound-variables-occurences" function which has an accumulator like trans
 Christian Urban <urbanc@in.tum.de> parents: 
2826diff
changeset | 531 | |
| 2973 
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
 Christian Urban <urbanc@in.tum.de> parents: 
2972diff
changeset | 532 | termination (eqvt) | 
| 2858 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 533 | by lexicographic_order | 
| 2827 
394664816e24
defined the "count-bound-variables-occurences" function which has an accumulator like trans
 Christian Urban <urbanc@in.tum.de> parents: 
2826diff
changeset | 534 | |
| 2858 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 535 | section {* De Bruijn Terms *}
 | 
| 2827 
394664816e24
defined the "count-bound-variables-occurences" function which has an accumulator like trans
 Christian Urban <urbanc@in.tum.de> parents: 
2826diff
changeset | 536 | |
| 2667 
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
 Christian Urban <urbanc@in.tum.de> parents: 
2666diff
changeset | 537 | nominal_datatype db = | 
| 
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
 Christian Urban <urbanc@in.tum.de> parents: 
2666diff
changeset | 538 | DBVar nat | 
| 
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
 Christian Urban <urbanc@in.tum.de> parents: 
2666diff
changeset | 539 | | DBApp db db | 
| 
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
 Christian Urban <urbanc@in.tum.de> parents: 
2666diff
changeset | 540 | | DBLam db | 
| 
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
 Christian Urban <urbanc@in.tum.de> parents: 
2666diff
changeset | 541 | |
| 2792 
c4ed08a7454a
Defined translation from nominal to de-Bruijn; with a freshness condition for the lambda case.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2791diff
changeset | 542 | instance db :: pure | 
| 
c4ed08a7454a
Defined translation from nominal to de-Bruijn; with a freshness condition for the lambda case.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2791diff
changeset | 543 | apply default | 
| 
c4ed08a7454a
Defined translation from nominal to de-Bruijn; with a freshness condition for the lambda case.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2791diff
changeset | 544 | apply (induct_tac x rule: db.induct) | 
| 
c4ed08a7454a
Defined translation from nominal to de-Bruijn; with a freshness condition for the lambda case.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2791diff
changeset | 545 | apply (simp_all add: permute_pure) | 
| 
c4ed08a7454a
Defined translation from nominal to de-Bruijn; with a freshness condition for the lambda case.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2791diff
changeset | 546 | done | 
| 2795 
929bd2dd1ab2
DeBruijn translation in a simplifier friendly way
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2794diff
changeset | 547 | |
| 
929bd2dd1ab2
DeBruijn translation in a simplifier friendly way
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2794diff
changeset | 548 | lemma fresh_at_list: "atom x \<sharp> xs \<longleftrightarrow> x \<notin> set xs" | 
| 
929bd2dd1ab2
DeBruijn translation in a simplifier friendly way
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2794diff
changeset | 549 | unfolding fresh_def supp_set[symmetric] | 
| 2845 
a99f488a96bb
Optimized proofs and removed some garbage.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2843diff
changeset | 550 | by (induct xs) (auto simp add: supp_of_finite_insert supp_at_base supp_set_empty) | 
| 2795 
929bd2dd1ab2
DeBruijn translation in a simplifier friendly way
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2794diff
changeset | 551 | |
| 2800 
6e518b436740
DB translation using index; easier to reason about.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2799diff
changeset | 552 | fun | 
| 
6e518b436740
DB translation using index; easier to reason about.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2799diff
changeset | 553 | vindex :: "name list \<Rightarrow> name \<Rightarrow> nat \<Rightarrow> db option" | 
| 2792 
c4ed08a7454a
Defined translation from nominal to de-Bruijn; with a freshness condition for the lambda case.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2791diff
changeset | 554 | where | 
| 2800 
6e518b436740
DB translation using index; easier to reason about.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2799diff
changeset | 555 | "vindex [] v n = None" | 
| 
6e518b436740
DB translation using index; easier to reason about.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2799diff
changeset | 556 | | "vindex (h # t) v n = (if v = h then (Some (DBVar n)) else (vindex t v (Suc n)))" | 
| 
6e518b436740
DB translation using index; easier to reason about.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2799diff
changeset | 557 | |
| 
6e518b436740
DB translation using index; easier to reason about.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2799diff
changeset | 558 | lemma vindex_eqvt[eqvt]: | 
| 
6e518b436740
DB translation using index; easier to reason about.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2799diff
changeset | 559 | "(p \<bullet> vindex l v n) = vindex (p \<bullet> l) (p \<bullet> v) (p \<bullet> n)" | 
| 
6e518b436740
DB translation using index; easier to reason about.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2799diff
changeset | 560 | by (induct l arbitrary: n) (simp_all add: permute_pure) | 
| 
6e518b436740
DB translation using index; easier to reason about.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2799diff
changeset | 561 | |
| 
6e518b436740
DB translation using index; easier to reason about.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2799diff
changeset | 562 | nominal_primrec | 
| 2822 
23befefc6e73
cleaned ups a bit the examples with the invariant framework; exported nominal_function_config datatype into separate structure and file
 Christian Urban <urbanc@in.tum.de> parents: 
2821diff
changeset | 563 | transdb :: "lam \<Rightarrow> name list \<Rightarrow> db option" | 
| 2800 
6e518b436740
DB translation using index; easier to reason about.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2799diff
changeset | 564 | where | 
| 2822 
23befefc6e73
cleaned ups a bit the examples with the invariant framework; exported nominal_function_config datatype into separate structure and file
 Christian Urban <urbanc@in.tum.de> parents: 
2821diff
changeset | 565 | "transdb (Var x) l = vindex l x 0" | 
| 2950 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 566 | | "transdb (App t1 t2) xs = | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 567 | Option.bind (transdb t1 xs) (\<lambda>d1. Option.bind (transdb t2 xs) (\<lambda>d2. Some (DBApp d1 d2)))" | 
| 2942 
fac8895b109a
Express trans_db with Option.map and Option.bind. Possibly mbind is a copy of bind?
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2941diff
changeset | 568 | | "x \<notin> set xs \<Longrightarrow> transdb (Lam [x].t) xs = Option.map DBLam (transdb t (x # xs))" | 
| 2822 
23befefc6e73
cleaned ups a bit the examples with the invariant framework; exported nominal_function_config datatype into separate structure and file
 Christian Urban <urbanc@in.tum.de> parents: 
2821diff
changeset | 569 | unfolding eqvt_def transdb_graph_def | 
| 2792 
c4ed08a7454a
Defined translation from nominal to de-Bruijn; with a freshness condition for the lambda case.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2791diff
changeset | 570 | apply (rule, perm_simp, rule) | 
| 2822 
23befefc6e73
cleaned ups a bit the examples with the invariant framework; exported nominal_function_config datatype into separate structure and file
 Christian Urban <urbanc@in.tum.de> parents: 
2821diff
changeset | 571 | apply(rule TrueI) | 
| 2792 
c4ed08a7454a
Defined translation from nominal to de-Bruijn; with a freshness condition for the lambda case.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2791diff
changeset | 572 | apply (case_tac x) | 
| 
c4ed08a7454a
Defined translation from nominal to de-Bruijn; with a freshness condition for the lambda case.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2791diff
changeset | 573 | apply (rule_tac y="a" and c="b" in lam.strong_exhaust) | 
| 2902 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 574 | apply (auto simp add: fresh_star_def fresh_at_list)[3] | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 575 | apply(simp_all) | 
| 2942 
fac8895b109a
Express trans_db with Option.map and Option.bind. Possibly mbind is a copy of bind?
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2941diff
changeset | 576 | apply(elim conjE) | 
| 2912 
3c363a5070a5
copied all work to Lambda.thy; had to derive a special version of fcb1 for concrete atom
 Christian Urban <urbanc@in.tum.de> parents: 
2902diff
changeset | 577 | apply (erule_tac c="xsa" in Abs_lst1_fcb2') | 
| 2902 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 578 | apply (simp add: pure_fresh) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 579 | apply(simp add: fresh_star_def fresh_at_list) | 
| 2942 
fac8895b109a
Express trans_db with Option.map and Option.bind. Possibly mbind is a copy of bind?
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2941diff
changeset | 580 | apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq eqvts eqvts_raw)+ | 
| 2792 
c4ed08a7454a
Defined translation from nominal to de-Bruijn; with a freshness condition for the lambda case.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2791diff
changeset | 581 | done | 
| 
c4ed08a7454a
Defined translation from nominal to de-Bruijn; with a freshness condition for the lambda case.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2791diff
changeset | 582 | |
| 2973 
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
 Christian Urban <urbanc@in.tum.de> parents: 
2972diff
changeset | 583 | termination (eqvt) | 
| 2858 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 584 | by lexicographic_order | 
| 2792 
c4ed08a7454a
Defined translation from nominal to de-Bruijn; with a freshness condition for the lambda case.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2791diff
changeset | 585 | |
| 2822 
23befefc6e73
cleaned ups a bit the examples with the invariant framework; exported nominal_function_config datatype into separate structure and file
 Christian Urban <urbanc@in.tum.de> parents: 
2821diff
changeset | 586 | lemma transdb_eqvt[eqvt]: | 
| 
23befefc6e73
cleaned ups a bit the examples with the invariant framework; exported nominal_function_config datatype into separate structure and file
 Christian Urban <urbanc@in.tum.de> parents: 
2821diff
changeset | 587 | "p \<bullet> transdb t l = transdb (p \<bullet>t) (p \<bullet>l)" | 
| 2942 
fac8895b109a
Express trans_db with Option.map and Option.bind. Possibly mbind is a copy of bind?
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2941diff
changeset | 588 | apply (nominal_induct t avoiding: l rule: lam.strong_induct) | 
| 2800 
6e518b436740
DB translation using index; easier to reason about.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2799diff
changeset | 589 | apply (simp add: vindex_eqvt) | 
| 
6e518b436740
DB translation using index; easier to reason about.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2799diff
changeset | 590 | apply (simp_all add: permute_pure) | 
| 2797 
6750964a69bf
equivariance of db_trans
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2796diff
changeset | 591 | apply (simp add: fresh_at_list) | 
| 2822 
23befefc6e73
cleaned ups a bit the examples with the invariant framework; exported nominal_function_config datatype into separate structure and file
 Christian Urban <urbanc@in.tum.de> parents: 
2821diff
changeset | 592 | apply (subst transdb.simps) | 
| 2797 
6750964a69bf
equivariance of db_trans
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2796diff
changeset | 593 | apply (simp add: fresh_at_list[symmetric]) | 
| 2800 
6e518b436740
DB translation using index; easier to reason about.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2799diff
changeset | 594 | apply (drule_tac x="name # l" in meta_spec) | 
| 
6e518b436740
DB translation using index; easier to reason about.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2799diff
changeset | 595 | apply auto | 
| 2797 
6750964a69bf
equivariance of db_trans
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2796diff
changeset | 596 | done | 
| 
6750964a69bf
equivariance of db_trans
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2796diff
changeset | 597 | |
| 2792 
c4ed08a7454a
Defined translation from nominal to de-Bruijn; with a freshness condition for the lambda case.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2791diff
changeset | 598 | lemma db_trans_test: | 
| 
c4ed08a7454a
Defined translation from nominal to de-Bruijn; with a freshness condition for the lambda case.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2791diff
changeset | 599 | assumes a: "y \<noteq> x" | 
| 2858 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 600 | shows "transdb (Lam [x]. Lam [y]. App (Var x) (Var y)) [] = | 
| 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 601 | Some (DBLam (DBLam (DBApp (DBVar 1) (DBVar 0))))" | 
| 2795 
929bd2dd1ab2
DeBruijn translation in a simplifier friendly way
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2794diff
changeset | 602 | using a by simp | 
| 2792 
c4ed08a7454a
Defined translation from nominal to de-Bruijn; with a freshness condition for the lambda case.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2791diff
changeset | 603 | |
| 2841 
f8d660de0cf7
Eval can be defined with additional freshness
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2840diff
changeset | 604 | lemma supp_subst: | 
| 2972 
84afb941df53
moved eqvt for Option.map
 Christian Urban <urbanc@in.tum.de> parents: 
2951diff
changeset | 605 |   shows "supp (t[x ::= s]) \<subseteq> (supp t - {atom x}) \<union> supp s"
 | 
| 
84afb941df53
moved eqvt for Option.map
 Christian Urban <urbanc@in.tum.de> parents: 
2951diff
changeset | 606 | by (induct t x s rule: subst.induct) (auto simp add: lam.supp supp_at_base) | 
| 2841 
f8d660de0cf7
Eval can be defined with additional freshness
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2840diff
changeset | 607 | |
| 
f8d660de0cf7
Eval can be defined with additional freshness
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2840diff
changeset | 608 | lemma var_fresh_subst: | 
| 
f8d660de0cf7
Eval can be defined with additional freshness
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2840diff
changeset | 609 | "atom x \<sharp> s \<Longrightarrow> atom x \<sharp> (t[x ::= s])" | 
| 
f8d660de0cf7
Eval can be defined with additional freshness
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2840diff
changeset | 610 | by (induct t x s rule: subst.induct) (auto simp add: lam.supp lam.fresh fresh_at_base) | 
| 
f8d660de0cf7
Eval can be defined with additional freshness
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2840diff
changeset | 611 | |
| 2834 | 612 | (* function that evaluates a lambda term *) | 
| 613 | nominal_primrec | |
| 614 | eval :: "lam \<Rightarrow> lam" and | |
| 2860 
25a7f421a3ba
added a test that every function must be of pt-sort
 Christian Urban <urbanc@in.tum.de> parents: 
2858diff
changeset | 615 | apply_subst :: "lam \<Rightarrow> lam \<Rightarrow> lam" | 
| 2834 | 616 | where | 
| 617 | "eval (Var x) = Var x" | |
| 618 | | "eval (Lam [x].t) = Lam [x].(eval t)" | |
| 2860 
25a7f421a3ba
added a test that every function must be of pt-sort
 Christian Urban <urbanc@in.tum.de> parents: 
2858diff
changeset | 619 | | "eval (App t1 t2) = apply_subst (eval t1) (eval t2)" | 
| 
25a7f421a3ba
added a test that every function must be of pt-sort
 Christian Urban <urbanc@in.tum.de> parents: 
2858diff
changeset | 620 | | "apply_subst (Var x) t2 = App (Var x) t2" | 
| 
25a7f421a3ba
added a test that every function must be of pt-sort
 Christian Urban <urbanc@in.tum.de> parents: 
2858diff
changeset | 621 | | "apply_subst (App t0 t1) t2 = App (App t0 t1) t2" | 
| 
25a7f421a3ba
added a test that every function must be of pt-sort
 Christian Urban <urbanc@in.tum.de> parents: 
2858diff
changeset | 622 | | "atom x \<sharp> t2 \<Longrightarrow> apply_subst (Lam [x].t1) t2 = eval (t1[x::= t2])" | 
| 2902 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 623 | apply(simp add: eval_apply_subst_graph_def eqvt_def) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 624 | apply(rule, perm_simp, rule) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 625 | apply(rule TrueI) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 626 | apply (case_tac x) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 627 | apply (case_tac a rule: lam.exhaust) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 628 | apply simp_all[3] | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 629 | apply blast | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 630 | apply (case_tac b) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 631 | apply (rule_tac y="a" and c="ba" in lam.strong_exhaust) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 632 | apply simp_all[3] | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 633 | apply blast | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 634 | apply blast | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 635 | apply (simp add: Abs1_eq_iff fresh_star_def) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 636 | apply(simp_all) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 637 | apply(erule_tac c="()" in Abs_lst1_fcb2) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 638 | apply (simp add: Abs_fresh_iff) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 639 | apply(simp add: fresh_star_def fresh_Unit) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 640 | apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 641 | apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 642 | apply(erule conjE) | 
| 2912 
3c363a5070a5
copied all work to Lambda.thy; had to derive a special version of fcb1 for concrete atom
 Christian Urban <urbanc@in.tum.de> parents: 
2902diff
changeset | 643 | apply(erule_tac c="t2a" in Abs_lst1_fcb2') | 
| 2902 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 644 | apply (erule fresh_eqvt_at) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 645 | apply (simp add: finite_supp) | 
| 
9c3f6a4d95d4
another change to the fcb2; this is needed in order to get all proofs through in Lambda.thy
 Christian Urban <urbanc@in.tum.de> parents: 
2891diff
changeset | 646 | apply (simp add: fresh_Inl var_fresh_subst) | 
| 2912 
3c363a5070a5
copied all work to Lambda.thy; had to derive a special version of fcb1 for concrete atom
 Christian Urban <urbanc@in.tum.de> parents: 
2902diff
changeset | 647 | apply(simp add: fresh_star_def) | 
| 2982 
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
 Christian Urban <urbanc@in.tum.de> parents: 
2975diff
changeset | 648 | apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq subst.eqvt) | 
| 
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
 Christian Urban <urbanc@in.tum.de> parents: 
2975diff
changeset | 649 | apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq subst.eqvt) | 
| 2860 
25a7f421a3ba
added a test that every function must be of pt-sort
 Christian Urban <urbanc@in.tum.de> parents: 
2858diff
changeset | 650 | done | 
| 
25a7f421a3ba
added a test that every function must be of pt-sort
 Christian Urban <urbanc@in.tum.de> parents: 
2858diff
changeset | 651 | |
| 
25a7f421a3ba
added a test that every function must be of pt-sort
 Christian Urban <urbanc@in.tum.de> parents: 
2858diff
changeset | 652 | |
| 
25a7f421a3ba
added a test that every function must be of pt-sort
 Christian Urban <urbanc@in.tum.de> parents: 
2858diff
changeset | 653 | (* a small test | 
| 2973 
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
 Christian Urban <urbanc@in.tum.de> parents: 
2972diff
changeset | 654 | termination (eqvt) sorry | 
| 2845 
a99f488a96bb
Optimized proofs and removed some garbage.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2843diff
changeset | 655 | |
| 2860 
25a7f421a3ba
added a test that every function must be of pt-sort
 Christian Urban <urbanc@in.tum.de> parents: 
2858diff
changeset | 656 | lemma | 
| 
25a7f421a3ba
added a test that every function must be of pt-sort
 Christian Urban <urbanc@in.tum.de> parents: 
2858diff
changeset | 657 | assumes "x \<noteq> y" | 
| 
25a7f421a3ba
added a test that every function must be of pt-sort
 Christian Urban <urbanc@in.tum.de> parents: 
2858diff
changeset | 658 | shows "eval (App (Lam [x].App (Var x) (Var x)) (Var y)) = App (Var y) (Var y)" | 
| 
25a7f421a3ba
added a test that every function must be of pt-sort
 Christian Urban <urbanc@in.tum.de> parents: 
2858diff
changeset | 659 | using assms | 
| 
25a7f421a3ba
added a test that every function must be of pt-sort
 Christian Urban <urbanc@in.tum.de> parents: 
2858diff
changeset | 660 | apply(simp add: lam.supp fresh_def supp_at_base) | 
| 
25a7f421a3ba
added a test that every function must be of pt-sort
 Christian Urban <urbanc@in.tum.de> parents: 
2858diff
changeset | 661 | done | 
| 
25a7f421a3ba
added a test that every function must be of pt-sort
 Christian Urban <urbanc@in.tum.de> parents: 
2858diff
changeset | 662 | *) | 
| 
25a7f421a3ba
added a test that every function must be of pt-sort
 Christian Urban <urbanc@in.tum.de> parents: 
2858diff
changeset | 663 | |
| 2852 | 664 | |
| 665 | text {* TODO: eqvt_at for the other side *}
 | |
| 666 | nominal_primrec q where | |
| 667 | "atom c \<sharp> (x, M) \<Longrightarrow> q (Lam [x]. M) (N :: lam) = Lam [x]. (Lam [c]. (App M (q (Var c) N)))" | |
| 668 | | "q (Var x) N = Var x" | |
| 669 | | "q (App l r) N = App l r" | |
| 670 | unfolding eqvt_def q_graph_def | |
| 671 | apply (rule, perm_simp, rule) | |
| 672 | apply (rule TrueI) | |
| 673 | apply (case_tac x) | |
| 674 | apply (rule_tac y="a" in lam.exhaust) | |
| 675 | apply simp_all | |
| 676 | apply blast | |
| 677 | apply blast | |
| 678 | apply (rule_tac x="(name, lam)" and ?'a="name" in obtain_fresh) | |
| 679 | apply blast | |
| 680 | apply clarify | |
| 681 | apply (rule_tac x="(x, xa, M, Ma, c, ca, Na)" and ?'a="name" in obtain_fresh) | |
| 682 | apply (subgoal_tac "eqvt_at q_sumC (Var ca, Na)") --"Could come from nominal_function?" | |
| 683 | apply (subgoal_tac "Lam [c]. App M (q_sumC (Var c, Na)) = Lam [a]. App M (q_sumC (Var a, Na))") | |
| 684 | apply (subgoal_tac "Lam [ca]. App Ma (q_sumC (Var ca, Na)) = Lam [a]. App Ma (q_sumC (Var a, Na))") | |
| 685 | apply (simp only:) | |
| 686 | apply (erule Abs_lst1_fcb) | |
| 687 | oops | |
| 688 | ||
| 2846 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 689 | text {* Working Examples *}
 | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 690 | |
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 691 | nominal_primrec | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 692 | map_term :: "(lam \<Rightarrow> lam) \<Rightarrow> lam \<Rightarrow> lam" | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 693 | where | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 694 | "eqvt f \<Longrightarrow> map_term f (Var x) = f (Var x)" | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 695 | | "eqvt f \<Longrightarrow> map_term f (App t1 t2) = App (f t1) (f t2)" | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 696 | | "eqvt f \<Longrightarrow> map_term f (Lam [x].t) = Lam [x].(f t)" | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 697 | | "\<not>eqvt f \<Longrightarrow> map_term f t = t" | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 698 | apply (simp add: eqvt_def map_term_graph_def) | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 699 | apply (rule, perm_simp, rule) | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 700 | apply(rule TrueI) | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 701 | apply (case_tac x, case_tac "eqvt a", case_tac b rule: lam.exhaust) | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 702 | apply auto | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 703 | apply (erule Abs_lst1_fcb) | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 704 | apply (simp_all add: Abs_fresh_iff fresh_fun_eqvt_app) | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 705 | apply (simp add: eqvt_def permute_fun_app_eq) | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 706 | done | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 707 | |
| 2973 
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
 Christian Urban <urbanc@in.tum.de> parents: 
2972diff
changeset | 708 | termination (eqvt) | 
| 2858 
de6b601c8d3d
added size-lemmas to simplifier; as a result termination can be proved by the standard lexicographic_order method
 Christian Urban <urbanc@in.tum.de> parents: 
2852diff
changeset | 709 | by lexicographic_order | 
| 2846 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 710 | |
| 2950 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 711 | |
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 712 | (* | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 713 | abbreviation | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 714 |   mbind :: "'a option => ('a => 'b option) => 'b option"  ("_ \<guillemotright>= _" [65,65] 65) 
 | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 715 | where | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 716 | "c \<guillemotright>= f \<equiv> case c of None => None | (Some v) => f v" | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 717 | |
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 718 | lemma mbind_eqvt: | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 719 | fixes c::"'a::pt option" | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 720 | shows "(p \<bullet> (c \<guillemotright>= f)) = ((p \<bullet> c) \<guillemotright>= (p \<bullet> f))" | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 721 | apply(cases c) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 722 | apply(simp_all) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 723 | apply(perm_simp) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 724 | apply(rule refl) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 725 | done | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 726 | |
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 727 | lemma mbind_eqvt_raw[eqvt_raw]: | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 728 | shows "(p \<bullet> option_case) \<equiv> option_case" | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 729 | apply(rule eq_reflection) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 730 | apply(rule ext)+ | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 731 | apply(case_tac xb) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 732 | apply(simp_all) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 733 | apply(rule_tac p="-p" in permute_boolE) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 734 | apply(perm_simp add: permute_minus_cancel) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 735 | apply(simp) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 736 | apply(rule_tac p="-p" in permute_boolE) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 737 | apply(perm_simp add: permute_minus_cancel) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 738 | apply(simp) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 739 | done | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 740 | |
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 741 | fun | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 742 | index :: "atom list \<Rightarrow> nat \<Rightarrow> atom \<Rightarrow> nat option" | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 743 | where | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 744 | "index [] n x = None" | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 745 | | "index (y # ys) n x = (if x = y then (Some n) else (index ys (n + 1) x))" | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 746 | |
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 747 | lemma [eqvt]: | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 748 | shows "(p \<bullet> index xs n x) = index (p \<bullet> xs) (p \<bullet> n) (p \<bullet> x)" | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 749 | apply(induct xs arbitrary: n) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 750 | apply(simp_all add: permute_pure) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 751 | done | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 752 | *) | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 753 | |
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 754 | (* | 
| 2846 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 755 | nominal_primrec | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 756 | trans2 :: "lam \<Rightarrow> atom list \<Rightarrow> db option" | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 757 | where | 
| 2950 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 758 | "trans2 (Var x) xs = (index xs 0 (atom x) \<guillemotright>= (\<lambda>n::nat. Some (DBVar n)))" | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 759 | | "trans2 (App t1 t2) xs = | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 760 | ((trans2 t1 xs) \<guillemotright>= (\<lambda>db1::db. (trans2 t2 xs) \<guillemotright>= (\<lambda>db2::db. Some (DBApp db1 db2))))" | 
| 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 761 | | "trans2 (Lam [x].t) xs = (trans2 t (atom x # xs) \<guillemotright>= (\<lambda>db::db. Some (DBLam db)))" | 
| 2846 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 762 | oops | 
| 2950 
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
 Christian Urban <urbanc@in.tum.de> parents: 
2948diff
changeset | 763 | *) | 
| 2846 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 764 | |
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 765 | nominal_primrec | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 766 | CPS :: "lam \<Rightarrow> (lam \<Rightarrow> lam) \<Rightarrow> lam" | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 767 | where | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 768 | "CPS (Var x) k = Var x" | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 769 | | "CPS (App M N) k = CPS M (\<lambda>m. CPS N (\<lambda>n. n))" | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 770 | oops | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 771 | |
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 772 | consts b :: name | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 773 | nominal_primrec | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 774 | Z :: "lam \<Rightarrow> (lam \<Rightarrow> lam) \<Rightarrow> lam" | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 775 | where | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 776 | "Z (App M N) k = Z M (%m. (Z N (%n.(App m n))))" | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 777 | | "Z (App M N) k = Z M (%m. (Z N (%n.(App (App m n) (Abs b (k (Var b)))))))" | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 778 | unfolding eqvt_def Z_graph_def | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 779 | apply (rule, perm_simp, rule) | 
| 
1d43d30e44c9
Move working examples before non-working ones
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2845diff
changeset | 780 | oops | 
| 2789 
32979078bfe9
functions involving if and case do not throw exceptions anymore; but eqvt_at assumption has now a precondition
 Christian Urban <urbanc@in.tum.de> parents: 
2787diff
changeset | 781 | |
| 2937 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 782 | lemma test: | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 783 | assumes "t = s" | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 784 | and "supp p \<sharp>* t" "supp p \<sharp>* x" | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 785 | and "(p \<bullet> t) = s \<Longrightarrow> (p \<bullet> x) = y" | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 786 | shows "x = y" | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 787 | using assms by (simp add: perm_supp_eq) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 788 | |
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 789 | lemma test2: | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 790 | assumes "cs \<subseteq> as \<union> bs" | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 791 | and "as \<sharp>* x" "bs \<sharp>* x" | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 792 | shows "cs \<sharp>* x" | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 793 | using assms | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 794 | by (auto simp add: fresh_star_def) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 795 | |
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 796 | lemma test3: | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 797 | assumes "cs \<subseteq> as" | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 798 | and "as \<sharp>* x" | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 799 | shows "cs \<sharp>* x" | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 800 | using assms | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 801 | by (auto simp add: fresh_star_def) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 802 | |
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 803 | |
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 804 | |
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 805 | nominal_primrec (invariant "\<lambda>(_, _, xs) y. atom ` fst ` set xs \<sharp>* y \<and> atom ` snd ` set xs \<sharp>* y") | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 806 | aux :: "lam \<Rightarrow> lam \<Rightarrow> (name \<times> name) list \<Rightarrow> bool" | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 807 | where | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 808 | "aux (Var x) (Var y) xs = ((x, y) \<in> set xs)" | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 809 | | "aux (App t1 t2) (App s1 s2) xs = (aux t1 s1 xs \<and> aux t2 s2 xs)" | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 810 | | "aux (Var x) (App t1 t2) xs = False" | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 811 | | "aux (Var x) (Lam [y].t) xs = False" | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 812 | | "aux (App t1 t2) (Var x) xs = False" | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 813 | | "aux (App t1 t2) (Lam [x].t) xs = False" | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 814 | | "aux (Lam [x].t) (Var y) xs = False" | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 815 | | "aux (Lam [x].t) (App t1 t2) xs = False" | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 816 | | "\<lbrakk>{atom x} \<sharp>* (s, xs); {atom y} \<sharp>* (t, xs); x \<noteq> y\<rbrakk> \<Longrightarrow> 
 | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 817 | aux (Lam [x].t) (Lam [y].s) xs = aux t s ((x, y) # xs)" | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 818 | apply (simp add: eqvt_def aux_graph_def) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 819 | apply (rule, perm_simp, rule) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 820 | apply(erule aux_graph.induct) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 821 | apply(simp_all add: fresh_star_def pure_fresh)[9] | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 822 | apply(case_tac x) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 823 | apply(simp) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 824 | apply(rule_tac y="a" and c="(b, c)" in lam.strong_exhaust) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 825 | apply(simp) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 826 | apply(rule_tac y="b" and c="c" in lam.strong_exhaust) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 827 | apply(metis)+ | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 828 | apply(simp) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 829 | apply(rule_tac y="b" and c="c" in lam.strong_exhaust) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 830 | apply(metis)+ | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 831 | apply(simp) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 832 | apply(rule_tac y="b" and c="(lam, c, name)" in lam.strong_exhaust) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 833 | apply(metis)+ | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 834 | apply(simp) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 835 | apply(drule_tac x="name" in meta_spec) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 836 | apply(drule_tac x="lama" in meta_spec) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 837 | apply(drule_tac x="c" in meta_spec) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 838 | apply(drule_tac x="namea" in meta_spec) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 839 | apply(drule_tac x="lam" in meta_spec) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 840 | apply(simp add: fresh_star_Pair) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 841 | apply(simp add: fresh_star_def fresh_at_base lam.fresh) | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 842 | apply(auto)[1] | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 843 | apply(simp_all)[44] | 
| 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 844 | apply(simp del: Product_Type.prod.inject) | 
| 2948 
b0b2adafb6d2
made the tests go through again
 Christian Urban <urbanc@in.tum.de> parents: 
2945diff
changeset | 845 | oops | 
| 2937 
a56d422e17f6
added an example that recurses over two arguments; the interesting proof-obligation is not yet done
 Christian Urban <urbanc@in.tum.de> parents: 
2912diff
changeset | 846 | |
| 2941 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 847 | lemma abs_same_binder: | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 848 | fixes t ta s sa :: "_ :: fs" | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 849 | assumes "sort_of (atom x) = sort_of (atom y)" | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 850 | shows "[[atom x]]lst. t = [[atom y]]lst. ta \<and> [[atom x]]lst. s = [[atom y]]lst. sa | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 851 | \<longleftrightarrow> [[atom x]]lst. (t, s) = [[atom y]]lst. (ta, sa)" | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 852 | by (cases "atom x = atom y") (auto simp add: Abs1_eq_iff assms fresh_Pair) | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 853 | |
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 854 | nominal_primrec | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 855 | aux2 :: "lam \<Rightarrow> lam \<Rightarrow> bool" | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 856 | where | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 857 | "aux2 (Var x) (Var y) = (x = y)" | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 858 | | "aux2 (App t1 t2) (App s1 s2) = (aux2 t1 s1 \<and> aux2 t2 s2)" | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 859 | | "aux2 (Var x) (App t1 t2) = False" | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 860 | | "aux2 (Var x) (Lam [y].t) = False" | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 861 | | "aux2 (App t1 t2) (Var x) = False" | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 862 | | "aux2 (App t1 t2) (Lam [x].t) = False" | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 863 | | "aux2 (Lam [x].t) (Var y) = False" | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 864 | | "aux2 (Lam [x].t) (App t1 t2) = False" | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 865 | | "x = y \<Longrightarrow> aux2 (Lam [x].t) (Lam [y].s) = aux2 t s" | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 866 | apply(simp add: eqvt_def aux2_graph_def) | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 867 | apply(rule, perm_simp, rule, rule) | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 868 | apply(case_tac x) | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 869 | apply(rule_tac y="a" and c="b" in lam.strong_exhaust) | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 870 | apply(rule_tac y="b" in lam.exhaust) | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 871 | apply(auto)[3] | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 872 | apply(rule_tac y="b" in lam.exhaust) | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 873 | apply(auto)[3] | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 874 | apply(rule_tac y="b" and c="(name, lam)" in lam.strong_exhaust) | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 875 | apply(auto)[3] | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 876 | apply(drule_tac x="name" in meta_spec) | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 877 | apply(drule_tac x="name" in meta_spec) | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 878 | apply(drule_tac x="lam" in meta_spec) | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 879 | apply(drule_tac x="(name \<leftrightarrow> namea) \<bullet> lama" in meta_spec) | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 880 | apply(simp add: Abs1_eq_iff fresh_star_def fresh_Pair_elim fresh_at_base lam.fresh flip_def) | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 881 | apply (metis Nominal2_Base.swap_commute fresh_permute_iff sort_of_atom_eq swap_atom_simps(2)) | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 882 | apply simp_all | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 883 | apply (simp add: abs_same_binder) | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 884 | apply (erule_tac c="()" in Abs_lst1_fcb2) | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 885 | apply (simp_all add: pure_fresh fresh_star_def eqvt_at_def) | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 886 | done | 
| 
40991ebcda12
Define a version of aux only for same binders. Completeness is fine.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2940diff
changeset | 887 | |
| 2940 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 888 | text {* tests of functions containing if and case *}
 | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 889 | |
| 2995 
6d2859aeebba
Comment out examples with 'True' that do not work because function still does not work
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2982diff
changeset | 890 | (*consts P :: "lam \<Rightarrow> bool" | 
| 2940 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 891 | |
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 892 | nominal_primrec | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 893 | A :: "lam => lam" | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 894 | where | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 895 | "A (App M N) = (if (True \<or> P M) then (A M) else (A N))" | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 896 | | "A (Var x) = (Var x)" | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 897 | | "A (App M N) = (if True then M else A N)" | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 898 | oops | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 899 | |
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 900 | nominal_primrec | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 901 | C :: "lam => lam" | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 902 | where | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 903 | "C (App M N) = (case (True \<or> P M) of True \<Rightarrow> (A M) | False \<Rightarrow> (A N))" | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 904 | | "C (Var x) = (Var x)" | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 905 | | "C (App M N) = (if True then M else C N)" | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 906 | oops | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 907 | |
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 908 | nominal_primrec | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 909 | A :: "lam => lam" | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 910 | where | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 911 | "A (Lam [x].M) = (Lam [x].M)" | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 912 | | "A (Var x) = (Var x)" | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 913 | | "A (App M N) = (if True then M else A N)" | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 914 | oops | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 915 | |
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 916 | nominal_primrec | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 917 | B :: "lam => lam" | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 918 | where | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 919 | "B (Lam [x].M) = (Lam [x].M)" | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 920 | | "B (Var x) = (Var x)" | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 921 | | "B (App M N) = (if True then M else (B N))" | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 922 | unfolding eqvt_def | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 923 | unfolding B_graph_def | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 924 | apply(perm_simp) | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 925 | apply(rule allI) | 
| 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 926 | apply(rule refl) | 
| 2995 
6d2859aeebba
Comment out examples with 'True' that do not work because function still does not work
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2982diff
changeset | 927 | oops*) | 
| 2940 
cc0605102f95
Move If / Let with 'True' to the end of Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2937diff
changeset | 928 | |
| 1594 | 929 | end | 
| 930 | ||
| 931 | ||
| 932 |