| author | Christian Urban <urbanc@in.tum.de> | 
| Mon, 26 Apr 2010 08:08:20 +0200 | |
| changeset 1949 | 0b692f37a771 | 
| parent 1947 | 51f411b1197d | 
| child 1950 | 7de54c9f81ac | 
| 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 | 
| 1773 
c0eac04ae3b4
added README and moved examples into separate directory
 Christian Urban <urbanc@in.tum.de> parents: 
1599diff
changeset | 2 | imports "../Parser" | 
| 1594 | 3 | begin | 
| 4 | ||
| 5 | atom_decl name | |
| 6 | ||
| 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 | 7 | nominal_datatype lam = | 
| 
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 | 8 | Var "name" | 
| 
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 | 9 | | App "lam" "lam" | 
| 
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 | 10 | | Lam x::"name" l::"lam" bind x in l | 
| 1594 | 11 | |
| 1845 
b7423c6b5564
deleted offending [eqvt]-attribute in Abs; Lambda works again, but there is now a problem in CoreHaskell
 Christian Urban <urbanc@in.tum.de> parents: 
1835diff
changeset | 12 | thm lam.fv | 
| 
b7423c6b5564
deleted offending [eqvt]-attribute in Abs; Lambda works again, but there is now a problem in CoreHaskell
 Christian Urban <urbanc@in.tum.de> parents: 
1835diff
changeset | 13 | thm lam.supp | 
| 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 | 14 | lemmas supp_fn' = lam.fv[simplified lam.supp] | 
| 
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 | 15 | |
| 1805 | 16 | declare lam.perm[eqvt] | 
| 17 | ||
| 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 | 18 | section {* Strong Induction Principles*}
 | 
| 1594 | 19 | |
| 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 | 20 | (* | 
| 
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 | 21 | Old way of establishing strong induction | 
| 
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 | 22 | principles by chosing a fresh name. | 
| 
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 | 23 | *) | 
| 1594 | 24 | lemma | 
| 25 | fixes c::"'a::fs" | |
| 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 | 26 | assumes a1: "\<And>name c. P c (Var name)" | 
| 
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 | 27 | and a2: "\<And>lam1 lam2 c. \<lbrakk>\<And>d. P d lam1; \<And>d. P d lam2\<rbrakk> \<Longrightarrow> P c (App lam1 lam2)" | 
| 
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 | 28 | and a3: "\<And>name lam c. \<lbrakk>atom name \<sharp> c; \<And>d. P d lam\<rbrakk> \<Longrightarrow> P c (Lam name lam)" | 
| 
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 | 29 | shows "P c lam" | 
| 1594 | 30 | proof - | 
| 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 | 31 | have "\<And>p. P c (p \<bullet> lam)" | 
| 
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 | 32 | apply(induct lam arbitrary: c rule: lam.induct) | 
| 1805 | 33 | apply(perm_simp) | 
| 1594 | 34 | apply(rule a1) | 
| 1805 | 35 | apply(perm_simp) | 
| 1594 | 36 | apply(rule a2) | 
| 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 | 37 | apply(assumption) | 
| 1594 | 38 | apply(assumption) | 
| 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 | 39 | apply(subgoal_tac "\<exists>new::name. (atom new) \<sharp> (c, Lam (p \<bullet> name) (p \<bullet> lam))") | 
| 1594 | 40 | defer | 
| 41 | apply(simp add: fresh_def) | |
| 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 | 42 | apply(rule_tac X="supp (c, Lam (p \<bullet> name) (p \<bullet> lam))" in obtain_at_base) | 
| 1594 | 43 | apply(simp add: supp_Pair finite_supp) | 
| 44 | apply(blast) | |
| 45 | apply(erule exE) | |
| 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 | 46 | apply(rule_tac t="p \<bullet> Lam name lam" and | 
| 
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 | 47 | s="(((p \<bullet> name) \<leftrightarrow> new) + p) \<bullet> Lam name lam" in subst) | 
| 
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 | 48 | apply(simp del: lam.perm) | 
| 
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 | 49 | apply(subst lam.perm) | 
| 
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 | 50 | apply(subst (2) lam.perm) | 
| 1594 | 51 | apply(rule flip_fresh_fresh) | 
| 52 | apply(simp add: fresh_def) | |
| 53 | apply(simp only: supp_fn') | |
| 54 | apply(simp) | |
| 55 | apply(simp add: fresh_Pair) | |
| 56 | apply(simp) | |
| 57 | apply(rule a3) | |
| 58 | apply(simp add: fresh_Pair) | |
| 59 | apply(drule_tac x="((p \<bullet> name) \<leftrightarrow> new) + p" in meta_spec) | |
| 60 | apply(simp) | |
| 61 | done | |
| 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 | 62 | then have "P c (0 \<bullet> lam)" by blast | 
| 
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 | 63 | then show "P c lam" by simp | 
| 1594 | 64 | qed | 
| 65 | ||
| 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 | 66 | (* | 
| 
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 | 67 | New way of establishing strong induction | 
| 
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 | 68 | principles by using a appropriate permutation. | 
| 
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 | 69 | *) | 
| 1594 | 70 | lemma | 
| 71 | fixes c::"'a::fs" | |
| 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 | 72 | assumes a1: "\<And>name c. P c (Var name)" | 
| 
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 | 73 | and a2: "\<And>lam1 lam2 c. \<lbrakk>\<And>d. P d lam1; \<And>d. P d lam2\<rbrakk> \<Longrightarrow> P c (App lam1 lam2)" | 
| 
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 | 74 | and a3: "\<And>name lam c. \<lbrakk>atom name \<sharp> c; \<And>d. P d lam\<rbrakk> \<Longrightarrow> P c (Lam name lam)" | 
| 
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 | 75 | shows "P c lam" | 
| 1594 | 76 | proof - | 
| 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 | 77 | have "\<And>p. P c (p \<bullet> lam)" | 
| 
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 | 78 | apply(induct lam arbitrary: c rule: lam.induct) | 
| 1805 | 79 | apply(perm_simp) | 
| 1594 | 80 | apply(rule a1) | 
| 1805 | 81 | apply(perm_simp) | 
| 1594 | 82 | apply(rule a2) | 
| 83 | apply(assumption) | |
| 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 | 84 | apply(assumption) | 
| 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 | 85 |     apply(subgoal_tac "\<exists>q. (q \<bullet> {p \<bullet> atom name}) \<sharp>* c \<and> supp (p \<bullet> Lam name lam) \<sharp>* q")
 | 
| 1594 | 86 | apply(erule exE) | 
| 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 | 87 | apply(rule_tac t="p \<bullet> Lam name lam" and | 
| 
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 | 88 | s="q \<bullet> p \<bullet> Lam name lam" in subst) | 
| 1594 | 89 | defer | 
| 1805 | 90 | apply(simp) | 
| 1594 | 91 | apply(rule a3) | 
| 92 | apply(simp add: eqvts fresh_star_def) | |
| 93 | apply(drule_tac x="q + p" in meta_spec) | |
| 94 | apply(simp) | |
| 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 | 95 | apply(rule at_set_avoiding2) | 
| 
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 | 96 | apply(simp add: finite_supp) | 
| 
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 | 97 | apply(simp add: finite_supp) | 
| 
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 | 98 | apply(simp add: finite_supp) | 
| 1805 | 99 | apply(perm_simp) | 
| 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 | 100 | apply(simp add: fresh_star_def fresh_def supp_fn') | 
| 
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 | 101 | apply(rule supp_perm_eq) | 
| 
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 | 102 | apply(simp) | 
| 
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 | 103 | done | 
| 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 | 104 | then have "P c (0 \<bullet> lam)" by blast | 
| 
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 | 105 | then show "P c lam" by simp | 
| 1594 | 106 | qed | 
| 107 | ||
| 1805 | 108 | section {* Typing *}
 | 
| 109 | ||
| 110 | nominal_datatype ty = | |
| 111 | TVar string | |
| 1810 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1805diff
changeset | 112 | | TFun ty ty | 
| 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1805diff
changeset | 113 | |
| 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1805diff
changeset | 114 | notation | 
| 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1805diff
changeset | 115 |  TFun ("_ \<rightarrow> _") 
 | 
| 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1805diff
changeset | 116 | |
| 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1805diff
changeset | 117 | declare ty.perm[eqvt] | 
| 1805 | 118 | |
| 119 | inductive | |
| 120 | valid :: "(name \<times> ty) list \<Rightarrow> bool" | |
| 121 | where | |
| 122 | "valid []" | |
| 123 | | "\<lbrakk>atom x \<sharp> Gamma; valid Gamma\<rbrakk> \<Longrightarrow> valid ((x, T)#Gamma)" | |
| 124 | ||
| 1828 | 125 | inductive | 
| 126 |   typing :: "(name\<times>ty) list \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _" [60,60,60] 60) 
 | |
| 127 | where | |
| 128 | t_Var[intro]: "\<lbrakk>valid \<Gamma>; (x, T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T" | |
| 1947 | 129 | | t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : T1 \<rightarrow> T2; \<Gamma> \<turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : T2" | 
| 1828 | 130 | | t_Lam[intro]: "\<lbrakk>atom x \<sharp> \<Gamma>; (x, T1) # \<Gamma> \<turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam x t : T1 \<rightarrow> T2" | 
| 131 | ||
| 1831 
16653e702d89
first working version of the automatic equivariance procedure
 Christian Urban <urbanc@in.tum.de> parents: 
1828diff
changeset | 132 | equivariance valid | 
| 
16653e702d89
first working version of the automatic equivariance procedure
 Christian Urban <urbanc@in.tum.de> parents: 
1828diff
changeset | 133 | equivariance typing | 
| 1816 
56cebe7f8e24
some small tunings (incompleted work in Lambda.thy)
 Christian Urban <urbanc@in.tum.de> parents: 
1814diff
changeset | 134 | |
| 1831 
16653e702d89
first working version of the automatic equivariance procedure
 Christian Urban <urbanc@in.tum.de> parents: 
1828diff
changeset | 135 | thm valid.eqvt | 
| 
16653e702d89
first working version of the automatic equivariance procedure
 Christian Urban <urbanc@in.tum.de> parents: 
1828diff
changeset | 136 | thm typing.eqvt | 
| 1811 
ae176476b525
implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1810diff
changeset | 137 | thm eqvts | 
| 
ae176476b525
implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1810diff
changeset | 138 | thm eqvts_raw | 
| 
ae176476b525
implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1810diff
changeset | 139 | |
| 1949 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 140 | thm typing.induct[of "\<Gamma>" "t" "T", no_vars] | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 141 | |
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 142 | lemma | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 143 | fixes c::"'a::fs" | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 144 | assumes a: "\<Gamma> \<turnstile> t : T" | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 145 | and a1: "\<And>\<Gamma> x T c. \<lbrakk>valid \<Gamma>; (x, T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> P c \<Gamma> (Var x) T" | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 146 | and a2: "\<And>\<Gamma> t1 T1 T2 t2 c. \<lbrakk>\<Gamma> \<turnstile> t1 : T1 \<rightarrow> T2; \<And>d. P d \<Gamma> t1 T1 \<rightarrow> T2; \<Gamma> \<turnstile> t2 : T1; \<And>d. P d \<Gamma> t2 T1\<rbrakk> | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 147 | \<Longrightarrow> P c \<Gamma> (App t1 t2) T2" | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 148 | and a3: "\<And>x \<Gamma> T1 t T2 c. \<lbrakk>atom x \<sharp> c; atom x \<sharp> \<Gamma>; (x, T1) # \<Gamma> \<turnstile> t : T2; \<And>d. P d ((x, T1) # \<Gamma>) t T2\<rbrakk> | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 149 | \<Longrightarrow> P c \<Gamma> (Lam x t) T1 \<rightarrow> T2" | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 150 | shows "P c \<Gamma> t T" | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 151 | proof - | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 152 | from a have "\<And>p c. P c (p \<bullet> \<Gamma>) (p \<bullet> t) (p \<bullet> T)" | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 153 | proof (induct) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 154 | case (t_Var \<Gamma> x T p c) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 155 | then show ?case | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 156 | apply - | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 157 | apply(perm_strict_simp) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 158 | apply(rule a1) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 159 | apply(rule_tac p="-p" in permute_boolE) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 160 | apply(perm_strict_simp add: permute_minus_cancel) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 161 | apply(assumption) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 162 | apply(rule_tac p="-p" in permute_boolE) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 163 | apply(perm_strict_simp add: permute_minus_cancel) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 164 | apply(assumption) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 165 | done | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 166 | next | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 167 | case (t_App \<Gamma> t1 T1 T2 t2 p c) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 168 | then show ?case | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 169 | apply - | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 170 | apply(perm_strict_simp) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 171 | apply(rule_tac ?T1.0="p \<bullet> T1" in a2) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 172 | apply(rule_tac p="-p" in permute_boolE) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 173 | apply(perm_strict_simp add: permute_minus_cancel) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 174 | apply(assumption) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 175 | apply(assumption) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 176 | apply(rule_tac p="-p" in permute_boolE) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 177 | apply(perm_strict_simp add: permute_minus_cancel) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 178 | apply(assumption) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 179 | apply(assumption) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 180 | done | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 181 | next | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 182 | case (t_Lam x \<Gamma> T1 t T2 p c) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 183 | then show ?case | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 184 | apply - | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 185 |       apply(subgoal_tac "\<exists>q. (q \<bullet> {p \<bullet> atom x}) \<sharp>* c \<and> 
 | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 186 | supp (p \<bullet> \<Gamma>, p \<bullet> Lam x t, p \<bullet> (T1 \<rightarrow> T2)) \<sharp>* q") | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 187 | apply(erule exE) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 188 | apply(rule_tac t="p \<bullet> \<Gamma>" and s="q \<bullet> p \<bullet> \<Gamma>" in subst) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 189 | apply(rule supp_perm_eq) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 190 | apply(simp add: supp_Pair fresh_star_union) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 191 | apply(rule_tac t="p \<bullet> Lam x t" and s="q \<bullet> p \<bullet> Lam x t" in subst) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 192 | apply(rule supp_perm_eq) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 193 | apply(simp add: supp_Pair fresh_star_union) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 194 | apply(rule_tac t="p \<bullet> (T1 \<rightarrow> T2)" and s="q \<bullet> p \<bullet> (T1 \<rightarrow> T2)" in subst) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 195 | apply(rule supp_perm_eq) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 196 | apply(simp add: supp_Pair fresh_star_union) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 197 | apply(simp add: eqvts) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 198 | apply(rule a3) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 199 | apply(simp add: fresh_star_def) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 200 | apply(rule_tac p="-q" in permute_boolE) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 201 | apply(perm_strict_simp add: permute_minus_cancel) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 202 | apply(rule_tac p="-p" in permute_boolE) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 203 | apply(perm_strict_simp add: permute_minus_cancel) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 204 | apply(assumption) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 205 | apply(rule_tac p="-q" in permute_boolE) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 206 | apply(perm_strict_simp add: permute_minus_cancel) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 207 | apply(rule_tac p="-p" in permute_boolE) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 208 | apply(perm_strict_simp add: permute_minus_cancel) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 209 | apply(assumption) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 210 | apply(drule_tac x="d" in meta_spec) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 211 | apply(drule_tac x="q + p" in meta_spec) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 212 | apply(simp) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 213 | apply(rule at_set_avoiding2) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 214 | apply(simp add: finite_supp) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 215 | apply(simp add: finite_supp) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 216 | apply(simp add: finite_supp) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 217 | apply(rule_tac p="-p" in permute_boolE) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 218 | apply(perm_strict_simp add: permute_minus_cancel) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 219 | (* supplied by the user *) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 220 | apply(simp add: fresh_star_prod) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 221 | apply(simp add: fresh_star_def) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 222 | |
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 223 | |
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 224 | done | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 225 | (* HERE *) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 226 | |
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 227 | |
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 228 | |
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 229 | |
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 230 | |
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 231 | |
| 1833 
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
 Christian Urban <urbanc@in.tum.de> parents: 
1831diff
changeset | 232 | |
| 
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
 Christian Urban <urbanc@in.tum.de> parents: 
1831diff
changeset | 233 | inductive | 
| 
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
 Christian Urban <urbanc@in.tum.de> parents: 
1831diff
changeset | 234 |   tt :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool)"  
 | 
| 
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
 Christian Urban <urbanc@in.tum.de> parents: 
1831diff
changeset | 235 |   for r :: "('a \<Rightarrow> 'a \<Rightarrow> bool)"
 | 
| 
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
 Christian Urban <urbanc@in.tum.de> parents: 
1831diff
changeset | 236 | where | 
| 1835 
636de31888a6
tuned and removed dead code
 Christian Urban <urbanc@in.tum.de> parents: 
1833diff
changeset | 237 | aa: "tt r a a" | 
| 
636de31888a6
tuned and removed dead code
 Christian Urban <urbanc@in.tum.de> parents: 
1833diff
changeset | 238 | | bb: "tt r a b ==> tt r a c" | 
| 1833 
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
 Christian Urban <urbanc@in.tum.de> parents: 
1831diff
changeset | 239 | |
| 1835 
636de31888a6
tuned and removed dead code
 Christian Urban <urbanc@in.tum.de> parents: 
1833diff
changeset | 240 | (* PROBLEM: derived eqvt-theorem does not conform with [eqvt] | 
| 1833 
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
 Christian Urban <urbanc@in.tum.de> parents: 
1831diff
changeset | 241 | equivariance tt | 
| 
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
 Christian Urban <urbanc@in.tum.de> parents: 
1831diff
changeset | 242 | *) | 
| 
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
 Christian Urban <urbanc@in.tum.de> parents: 
1831diff
changeset | 243 | |
| 1810 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1805diff
changeset | 244 | |
| 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1805diff
changeset | 245 | inductive | 
| 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1805diff
changeset | 246 | alpha_lam_raw' | 
| 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1805diff
changeset | 247 | where | 
| 1861 
226b797868dc
some tuning of eqvt-infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
1845diff
changeset | 248 | a1: "name = namea \<Longrightarrow> alpha_lam_raw' (Var_raw name) (Var_raw namea)" | 
| 
226b797868dc
some tuning of eqvt-infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
1845diff
changeset | 249 | | a2: "\<lbrakk>alpha_lam_raw' lam_raw1 lam_raw1a; alpha_lam_raw' lam_raw2 lam_raw2a\<rbrakk> \<Longrightarrow> | 
| 1810 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1805diff
changeset | 250 | alpha_lam_raw' (App_raw lam_raw1 lam_raw2) (App_raw lam_raw1a lam_raw2a)" | 
| 1865 
b71b838b0a75
Finished proof in Lambda.thy
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1863diff
changeset | 251 | | a3: "\<exists>pi. ({atom name}, lam_raw) \<approx>gen alpha_lam_raw' fv_lam_raw pi ({atom namea}, lam_rawa) \<Longrightarrow>
 | 
| 1810 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1805diff
changeset | 252 | alpha_lam_raw' (Lam_raw name lam_raw) (Lam_raw namea lam_rawa)" | 
| 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1805diff
changeset | 253 | |
| 1831 
16653e702d89
first working version of the automatic equivariance procedure
 Christian Urban <urbanc@in.tum.de> parents: 
1828diff
changeset | 254 | declare permute_lam_raw.simps[eqvt] | 
| 1861 
226b797868dc
some tuning of eqvt-infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
1845diff
changeset | 255 | declare alpha_gen_eqvt[eqvt] | 
| 1866 
6d4e4bf9bce6
automatic proofs for equivariance of alphas
 Christian Urban <urbanc@in.tum.de> parents: 
1865diff
changeset | 256 | equivariance alpha_lam_raw' | 
| 1947 | 257 | |
| 1861 
226b797868dc
some tuning of eqvt-infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
1845diff
changeset | 258 | thm eqvts_raw | 
| 
226b797868dc
some tuning of eqvt-infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
1845diff
changeset | 259 | |
| 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 | 260 | section {* size function *}
 | 
| 
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 | 261 | |
| 
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 | 262 | lemma size_eqvt_raw: | 
| 
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 | 263 | fixes t::"lam_raw" | 
| 
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 | 264 | shows "size (pi \<bullet> t) = size t" | 
| 
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 | 265 | apply (induct rule: lam_raw.inducts) | 
| 
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 | 266 | apply simp_all | 
| 
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 | 267 | done | 
| 
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 | 268 | |
| 
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 | 269 | instantiation lam :: size | 
| 
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 | 270 | begin | 
| 
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 | 271 | |
| 
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 | 272 | quotient_definition | 
| 
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 | 273 | "size_lam :: lam \<Rightarrow> nat" | 
| 
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 | 274 | is | 
| 
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 | 275 | "size :: lam_raw \<Rightarrow> nat" | 
| 
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 | 276 | |
| 
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 | 277 | lemma size_rsp: | 
| 
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 | 278 | "alpha_lam_raw x y \<Longrightarrow> size x = size y" | 
| 
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 | 279 | apply (induct rule: alpha_lam_raw.inducts) | 
| 
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 | 280 | apply (simp_all only: lam_raw.size) | 
| 
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 | 281 | apply (simp_all only: alphas) | 
| 
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 | 282 | apply clarify | 
| 
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 | 283 | apply (simp_all only: size_eqvt_raw) | 
| 
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 | 284 | done | 
| 
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 | 285 | |
| 
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 | 286 | lemma [quot_respect]: | 
| 
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 | 287 | "(alpha_lam_raw ===> op =) size size" | 
| 
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 | 288 | by (simp_all add: size_rsp) | 
| 
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 | 289 | |
| 
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 | 290 | lemma [quot_preserve]: | 
| 
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 | 291 | "(rep_lam ---> id) size = size" | 
| 
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 | 292 | by (simp_all add: size_lam_def) | 
| 
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 | 293 | |
| 
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 | 294 | instance | 
| 
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 | 295 | by default | 
| 
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 | 296 | |
| 
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 | 297 | end | 
| 
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 | 298 | |
| 
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 | 299 | lemmas size_lam[simp] = | 
| 
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 | 300 | lam_raw.size(4)[quot_lifted] | 
| 
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 | 301 | lam_raw.size(5)[quot_lifted] | 
| 
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 | 302 | lam_raw.size(6)[quot_lifted] | 
| 
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 | 303 | |
| 
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 | 304 | (* is this needed? *) | 
| 
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 | 305 | lemma [measure_function]: | 
| 
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 | 306 | "is_measure (size::lam\<Rightarrow>nat)" | 
| 
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 | 307 | by (rule is_measure_trivial) | 
| 
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 | 308 | |
| 
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 | 309 | section {* Matching *}
 | 
| 
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 | 310 | |
| 
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 | 311 | definition | 
| 
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 | 312 |   MATCH :: "('c::pt \<Rightarrow> (bool * 'a::pt * 'b::pt)) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b"
 | 
| 
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 | 313 | where | 
| 
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 | 314 | "MATCH M d x \<equiv> if (\<exists>!r. \<exists>q. M q = (True, x, r)) then (THE r. \<exists>q. M q = (True, x, r)) else d" | 
| 
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 | 315 | |
| 
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 | 316 | (* | 
| 
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 | 317 | lemma MATCH_eqvt: | 
| 
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 | 318 | shows "p \<bullet> (MATCH M d x) = MATCH (p \<bullet> M) (p \<bullet> d) (p \<bullet> x)" | 
| 
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 | 319 | unfolding MATCH_def | 
| 
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 | 320 | apply(perm_simp the_eqvt) | 
| 
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 | 321 | apply (tactic {* Nominal_Permeq.eqvt_tac @{context} 1 *})
 | 
| 
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 | 322 | apply(simp) | 
| 
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 | 323 | thm eqvts_raw | 
| 
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 | 324 | apply(subst if_eqvt) | 
| 
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 | 325 | apply(subst ex1_eqvt) | 
| 
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 | 326 | apply(subst permute_fun_def) | 
| 
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 | 327 | apply(subst ex_eqvt) | 
| 
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 | 328 | apply(subst permute_fun_def) | 
| 
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 | 329 | apply(subst eq_eqvt) | 
| 
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 | 330 | apply(subst permute_fun_app_eq[where f="M"]) | 
| 
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 | 331 | apply(simp only: permute_minus_cancel) | 
| 
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 | 332 | apply(subst permute_prod.simps) | 
| 
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 | 333 | apply(subst permute_prod.simps) | 
| 
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 | 334 | apply(simp only: permute_minus_cancel) | 
| 
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 | 335 | apply(simp only: permute_bool_def) | 
| 
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 | 336 | apply(simp) | 
| 
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 | 337 | apply(subst ex1_eqvt) | 
| 
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 | 338 | apply(subst permute_fun_def) | 
| 
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 | 339 | apply(subst ex_eqvt) | 
| 
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 | 340 | apply(subst permute_fun_def) | 
| 
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 | 341 | apply(subst eq_eqvt) | 
| 
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 | 342 | |
| 
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 | 343 | apply(simp only: eqvts) | 
| 
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 | 344 | apply(simp) | 
| 
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 | 345 | apply(subgoal_tac "(p \<bullet> (\<exists>!r. \<exists>q. M q = (True, x, r))) = (\<exists>!r. \<exists>q. (p \<bullet> M) q = (True, p \<bullet> x, r))") | 
| 
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 | 346 | apply(drule sym) | 
| 
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 | 347 | apply(simp) | 
| 
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 | 348 | apply(rule impI) | 
| 
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 | 349 | apply(simp add: perm_bool) | 
| 
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 | 350 | apply(rule trans) | 
| 
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 | 351 | apply(rule pt_the_eqvt[OF pta at]) | 
| 
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 | 352 | apply(assumption) | 
| 
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 | 353 | apply(simp add: pt_ex_eqvt[OF pt at]) | 
| 
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 | 354 | apply(simp add: pt_eq_eqvt[OF ptb at]) | 
| 
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 | 355 | apply(rule cheat) | 
| 
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 | 356 | apply(rule trans) | 
| 
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 | 357 | apply(rule pt_ex1_eqvt) | 
| 
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 | 358 | apply(rule pta) | 
| 
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 | 359 | apply(rule at) | 
| 
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 | 360 | apply(simp add: pt_ex_eqvt[OF pt at]) | 
| 
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 | 361 | apply(simp add: pt_eq_eqvt[OF ptb at]) | 
| 
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 | 362 | apply(subst pt_pi_rev[OF pta at]) | 
| 
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 | 363 | apply(subst pt_fun_app_eq[OF pt at]) | 
| 
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 | 364 | apply(subst pt_pi_rev[OF pt at]) | 
| 
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 | 365 | apply(simp) | 
| 
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 | 366 | done | 
| 
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 | 367 | |
| 
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 | 368 | lemma MATCH_cng: | 
| 
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 | 369 | assumes a: "M1 = M2" "d1 = d2" | 
| 
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 | 370 | shows "MATCH M1 d1 x = MATCH M2 d2 x" | 
| 
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 | 371 | using a by simp | 
| 
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 | 372 | |
| 
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 | 373 | lemma MATCH_eq: | 
| 
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 | 374 | assumes a: "t = l x" "G x" "\<And>x'. t = l x' \<Longrightarrow> G x' \<Longrightarrow> r x' = r x" | 
| 
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 | 375 | shows "MATCH (\<lambda>x. (G x, l x, r x)) d t = r x" | 
| 
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 | 376 | using a | 
| 
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 | 377 | unfolding MATCH_def | 
| 
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 | 378 | apply(subst if_P) | 
| 
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 | 379 | apply(rule_tac a="r x" in ex1I) | 
| 
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 | 380 | apply(rule_tac x="x" in exI) | 
| 
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 | 381 | apply(blast) | 
| 
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 | 382 | apply(erule exE) | 
| 
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 | 383 | apply(drule_tac x="q" in meta_spec) | 
| 
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 | 384 | apply(auto)[1] | 
| 
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 | 385 | apply(rule the_equality) | 
| 
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 | 386 | apply(blast) | 
| 
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 | 387 | apply(erule exE) | 
| 
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 | 388 | apply(drule_tac x="q" in meta_spec) | 
| 
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 | 389 | apply(auto)[1] | 
| 
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 | 390 | done | 
| 
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 | 391 | |
| 
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 | 392 | lemma MATCH_eq2: | 
| 
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 | 393 | assumes a: "t = l x1 x2" "G x1 x2" "\<And>x1' x2'. t = l x1' x2' \<Longrightarrow> G x1' x2' \<Longrightarrow> r x1' x2' = r x1 x2" | 
| 
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 | 394 | shows "MATCH (\<lambda>(x1,x2). (G x1 x2, l x1 x2, r x1 x2)) d t = r x1 x2" | 
| 
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 | 395 | sorry | 
| 
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 | 396 | |
| 
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 | 397 | lemma MATCH_neq: | 
| 
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 | 398 | assumes a: "\<And>x. t = l x \<Longrightarrow> G x \<Longrightarrow> False" | 
| 
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 | 399 | shows "MATCH (\<lambda>x. (G x, l x, r x)) d t = d" | 
| 
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 | 400 | using a | 
| 
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 | 401 | unfolding MATCH_def | 
| 
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 | 402 | apply(subst if_not_P) | 
| 
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 | 403 | apply(blast) | 
| 
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 | 404 | apply(rule refl) | 
| 
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 | 405 | done | 
| 
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 | 406 | |
| 
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 | 407 | lemma MATCH_neq2: | 
| 
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 | 408 | assumes a: "\<And>x1 x2. t = l x1 x2 \<Longrightarrow> G x1 x2 \<Longrightarrow> False" | 
| 
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 | 409 | shows "MATCH (\<lambda>(x1,x2). (G x1 x2, l x1 x2, r x1 x2)) d t = d" | 
| 
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 | 410 | using a | 
| 
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 | 411 | unfolding MATCH_def | 
| 
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 | 412 | apply(subst if_not_P) | 
| 
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 | 413 | apply(auto) | 
| 
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 | 414 | done | 
| 
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 | 415 | *) | 
| 
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 | 416 | |
| 1947 | 417 | |
| 1949 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 418 | ML {*
 | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 419 | |
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 420 | fun prove_strong_ind (pred_name, avoids) ctxt = | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 421 | Proof.theorem NONE (K I) [] ctxt | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 422 | |
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 423 | local structure P = OuterParse and K = OuterKeyword in | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 424 | |
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 425 | val _ = | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 426 | OuterSyntax.local_theory_to_proof "nominal_inductive" | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 427 | "proves strong induction theorem for inductive predicate involving nominal datatypes" K.thy_goal | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 428 | (P.xname -- (Scan.optional (P.$$$ "avoids" |-- P.enum1 "|" (P.name -- | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 429 | (P.$$$ ":" |-- P.and_list1 P.term))) []) >> prove_strong_ind) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 430 | |
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 431 | end; | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 432 | |
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 433 | *} | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 434 | |
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 435 | nominal_inductive typing | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 436 | |
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 437 | |
| 1947 | 438 | |
| 1594 | 439 | end | 
| 440 | ||
| 441 | ||
| 442 |