| author | Christian Urban <urbanc@in.tum.de> | 
| Sun, 13 Jun 2010 13:42:37 +0200 | |
| changeset 2232 | f49b5dfabd59 | 
| parent 2213 | 231a20534950 | 
| child 2312 | ad03df7e8056 | 
| permissions | -rw-r--r-- | 
| 2083 
9568f9f31822
tuned file names for examples
 Christian Urban <urbanc@in.tum.de> parents: 
2064diff
changeset | 1 | theory CoreHaskell | 
| 2044 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 2 | imports "../NewParser" | 
| 1601 | 3 | begin | 
| 4 | ||
| 5 | (* core haskell *) | |
| 6 | ||
| 2100 
705dc7532ee3
added comment about bind_set
 Christian Urban <urbanc@in.tum.de> parents: 
2095diff
changeset | 7 | (* at the moment it is hard coded that shallow binders | 
| 
705dc7532ee3
added comment about bind_set
 Christian Urban <urbanc@in.tum.de> parents: 
2095diff
changeset | 8 | need to use bind_set *) | 
| 
705dc7532ee3
added comment about bind_set
 Christian Urban <urbanc@in.tum.de> parents: 
2095diff
changeset | 9 | |
| 1601 | 10 | atom_decl var | 
| 1698 
69472e74bd3b
Update according to paper
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1695diff
changeset | 11 | atom_decl cvar | 
| 1601 | 12 | atom_decl tvar | 
| 13 | ||
| 14 | nominal_datatype tkind = | |
| 15 | KStar | |
| 16 | | KFun "tkind" "tkind" | |
| 17 | and ckind = | |
| 18 | CKEq "ty" "ty" | |
| 19 | and ty = | |
| 20 | TVar "tvar" | |
| 1684 
b9e4c812d2c6
Core Haskell can now use proper strings.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1667diff
changeset | 21 | | TC "string" | 
| 1601 | 22 | | TApp "ty" "ty" | 
| 1684 
b9e4c812d2c6
Core Haskell can now use proper strings.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1667diff
changeset | 23 | | TFun "string" "ty_lst" | 
| 2095 
ae94bae5bb93
Restore set bindings in CoreHaskell
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2083diff
changeset | 24 | | TAll tv::"tvar" "tkind" T::"ty" bind_set tv in T | 
| 1698 
69472e74bd3b
Update according to paper
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1695diff
changeset | 25 | | TEq "ckind" "ty" | 
| 1606 
75403378068b
Initial list unfoldings in Core Haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1601diff
changeset | 26 | and ty_lst = | 
| 
75403378068b
Initial list unfoldings in Core Haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1601diff
changeset | 27 | TsNil | 
| 
75403378068b
Initial list unfoldings in Core Haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1601diff
changeset | 28 | | TsCons "ty" "ty_lst" | 
| 1601 | 29 | and co = | 
| 1699 
2dca07aca287
small changes in the core-haskell spec
 Christian Urban <urbanc@in.tum.de> parents: 
1698diff
changeset | 30 | CVar "cvar" | 
| 1698 
69472e74bd3b
Update according to paper
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1695diff
changeset | 31 | | CConst "string" | 
| 1601 | 32 | | CApp "co" "co" | 
| 1684 
b9e4c812d2c6
Core Haskell can now use proper strings.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1667diff
changeset | 33 | | CFun "string" "co_lst" | 
| 2095 
ae94bae5bb93
Restore set bindings in CoreHaskell
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2083diff
changeset | 34 | | CAll cv::"cvar" "ckind" C::"co" bind_set cv in C | 
| 1698 
69472e74bd3b
Update according to paper
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1695diff
changeset | 35 | | CEq "ckind" "co" | 
| 
69472e74bd3b
Update according to paper
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1695diff
changeset | 36 | | CRefl "ty" | 
| 1601 | 37 | | CSym "co" | 
| 1698 
69472e74bd3b
Update according to paper
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1695diff
changeset | 38 | | CCir "co" "co" | 
| 
69472e74bd3b
Update according to paper
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1695diff
changeset | 39 | | CAt "co" "ty" | 
| 1601 | 40 | | CLeft "co" | 
| 41 | | CRight "co" | |
| 1698 
69472e74bd3b
Update according to paper
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1695diff
changeset | 42 | | CSim "co" "co" | 
| 1601 | 43 | | CRightc "co" | 
| 44 | | CLeftc "co" | |
| 45 | | CCoe "co" "co" | |
| 1606 
75403378068b
Initial list unfoldings in Core Haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1601diff
changeset | 46 | and co_lst = | 
| 
75403378068b
Initial list unfoldings in Core Haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1601diff
changeset | 47 | CsNil | 
| 
75403378068b
Initial list unfoldings in Core Haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1601diff
changeset | 48 | | CsCons "co" "co_lst" | 
| 1609 
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1606diff
changeset | 49 | and trm = | 
| 1601 | 50 | Var "var" | 
| 1699 
2dca07aca287
small changes in the core-haskell spec
 Christian Urban <urbanc@in.tum.de> parents: 
1698diff
changeset | 51 | | K "string" | 
| 2095 
ae94bae5bb93
Restore set bindings in CoreHaskell
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2083diff
changeset | 52 | | LAMT tv::"tvar" "tkind" t::"trm" bind_set tv in t | 
| 
ae94bae5bb93
Restore set bindings in CoreHaskell
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2083diff
changeset | 53 | | LAMC cv::"cvar" "ckind" t::"trm" bind_set cv in t | 
| 1698 
69472e74bd3b
Update according to paper
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1695diff
changeset | 54 | | AppT "trm" "ty" | 
| 
69472e74bd3b
Update according to paper
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1695diff
changeset | 55 | | AppC "trm" "co" | 
| 2095 
ae94bae5bb93
Restore set bindings in CoreHaskell
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2083diff
changeset | 56 | | Lam v::"var" "ty" t::"trm" bind_set v in t | 
| 1601 | 57 | | App "trm" "trm" | 
| 2095 
ae94bae5bb93
Restore set bindings in CoreHaskell
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2083diff
changeset | 58 | | Let x::"var" "ty" "trm" t::"trm" bind_set x in t | 
| 1609 
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1606diff
changeset | 59 | | Case "trm" "assoc_lst" | 
| 
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1606diff
changeset | 60 | | Cast "trm" "ty" --"ty is supposed to be a coercion type only" | 
| 
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1606diff
changeset | 61 | and assoc_lst = | 
| 
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1606diff
changeset | 62 | ANil | 
| 
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1606diff
changeset | 63 | | ACons p::"pat" t::"trm" "assoc_lst" bind "bv p" in t | 
| 1601 | 64 | and pat = | 
| 1700 | 65 | Kpat "string" "tvars" "cvars" "vars" | 
| 66 | and vars = | |
| 67 | VsNil | |
| 68 | | VsCons "var" "ty" "vars" | |
| 69 | and tvars = | |
| 70 | TvsNil | |
| 71 | | TvsCons "tvar" "tkind" "tvars" | |
| 72 | and cvars = | |
| 73 | CvsNil | |
| 74 | | CvsCons "cvar" "ckind" "cvars" | |
| 1601 | 75 | binder | 
| 1695 | 76 | bv :: "pat \<Rightarrow> atom list" | 
| 1700 | 77 | and bv_vs :: "vars \<Rightarrow> atom list" | 
| 78 | and bv_tvs :: "tvars \<Rightarrow> atom list" | |
| 79 | and bv_cvs :: "cvars \<Rightarrow> atom list" | |
| 1601 | 80 | where | 
| 1700 | 81 | "bv (Kpat s tvts tvcs vs) = append (bv_tvs tvts) (append (bv_cvs tvcs) (bv_vs vs))" | 
| 82 | | "bv_vs VsNil = []" | |
| 83 | | "bv_vs (VsCons v k t) = (atom v) # bv_vs t" | |
| 84 | | "bv_tvs TvsNil = []" | |
| 85 | | "bv_tvs (TvsCons v k t) = (atom v) # bv_tvs t" | |
| 86 | | "bv_cvs CvsNil = []" | |
| 2213 
231a20534950
improved abstract, some tuning
 Christian Urban <urbanc@in.tum.de> parents: 
2142diff
changeset | 87 | | "bv_cvs (CvsCons v k t) = (atom v) # bv_cvs t" | 
| 
231a20534950
improved abstract, some tuning
 Christian Urban <urbanc@in.tum.de> parents: 
2142diff
changeset | 88 | |
| 
231a20534950
improved abstract, some tuning
 Christian Urban <urbanc@in.tum.de> parents: 
2142diff
changeset | 89 | |
| 1621 
a40dbea68d0b
Core Haskell experiments.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1615diff
changeset | 90 | |
| 1700 | 91 | lemmas fv_supp=tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vars_tvars_cvars.supp(1-9,11,13,15) | 
| 92 | lemmas supp=tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vars_tvars_cvars.fv[simplified fv_supp] | |
| 93 | lemmas perm=tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vars_tvars_cvars.perm | |
| 94 | lemmas eq_iff=tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vars_tvars_cvars.eq_iff | |
| 95 | lemmas inducts=tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vars_tvars_cvars.inducts | |
| 1634 
b6656b707a8d
Solved one of the strong-induction goals.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1632diff
changeset | 96 | |
| 1700 | 97 | lemmas alpha_inducts=alpha_tkind_raw_alpha_ckind_raw_alpha_ty_raw_alpha_ty_lst_raw_alpha_co_raw_alpha_co_lst_raw_alpha_trm_raw_alpha_assoc_lst_raw_alpha_pat_raw_alpha_vars_raw_alpha_tvars_raw_alpha_cvars_raw_alpha_bv_raw_alpha_bv_vs_raw_alpha_bv_tvs_raw_alpha_bv_cvs_raw.inducts | 
| 98 | lemmas alpha_intros=alpha_tkind_raw_alpha_ckind_raw_alpha_ty_raw_alpha_ty_lst_raw_alpha_co_raw_alpha_co_lst_raw_alpha_trm_raw_alpha_assoc_lst_raw_alpha_pat_raw_alpha_vars_raw_alpha_tvars_raw_alpha_cvars_raw_alpha_bv_raw_alpha_bv_vs_raw_alpha_bv_tvs_raw_alpha_bv_cvs_raw.intros | |
| 1647 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 99 | |
| 1634 
b6656b707a8d
Solved one of the strong-induction goals.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1632diff
changeset | 100 | lemma fresh_star_minus_perm: "as \<sharp>* - p = as \<sharp>* (p :: perm)" | 
| 
b6656b707a8d
Solved one of the strong-induction goals.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1632diff
changeset | 101 | unfolding fresh_star_def Ball_def | 
| 
b6656b707a8d
Solved one of the strong-induction goals.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1632diff
changeset | 102 | by auto (simp_all add: fresh_minus_perm) | 
| 1615 
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1609diff
changeset | 103 | |
| 1700 | 104 | primrec permute_bv_vs_raw | 
| 105 | where "permute_bv_vs_raw p VsNil_raw = VsNil_raw" | |
| 106 | | "permute_bv_vs_raw p (VsCons_raw v t l) = VsCons_raw (p \<bullet> v) t (permute_bv_vs_raw p l)" | |
| 107 | primrec permute_bv_cvs_raw | |
| 108 | where "permute_bv_cvs_raw p CvsNil_raw = CvsNil_raw" | |
| 109 | | "permute_bv_cvs_raw p (CvsCons_raw v t l) = CvsCons_raw (p \<bullet> v) t (permute_bv_cvs_raw p l)" | |
| 110 | primrec permute_bv_tvs_raw | |
| 111 | where "permute_bv_tvs_raw p TvsNil_raw = TvsNil_raw" | |
| 112 | | "permute_bv_tvs_raw p (TvsCons_raw v t l) = TvsCons_raw (p \<bullet> v) t (permute_bv_tvs_raw p l)" | |
| 1647 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 113 | primrec permute_bv_raw | 
| 1700 | 114 | where "permute_bv_raw p (Kpat_raw c l1 l2 l3) = | 
| 115 | Kpat_raw c (permute_bv_tvs_raw p l1) (permute_bv_cvs_raw p l2) (permute_bv_vs_raw p l3)" | |
| 1647 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 116 | |
| 1700 | 117 | quotient_definition "permute_bv_vs :: perm \<Rightarrow> vars \<Rightarrow> vars" | 
| 118 | is "permute_bv_vs_raw" | |
| 119 | quotient_definition "permute_bv_cvs :: perm \<Rightarrow> cvars \<Rightarrow> cvars" | |
| 120 | is "permute_bv_cvs_raw" | |
| 121 | quotient_definition "permute_bv_tvs :: perm \<Rightarrow> tvars \<Rightarrow> tvars" | |
| 122 | is "permute_bv_tvs_raw" | |
| 1647 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 123 | quotient_definition "permute_bv :: perm \<Rightarrow> pat \<Rightarrow> pat" | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 124 | is "permute_bv_raw" | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 125 | |
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 126 | lemma rsp_pre: | 
| 1700 | 127 | "alpha_tvars_raw d a \<Longrightarrow> alpha_tvars_raw (permute_bv_tvs_raw x d) (permute_bv_tvs_raw x a)" | 
| 128 | "alpha_cvars_raw e b \<Longrightarrow> alpha_cvars_raw (permute_bv_cvs_raw x e) (permute_bv_cvs_raw x b)" | |
| 129 | "alpha_vars_raw f c \<Longrightarrow> alpha_vars_raw (permute_bv_vs_raw x f) (permute_bv_vs_raw x c)" | |
| 1647 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 130 | apply (erule_tac [!] alpha_inducts) | 
| 2044 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 131 | apply (simp_all only: alpha_intros perm permute_bv_tvs_raw.simps permute_bv_cvs_raw.simps permute_bv_vs_raw.simps) | 
| 1647 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 132 | done | 
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 133 | |
| 1647 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 134 | lemma [quot_respect]: | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 135 | "(op = ===> alpha_pat_raw ===> alpha_pat_raw) permute_bv_raw permute_bv_raw" | 
| 1700 | 136 | "(op = ===> alpha_tvars_raw ===> alpha_tvars_raw) permute_bv_tvs_raw permute_bv_tvs_raw" | 
| 137 | "(op = ===> alpha_cvars_raw ===> alpha_cvars_raw) permute_bv_cvs_raw permute_bv_cvs_raw" | |
| 138 | "(op = ===> alpha_vars_raw ===> alpha_vars_raw) permute_bv_vs_raw permute_bv_vs_raw" | |
| 1647 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 139 | apply (simp_all add: rsp_pre) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 140 | apply clarify | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 141 | apply (erule_tac alpha_inducts) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 142 | apply (simp_all) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 143 | apply (rule alpha_intros) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 144 | apply (simp_all add: rsp_pre) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 145 | done | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 146 | |
| 1684 
b9e4c812d2c6
Core Haskell can now use proper strings.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1667diff
changeset | 147 | thm permute_bv_raw.simps[no_vars] | 
| 1700 | 148 | permute_bv_vs_raw.simps[quot_lifted] | 
| 149 | permute_bv_cvs_raw.simps[quot_lifted] | |
| 150 | permute_bv_tvs_raw.simps[quot_lifted] | |
| 1684 
b9e4c812d2c6
Core Haskell can now use proper strings.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1667diff
changeset | 151 | |
| 
b9e4c812d2c6
Core Haskell can now use proper strings.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1667diff
changeset | 152 | lemma permute_bv_pre: | 
| 1700 | 153 | "permute_bv p (Kpat c l1 l2 l3) = | 
| 154 | Kpat c (permute_bv_tvs p l1) (permute_bv_cvs p l2) (permute_bv_vs p l3)" | |
| 1684 
b9e4c812d2c6
Core Haskell can now use proper strings.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1667diff
changeset | 155 | by (lifting permute_bv_raw.simps) | 
| 
b9e4c812d2c6
Core Haskell can now use proper strings.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1667diff
changeset | 156 | |
| 
b9e4c812d2c6
Core Haskell can now use proper strings.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1667diff
changeset | 157 | lemmas permute_bv[simp] = | 
| 
b9e4c812d2c6
Core Haskell can now use proper strings.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1667diff
changeset | 158 | permute_bv_pre | 
| 1700 | 159 | permute_bv_vs_raw.simps[quot_lifted] | 
| 160 | permute_bv_cvs_raw.simps[quot_lifted] | |
| 161 | permute_bv_tvs_raw.simps[quot_lifted] | |
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 162 | |
| 1647 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 163 | lemma perm_bv1: | 
| 1700 | 164 | "p \<bullet> bv_cvs b = bv_cvs (permute_bv_cvs p b)" | 
| 165 | "p \<bullet> bv_tvs c = bv_tvs (permute_bv_tvs p c)" | |
| 166 | "p \<bullet> bv_vs d = bv_vs (permute_bv_vs p d)" | |
| 1647 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 167 | apply(induct b rule: inducts(12)) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 168 | apply(simp_all add:permute_bv eqvts) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 169 | apply(induct c rule: inducts(11)) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 170 | apply(simp_all add:permute_bv eqvts) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 171 | apply(induct d rule: inducts(10)) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 172 | apply(simp_all add:permute_bv eqvts) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 173 | done | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 174 | |
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 175 | lemma perm_bv2: | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 176 | "p \<bullet> bv l = bv (permute_bv p l)" | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 177 | apply(induct l rule: inducts(9)) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 178 | apply(simp_all add:permute_bv) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 179 | apply(simp add: perm_bv1[symmetric]) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 180 | apply(simp add: eqvts) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 181 | done | 
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 182 | |
| 1648 | 183 | lemma alpha_perm_bn1: | 
| 2044 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 184 | "alpha_bv_tvs tvars (permute_bv_tvs q tvars)" | 
| 1700 | 185 | "alpha_bv_cvs cvars (permute_bv_cvs q cvars)" | 
| 186 | "alpha_bv_vs vars (permute_bv_vs q vars)" | |
| 187 | apply(induct tvars rule: inducts(11)) | |
| 1648 | 188 | apply(simp_all add:permute_bv eqvts eq_iff) | 
| 1700 | 189 | apply(induct cvars rule: inducts(12)) | 
| 1648 | 190 | apply(simp_all add:permute_bv eqvts eq_iff) | 
| 1700 | 191 | apply(induct vars rule: inducts(10)) | 
| 1648 | 192 | apply(simp_all add:permute_bv eqvts eq_iff) | 
| 193 | done | |
| 194 | ||
| 195 | lemma alpha_perm_bn: | |
| 2142 
c39d4fe31100
moved the exporting part into the parser (this is still a hack); re-added CoreHaskell again to the examples - there seems to be a problem with the variable name pat
 Christian Urban <urbanc@in.tum.de> parents: 
2120diff
changeset | 196 | "alpha_bv pt (permute_bv q pt)" | 
| 
c39d4fe31100
moved the exporting part into the parser (this is still a hack); re-added CoreHaskell again to the examples - there seems to be a problem with the variable name pat
 Christian Urban <urbanc@in.tum.de> parents: 
2120diff
changeset | 197 | apply(induct pt rule: inducts(9)) | 
| 1648 | 198 | apply(simp_all add:permute_bv eqvts eq_iff alpha_perm_bn1) | 
| 199 | done | |
| 200 | ||
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 201 | lemma ACons_subst: | 
| 2142 
c39d4fe31100
moved the exporting part into the parser (this is still a hack); re-added CoreHaskell again to the examples - there seems to be a problem with the variable name pat
 Christian Urban <urbanc@in.tum.de> parents: 
2120diff
changeset | 202 | "supp (Abs_lst (bv pt) trm) \<sharp>* q \<Longrightarrow> (ACons pt trm al) = ACons (permute_bv q pt) (q \<bullet> trm) al" | 
| 1648 | 203 | apply (simp only: eq_iff) | 
| 2044 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 204 | apply (simp add: alpha_perm_bn) | 
| 1648 | 205 | apply (rule_tac x="q" in exI) | 
| 206 | apply (simp add: alphas) | |
| 207 | apply (simp add: perm_bv2[symmetric]) | |
| 1658 | 208 | apply (simp add: supp_abs) | 
| 1648 | 209 | apply (simp add: fv_supp) | 
| 2104 
2205b572bc9b
Declare alpha_gen_eqvt as eqvt and change the proofs that used 'eqvts[symmetric]'
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2100diff
changeset | 210 | apply (simp add: supp_eqvt[symmetric] set_eqvt[symmetric] Diff_eqvt[symmetric]) | 
| 1648 | 211 | apply (rule supp_perm_eq[symmetric]) | 
| 212 | apply (subst supp_finite_atom_set) | |
| 213 | apply (rule finite_Diff) | |
| 214 | apply (simp add: finite_supp) | |
| 215 | apply (assumption) | |
| 216 | done | |
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 217 | |
| 1647 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 218 | lemma permute_bv_zero1: | 
| 1700 | 219 | "permute_bv_cvs 0 b = b" | 
| 220 | "permute_bv_tvs 0 c = c" | |
| 221 | "permute_bv_vs 0 d = d" | |
| 1647 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 222 | apply(induct b rule: inducts(12)) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 223 | apply(simp_all add:permute_bv eqvts) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 224 | apply(induct c rule: inducts(11)) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 225 | apply(simp_all add:permute_bv eqvts) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 226 | apply(induct d rule: inducts(10)) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 227 | apply(simp_all add:permute_bv eqvts) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 228 | done | 
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 229 | |
| 1647 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 230 | lemma permute_bv_zero2: | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 231 | "permute_bv 0 a = a" | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 232 | apply(induct a rule: inducts(9)) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 233 | apply(simp_all add:permute_bv eqvts permute_bv_zero1) | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 234 | done | 
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 235 | |
| 1700 | 236 | lemma fv_alpha1: "fv_bv_tvs x \<sharp>* pa \<Longrightarrow> alpha_bv_tvs (pa \<bullet> x) x" | 
| 2044 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 237 | apply (induct x rule: inducts(11)) | 
| 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 238 | apply (simp_all add: eq_iff fresh_star_union) | 
| 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 239 | done | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 240 | |
| 1700 | 241 | lemma fv_alpha2: "fv_bv_cvs x \<sharp>* pa \<Longrightarrow> alpha_bv_cvs (pa \<bullet> x) x" | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 242 | apply (induct x rule: inducts(12)) | 
| 2044 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 243 | apply (rule TrueI)+ | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 244 | apply (simp_all add: eq_iff fresh_star_union) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 245 | apply (subst supp_perm_eq) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 246 | apply (simp_all add: fv_supp) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 247 | done | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 248 | |
| 1700 | 249 | lemma fv_alpha3: "fv_bv_vs x \<sharp>* pa \<Longrightarrow> alpha_bv_vs (pa \<bullet> x) x" | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 250 | apply (induct x rule: inducts(10)) | 
| 2044 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 251 | apply (rule TrueI)+ | 
| 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 252 | apply (simp_all add: fresh_star_union eq_iff) | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 253 | apply (subst supp_perm_eq) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 254 | apply (simp_all add: fv_supp) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 255 | done | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 256 | |
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 257 | lemma fv_alpha: "fv_bv x \<sharp>* pa \<Longrightarrow> alpha_bv (pa \<bullet> x) x" | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 258 | apply (induct x rule: inducts(9)) | 
| 2044 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 259 | apply (rule TrueI)+ | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 260 | apply (simp_all add: eq_iff fresh_star_union) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 261 | apply (simp add: fv_alpha1 fv_alpha2 fv_alpha3) | 
| 2044 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 262 | apply (subst supp_perm_eq) | 
| 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 263 | apply (simp_all add: fv_supp) | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 264 | done | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 265 | |
| 1700 | 266 | lemma fin1: "finite (fv_bv_tvs x)" | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 267 | apply (induct x rule: inducts(11)) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 268 | apply (simp_all add: fv_supp finite_supp) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 269 | done | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 270 | |
| 1700 | 271 | lemma fin2: "finite (fv_bv_cvs x)" | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 272 | apply (induct x rule: inducts(12)) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 273 | apply (simp_all add: fv_supp finite_supp) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 274 | done | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 275 | |
| 1700 | 276 | lemma fin3: "finite (fv_bv_vs x)" | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 277 | apply (induct x rule: inducts(10)) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 278 | apply (simp_all add: fv_supp finite_supp) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 279 | done | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 280 | |
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 281 | lemma fin_fv_bv: "finite (fv_bv x)" | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 282 | apply (induct x rule: inducts(9)) | 
| 2044 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 283 | apply (rule TrueI)+ | 
| 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 284 | defer | 
| 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 285 | apply (rule TrueI)+ | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 286 | apply (simp add: fin1 fin2 fin3) | 
| 2044 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 287 | apply (rule finite_supp) | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 288 | done | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 289 | |
| 1700 | 290 | lemma finb1: "finite (set (bv_tvs x))" | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 291 | apply (induct x rule: inducts(11)) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 292 | apply (simp_all add: fv_supp finite_supp) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 293 | done | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 294 | |
| 1700 | 295 | lemma finb2: "finite (set (bv_cvs x))" | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 296 | apply (induct x rule: inducts(12)) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 297 | apply (simp_all add: fv_supp finite_supp) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 298 | done | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 299 | |
| 1700 | 300 | lemma finb3: "finite (set (bv_vs x))" | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 301 | apply (induct x rule: inducts(10)) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 302 | apply (simp_all add: fv_supp finite_supp) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 303 | done | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 304 | |
| 1695 | 305 | lemma fin_bv: "finite (set (bv x))" | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 306 | apply (induct x rule: inducts(9)) | 
| 2044 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 307 | apply (simp_all add: finb1 finb2 finb3) | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 308 | done | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 309 | |
| 1701 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 310 | lemma strong_induction_principle: | 
| 1631 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 311 | assumes a01: "\<And>b. P1 b KStar" | 
| 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 312 | and a02: "\<And>tkind1 tkind2 b. \<lbrakk>\<And>c. P1 c tkind1; \<And>c. P1 c tkind2\<rbrakk> \<Longrightarrow> P1 b (KFun tkind1 tkind2)" | 
| 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 313 | and a03: "\<And>ty1 ty2 b. \<lbrakk>\<And>c. P3 c ty1; \<And>c. P3 c ty2\<rbrakk> \<Longrightarrow> P2 b (CKEq ty1 ty2)" | 
| 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 314 | and a04: "\<And>tvar b. P3 b (TVar tvar)" | 
| 1684 
b9e4c812d2c6
Core Haskell can now use proper strings.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1667diff
changeset | 315 | and a05: "\<And>string b. P3 b (TC string)" | 
| 1631 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 316 | and a06: "\<And>ty1 ty2 b. \<lbrakk>\<And>c. P3 c ty1; \<And>c. P3 c ty2\<rbrakk> \<Longrightarrow> P3 b (TApp ty1 ty2)" | 
| 1684 
b9e4c812d2c6
Core Haskell can now use proper strings.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1667diff
changeset | 317 | and a07: "\<And>string ty_lst b. \<lbrakk>\<And>c. P4 c ty_lst\<rbrakk> \<Longrightarrow> P3 b (TFun string ty_lst)" | 
| 1632 
68c8666453f7
Started proving strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1631diff
changeset | 318 | and a08: "\<And>tvar tkind ty b. \<lbrakk>\<And>c. P1 c tkind; \<And>c. P3 c ty; atom tvar \<sharp> b\<rbrakk> | 
| 
68c8666453f7
Started proving strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1631diff
changeset | 319 | \<Longrightarrow> P3 b (TAll tvar tkind ty)" | 
| 1701 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 320 | and a09: "\<And>ck ty b. \<lbrakk>\<And>c. P2 c ck; \<And>c. P3 c ty\<rbrakk> \<Longrightarrow> P3 b (TEq ck ty)" | 
| 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 321 | and a10: "\<And>b. P4 b TsNil" | 
| 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 322 | and a11: "\<And>ty ty_lst b. \<lbrakk>\<And>c. P3 c ty; \<And>c. P4 c ty_lst\<rbrakk> \<Longrightarrow> P4 b (TsCons ty ty_lst)" | 
| 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 323 | and a12: "\<And>string b. P5 b (CVar string)" | 
| 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 324 | and a12a:"\<And>str b. P5 b (CConst str)" | 
| 1631 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 325 | and a13: "\<And>co1 co2 b. \<lbrakk>\<And>c. P5 c co1; \<And>c. P5 c co2\<rbrakk> \<Longrightarrow> P5 b (CApp co1 co2)" | 
| 1684 
b9e4c812d2c6
Core Haskell can now use proper strings.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1667diff
changeset | 326 | and a14: "\<And>string co_lst b. \<lbrakk>\<And>c. P6 c co_lst\<rbrakk> \<Longrightarrow> P5 b (CFun string co_lst)" | 
| 1632 
68c8666453f7
Started proving strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1631diff
changeset | 327 | and a15: "\<And>tvar ckind co b. \<lbrakk>\<And>c. P2 c ckind; \<And>c. P5 c co; atom tvar \<sharp> b\<rbrakk> | 
| 
68c8666453f7
Started proving strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1631diff
changeset | 328 | \<Longrightarrow> P5 b (CAll tvar ckind co)" | 
| 1701 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 329 | and a16: "\<And>ck co b. \<lbrakk>\<And>c. P2 c ck; \<And>c. P5 c co\<rbrakk> \<Longrightarrow> P5 b (CEq ck co)" | 
| 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 330 | and a17: "\<And>ty b. \<lbrakk>\<And>c. P3 c ty\<rbrakk> \<Longrightarrow> P5 b (CRefl ty)" | 
| 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 331 | and a17a: "\<And>co b. \<lbrakk>\<And>c. P5 c co\<rbrakk> \<Longrightarrow> P5 b (CSym co)" | 
| 1631 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 332 | and a18: "\<And>co1 co2 b. \<lbrakk>\<And>c. P5 c co1; \<And>c. P5 c co2\<rbrakk> \<Longrightarrow> P5 b (CCir co1 co2)" | 
| 1701 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 333 | and a18a:"\<And>co ty b. \<lbrakk>\<And>c. P5 c co; \<And>c. P3 c ty\<rbrakk> \<Longrightarrow> P5 b (CAt co ty)" | 
| 1631 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 334 | and a19: "\<And>co b. \<lbrakk>\<And>c. P5 c co\<rbrakk> \<Longrightarrow> P5 b (CLeft co)" | 
| 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 335 | and a20: "\<And>co b. \<lbrakk>\<And>c. P5 c co\<rbrakk> \<Longrightarrow> P5 b (CRight co)" | 
| 1701 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 336 | and a21: "\<And>co1 co2 b. \<lbrakk>\<And>c. P5 c co1; \<And>c. P5 c co2\<rbrakk> \<Longrightarrow> P5 b (CSim co1 co2)" | 
| 1631 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 337 | and a22: "\<And>co b. \<lbrakk>\<And>c. P5 c co\<rbrakk> \<Longrightarrow> P5 b (CRightc co)" | 
| 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 338 | and a23: "\<And>co b. \<lbrakk>\<And>c. P5 c co\<rbrakk> \<Longrightarrow> P5 b (CLeftc co)" | 
| 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 339 | and a24: "\<And>co1 co2 b. \<lbrakk>\<And>c. P5 c co1; \<And>c. P5 c co2\<rbrakk> \<Longrightarrow> P5 b (CCoe co1 co2)" | 
| 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 340 | and a25: "\<And>b. P6 b CsNil" | 
| 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 341 | and a26: "\<And>co co_lst b. \<lbrakk>\<And>c. P5 c co; \<And>c. P6 c co_lst\<rbrakk> \<Longrightarrow> P6 b (CsCons co co_lst)" | 
| 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 342 | and a27: "\<And>var b. P7 b (Var var)" | 
| 1701 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 343 | and a28: "\<And>string b. P7 b (K string)" | 
| 1632 
68c8666453f7
Started proving strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1631diff
changeset | 344 | and a29: "\<And>tvar tkind trm b. \<lbrakk>\<And>c. P1 c tkind; \<And>c. P7 c trm; atom tvar \<sharp> b\<rbrakk> | 
| 
68c8666453f7
Started proving strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1631diff
changeset | 345 | \<Longrightarrow> P7 b (LAMT tvar tkind trm)" | 
| 
68c8666453f7
Started proving strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1631diff
changeset | 346 | and a30: "\<And>tvar ckind trm b. \<lbrakk>\<And>c. P2 c ckind; \<And>c. P7 c trm; atom tvar \<sharp> b\<rbrakk> | 
| 
68c8666453f7
Started proving strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1631diff
changeset | 347 | \<Longrightarrow> P7 b (LAMC tvar ckind trm)" | 
| 1701 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 348 | and a31: "\<And>trm ty b. \<lbrakk>\<And>c. P7 c trm; \<And>c. P3 c ty\<rbrakk> \<Longrightarrow> P7 b (AppT trm ty)" | 
| 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 349 | and a31a:"\<And>trm co b. \<lbrakk>\<And>c. P7 c trm; \<And>c. P5 c co\<rbrakk> \<Longrightarrow> P7 b (AppC trm co)" | 
| 1632 
68c8666453f7
Started proving strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1631diff
changeset | 350 | and a32: "\<And>var ty trm b. \<lbrakk>\<And>c. P3 c ty; \<And>c. P7 c trm; atom var \<sharp> b\<rbrakk> \<Longrightarrow> P7 b (Lam var ty trm)" | 
| 1631 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 351 | and a33: "\<And>trm1 trm2 b. \<lbrakk>\<And>c. P7 c trm1; \<And>c. P7 c trm2\<rbrakk> \<Longrightarrow> P7 b (App trm1 trm2)" | 
| 1632 
68c8666453f7
Started proving strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1631diff
changeset | 352 | and a34: "\<And>var ty trm1 trm2 b. \<lbrakk>\<And>c. P3 c ty; \<And>c. P7 c trm1; \<And>c. P7 c trm2; atom var \<sharp> b\<rbrakk> | 
| 1631 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 353 | \<Longrightarrow> P7 b (Let var ty trm1 trm2)" | 
| 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 354 | and a35: "\<And>trm assoc_lst b. \<lbrakk>\<And>c. P7 c trm; \<And>c. P8 c assoc_lst\<rbrakk> \<Longrightarrow> P7 b (Case trm assoc_lst)" | 
| 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 355 | and a36: "\<And>trm ty b. \<lbrakk>\<And>c. P7 c trm; \<And>c. P3 c ty\<rbrakk> \<Longrightarrow> P7 b (Cast trm ty)" | 
| 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 356 | and a37: "\<And>b. P8 b ANil" | 
| 2142 
c39d4fe31100
moved the exporting part into the parser (this is still a hack); re-added CoreHaskell again to the examples - there seems to be a problem with the variable name pat
 Christian Urban <urbanc@in.tum.de> parents: 
2120diff
changeset | 357 | and a38: "\<And>pt trm assoc_lst b. \<lbrakk>\<And>c. P9 c pt; \<And>c. P7 c trm; \<And>c. P8 c assoc_lst; set (bv (pt)) \<sharp>* b\<rbrakk> | 
| 
c39d4fe31100
moved the exporting part into the parser (this is still a hack); re-added CoreHaskell again to the examples - there seems to be a problem with the variable name pat
 Christian Urban <urbanc@in.tum.de> parents: 
2120diff
changeset | 358 | \<Longrightarrow> P8 b (ACons pt trm assoc_lst)" | 
| 1700 | 359 | and a39: "\<And>string tvars cvars vars b. \<lbrakk>\<And>c. P11 c tvars; \<And>c. P12 c cvars; \<And>c. P10 c vars\<rbrakk> | 
| 1701 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 360 | \<Longrightarrow> P9 b (Kpat string tvars cvars vars)" | 
| 1700 | 361 | and a40: "\<And>b. P10 b VsNil" | 
| 362 | and a41: "\<And>var ty vars b. \<lbrakk>\<And>c. P3 c ty; \<And>c. P10 c vars\<rbrakk> \<Longrightarrow> P10 b (VsCons var ty vars)" | |
| 363 | and a42: "\<And>b. P11 b TvsNil" | |
| 364 | and a43: "\<And>tvar tkind tvars b. \<lbrakk>\<And>c. P1 c tkind; \<And>c. P11 c tvars\<rbrakk> | |
| 365 | \<Longrightarrow> P11 b (TvsCons tvar tkind tvars)" | |
| 366 | and a44: "\<And>b. P12 b CvsNil" | |
| 367 | and a45: "\<And>tvar ckind cvars b. \<lbrakk>\<And>c. P2 c ckind; \<And>c. P12 c cvars\<rbrakk> | |
| 368 | \<Longrightarrow> P12 b (CvsCons tvar ckind cvars)" | |
| 1632 
68c8666453f7
Started proving strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1631diff
changeset | 369 | shows "P1 (a :: 'a :: pt) tkind \<and> | 
| 1631 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 370 | P2 (b :: 'b :: pt) ckind \<and> | 
| 1634 
b6656b707a8d
Solved one of the strong-induction goals.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1632diff
changeset | 371 |          P3 (c :: 'c :: {pt,fs}) ty \<and>
 | 
| 1631 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 372 | P4 (d :: 'd :: pt) ty_lst \<and> | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 373 |          P5 (e :: 'e :: {pt,fs}) co \<and>
 | 
| 1631 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 374 | P6 (f :: 'f :: pt) co_lst \<and> | 
| 1645 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 375 |          P7 (g :: 'g :: {pt,fs}) trm \<and>
 | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 376 |          P8 (h :: 'h :: {pt,fs}) assoc_lst \<and>
 | 
| 2142 
c39d4fe31100
moved the exporting part into the parser (this is still a hack); re-added CoreHaskell again to the examples - there seems to be a problem with the variable name pat
 Christian Urban <urbanc@in.tum.de> parents: 
2120diff
changeset | 377 | P9 (i :: 'i :: pt) pt \<and> | 
| 1700 | 378 | P10 (j :: 'j :: pt) vars \<and> | 
| 379 | P11 (k :: 'k :: pt) tvars \<and> | |
| 380 | P12 (l :: 'l :: pt) cvars" | |
| 1631 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 381 | proof - | 
| 2142 
c39d4fe31100
moved the exporting part into the parser (this is still a hack); re-added CoreHaskell again to the examples - there seems to be a problem with the variable name pat
 Christian Urban <urbanc@in.tum.de> parents: 
2120diff
changeset | 382 | have a1: "(\<forall>p a. P1 a (p \<bullet> tkind))" and "(\<forall>p b. P2 b (p \<bullet> ckind))" and "(\<forall>p c. P3 c (p \<bullet> ty))" and "(\<forall>p d. P4 d (p \<bullet> ty_lst))" and "(\<forall>p e. P5 e (p \<bullet> co))" and " (\<forall>p f. P6 f (p \<bullet> co_lst))" and "(\<forall>p g. P7 g (p \<bullet> trm))" and "(\<forall>p h. P8 h (p \<bullet> assoc_lst))" and a1:"(\<forall>p q i. P9 i (permute_bv p (q \<bullet> pt)))" and a2:"(\<forall>p q j. P10 j (permute_bv_vs q (p \<bullet> vars)))" and a3:"(\<forall>p q k. P11 k ( permute_bv_tvs q (p \<bullet> tvars)))" and a4:"(\<forall>p q l. P12 l (permute_bv_cvs q (p \<bullet> cvars)))" | 
| 1700 | 383 | apply (induct rule: tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vars_tvars_cvars.inducts) | 
| 1632 
68c8666453f7
Started proving strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1631diff
changeset | 384 |     apply (tactic {* ALLGOALS (REPEAT o rtac allI) *})
 | 
| 
68c8666453f7
Started proving strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1631diff
changeset | 385 |     apply (tactic {* ALLGOALS (TRY o SOLVED' (simp_tac @{simpset} THEN_ALL_NEW resolve_tac @{thms assms} THEN_ALL_NEW asm_full_simp_tac @{simpset})) *})
 | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 386 | |
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 387 | (* GOAL1 *) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 388 | apply(subgoal_tac "\<exists>pa. ((pa \<bullet> (atom (p \<bullet> tvar))) \<sharp> c \<and> | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 389 |                        supp (Abs (p \<bullet> {atom tvar}) (p \<bullet> ty)) \<sharp>* pa)")
 | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 390 | apply clarify | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 391 | apply (simp only: perm) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 392 | apply(rule_tac t="TAll (p \<bullet> tvar) (p \<bullet> tkind) (p \<bullet> ty)" | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 393 | and s="TAll (pa \<bullet> p \<bullet> tvar) (p \<bullet> tkind) (pa \<bullet> p \<bullet> ty)" in subst) | 
| 2044 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 394 | apply (simp add: eq_iff) | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 395 | apply (rule_tac x="-pa" in exI) | 
| 1658 | 396 | apply (simp add: alphas eqvts eqvts_raw supp_abs fv_supp) | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 397 |     apply (rule_tac t="supp (pa \<bullet> p \<bullet> ty) - {atom (pa \<bullet> p \<bullet> tvar)}"
 | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 398 |                 and s="pa \<bullet> (p \<bullet> supp ty - {p \<bullet> atom tvar})" in subst)
 | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 399 | apply (simp add: eqvts) | 
| 2104 
2205b572bc9b
Declare alpha_gen_eqvt as eqvt and change the proofs that used 'eqvts[symmetric]'
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2100diff
changeset | 400 | apply (simp add: supp_eqvt[symmetric] atom_eqvt[symmetric]) | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 401 | apply (rule conjI) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 402 | apply (rule supp_perm_eq) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 403 | apply (simp add: eqvts) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 404 | apply (subst supp_finite_atom_set) | 
| 2104 
2205b572bc9b
Declare alpha_gen_eqvt as eqvt and change the proofs that used 'eqvts[symmetric]'
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2100diff
changeset | 405 | apply (simp add: supp_eqvt[symmetric] atom_eqvt[symmetric] Diff_eqvt[symmetric] finite_eqvt[symmetric] finite_supp True_eqvt) | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 406 | apply (simp add: eqvts) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 407 | apply (subst supp_perm_eq) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 408 | apply (subst supp_finite_atom_set) | 
| 2104 
2205b572bc9b
Declare alpha_gen_eqvt as eqvt and change the proofs that used 'eqvts[symmetric]'
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2100diff
changeset | 409 | apply (simp add: supp_eqvt[symmetric] atom_eqvt[symmetric] Diff_eqvt[symmetric] finite_eqvt[symmetric] finite_supp True_eqvt) | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 410 | apply assumption | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 411 | apply (simp add: fresh_star_minus_perm) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 412 | apply (rule a08) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 413 | apply simp | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 414 | apply(rotate_tac 1) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 415 | apply(erule_tac x="(pa + p)" in allE) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 416 | apply simp | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 417 | apply (simp add: eqvts eqvts_raw) | 
| 1634 
b6656b707a8d
Solved one of the strong-induction goals.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1632diff
changeset | 418 | apply (rule at_set_avoiding2_atom) | 
| 
b6656b707a8d
Solved one of the strong-induction goals.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1632diff
changeset | 419 | apply (simp add: finite_supp) | 
| 
b6656b707a8d
Solved one of the strong-induction goals.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1632diff
changeset | 420 | apply (simp add: finite_supp) | 
| 
b6656b707a8d
Solved one of the strong-induction goals.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1632diff
changeset | 421 | apply (simp add: fresh_def) | 
| 1658 | 422 | apply (simp only: supp_abs eqvts) | 
| 1634 
b6656b707a8d
Solved one of the strong-induction goals.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1632diff
changeset | 423 | apply blast | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 424 | |
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 425 | (* GOAL2 *) | 
| 1701 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 426 | apply(subgoal_tac "\<exists>pa. ((pa \<bullet> (atom (p \<bullet> cvar))) \<sharp> e \<and> | 
| 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 427 |                        supp (Abs (p \<bullet> {atom cvar}) (p \<bullet> co)) \<sharp>* pa)")
 | 
| 1634 
b6656b707a8d
Solved one of the strong-induction goals.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1632diff
changeset | 428 | apply clarify | 
| 
b6656b707a8d
Solved one of the strong-induction goals.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1632diff
changeset | 429 | apply (simp only: perm) | 
| 1701 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 430 | apply(rule_tac t="CAll (p \<bullet> cvar) (p \<bullet> ckind) (p \<bullet> co)" | 
| 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 431 | and s="CAll (pa \<bullet> p \<bullet> cvar) (p \<bullet> ckind) (pa \<bullet> p \<bullet> co)" in subst) | 
| 2044 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 432 | apply (simp add: eq_iff) | 
| 1634 
b6656b707a8d
Solved one of the strong-induction goals.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1632diff
changeset | 433 | apply (rule_tac x="-pa" in exI) | 
| 1658 | 434 | apply (simp add: alphas eqvts eqvts_raw supp_abs fv_supp) | 
| 1701 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 435 |     apply (rule_tac t="supp (pa \<bullet> p \<bullet> co) - {atom (pa \<bullet> p \<bullet> cvar)}"
 | 
| 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 436 |                 and s="pa \<bullet> (p \<bullet> supp co - {p \<bullet> atom cvar})" in subst)
 | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 437 | apply (simp add: eqvts) | 
| 2104 
2205b572bc9b
Declare alpha_gen_eqvt as eqvt and change the proofs that used 'eqvts[symmetric]'
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2100diff
changeset | 438 | apply (simp add: supp_eqvt[symmetric] atom_eqvt[symmetric]) | 
| 1634 
b6656b707a8d
Solved one of the strong-induction goals.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1632diff
changeset | 439 | apply (rule conjI) | 
| 
b6656b707a8d
Solved one of the strong-induction goals.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1632diff
changeset | 440 | apply (rule supp_perm_eq) | 
| 
b6656b707a8d
Solved one of the strong-induction goals.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1632diff
changeset | 441 | apply (simp add: eqvts) | 
| 
b6656b707a8d
Solved one of the strong-induction goals.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1632diff
changeset | 442 | apply (subst supp_finite_atom_set) | 
| 2104 
2205b572bc9b
Declare alpha_gen_eqvt as eqvt and change the proofs that used 'eqvts[symmetric]'
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2100diff
changeset | 443 | apply (simp add: supp_eqvt[symmetric] atom_eqvt[symmetric] Diff_eqvt[symmetric] finite_eqvt[symmetric] finite_supp True_eqvt) | 
| 1634 
b6656b707a8d
Solved one of the strong-induction goals.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1632diff
changeset | 444 | apply (simp add: eqvts) | 
| 
b6656b707a8d
Solved one of the strong-induction goals.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1632diff
changeset | 445 | apply (subst supp_perm_eq) | 
| 
b6656b707a8d
Solved one of the strong-induction goals.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1632diff
changeset | 446 | apply (subst supp_finite_atom_set) | 
| 2104 
2205b572bc9b
Declare alpha_gen_eqvt as eqvt and change the proofs that used 'eqvts[symmetric]'
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2100diff
changeset | 447 | apply (simp add: supp_eqvt[symmetric] atom_eqvt[symmetric] Diff_eqvt[symmetric] finite_eqvt[symmetric] finite_supp True_eqvt) | 
| 1634 
b6656b707a8d
Solved one of the strong-induction goals.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1632diff
changeset | 448 | apply assumption | 
| 
b6656b707a8d
Solved one of the strong-induction goals.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1632diff
changeset | 449 | apply (simp add: fresh_star_minus_perm) | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 450 | apply (rule a15) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 451 | apply simp | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 452 | apply(rotate_tac 1) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 453 | apply(erule_tac x="(pa + p)" in allE) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 454 | apply simp | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 455 | apply (simp add: eqvts eqvts_raw) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 456 | apply (rule at_set_avoiding2_atom) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 457 | apply (simp add: finite_supp) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 458 | apply (simp add: finite_supp) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 459 | apply (simp add: fresh_def) | 
| 1658 | 460 | apply (simp only: supp_abs eqvts) | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 461 | apply blast | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 462 | |
| 1645 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 463 | |
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 464 | (* GOAL3 a copy-and-paste of Goal2 with consts and variable names changed *) | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 465 | apply(subgoal_tac "\<exists>pa. ((pa \<bullet> (atom (p \<bullet> tvar))) \<sharp> g \<and> | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 466 |                        supp (Abs (p \<bullet> {atom tvar}) (p \<bullet> trm)) \<sharp>* pa)")
 | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 467 | apply clarify | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 468 | apply (simp only: perm) | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 469 | apply(rule_tac t="LAMT (p \<bullet> tvar) (p \<bullet> tkind) (p \<bullet> trm)" | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 470 | and s="LAMT (pa \<bullet> p \<bullet> tvar) (p \<bullet> tkind) (pa \<bullet> p \<bullet> trm)" in subst) | 
| 2044 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 471 | apply (simp add: eq_iff) | 
| 1645 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 472 | apply (rule_tac x="-pa" in exI) | 
| 1658 | 473 | apply (simp add: alphas eqvts eqvts_raw supp_abs fv_supp) | 
| 1645 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 474 |     apply (rule_tac t="supp (pa \<bullet> p \<bullet> trm) - {atom (pa \<bullet> p \<bullet> tvar)}"
 | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 475 |                 and s="pa \<bullet> (p \<bullet> supp trm - {p \<bullet> atom tvar})" in subst)
 | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 476 | apply (simp add: eqvts) | 
| 2104 
2205b572bc9b
Declare alpha_gen_eqvt as eqvt and change the proofs that used 'eqvts[symmetric]'
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2100diff
changeset | 477 | apply (simp add: supp_eqvt[symmetric] atom_eqvt[symmetric]) | 
| 1645 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 478 | apply (rule conjI) | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 479 | apply (rule supp_perm_eq) | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 480 | apply (simp add: eqvts) | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 481 | apply (subst supp_finite_atom_set) | 
| 2104 
2205b572bc9b
Declare alpha_gen_eqvt as eqvt and change the proofs that used 'eqvts[symmetric]'
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2100diff
changeset | 482 | apply (simp add: supp_eqvt[symmetric] atom_eqvt[symmetric] Diff_eqvt[symmetric] finite_eqvt[symmetric] finite_supp True_eqvt) | 
| 1645 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 483 | apply (simp add: eqvts) | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 484 | apply (subst supp_perm_eq) | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 485 | apply (subst supp_finite_atom_set) | 
| 2104 
2205b572bc9b
Declare alpha_gen_eqvt as eqvt and change the proofs that used 'eqvts[symmetric]'
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2100diff
changeset | 486 | apply (simp add: supp_eqvt[symmetric] atom_eqvt[symmetric] Diff_eqvt[symmetric] finite_eqvt[symmetric] finite_supp True_eqvt) | 
| 1645 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 487 | apply assumption | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 488 | apply (simp add: fresh_star_minus_perm) | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 489 | apply (rule a29) | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 490 | apply simp | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 491 | apply(rotate_tac 1) | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 492 | apply(erule_tac x="(pa + p)" in allE) | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 493 | apply simp | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 494 | apply (simp add: eqvts eqvts_raw) | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 495 | apply (rule at_set_avoiding2_atom) | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 496 | apply (simp add: finite_supp) | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 497 | apply (simp add: finite_supp) | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 498 | apply (simp add: fresh_def) | 
| 1658 | 499 | apply (simp only: supp_abs eqvts) | 
| 1645 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 500 | apply blast | 
| 
bde8da26093e
One more copy-and-paste in core-haskell.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1635diff
changeset | 501 | |
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 502 | (* GOAL4 a copy-and-paste *) | 
| 1701 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 503 | apply(subgoal_tac "\<exists>pa. ((pa \<bullet> (atom (p \<bullet> cvar))) \<sharp> g \<and> | 
| 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 504 |                        supp (Abs (p \<bullet> {atom cvar}) (p \<bullet> trm)) \<sharp>* pa)")
 | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 505 | apply clarify | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 506 | apply (simp only: perm) | 
| 1701 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 507 | apply(rule_tac t="LAMC (p \<bullet> cvar) (p \<bullet> ckind) (p \<bullet> trm)" | 
| 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 508 | and s="LAMC (pa \<bullet> p \<bullet> cvar) (p \<bullet> ckind) (pa \<bullet> p \<bullet> trm)" in subst) | 
| 2044 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 509 | apply (simp add: eq_iff) | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 510 | apply (rule_tac x="-pa" in exI) | 
| 1658 | 511 | apply (simp add: alphas eqvts eqvts_raw supp_abs fv_supp) | 
| 1701 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 512 |     apply (rule_tac t="supp (pa \<bullet> p \<bullet> trm) - {atom (pa \<bullet> p \<bullet> cvar)}"
 | 
| 
58abc09d6f9c
Updated strong induction to modified definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1700diff
changeset | 513 |                 and s="pa \<bullet> (p \<bullet> supp trm - {p \<bullet> atom cvar})" in subst)
 | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 514 | apply (simp add: eqvts) | 
| 2104 
2205b572bc9b
Declare alpha_gen_eqvt as eqvt and change the proofs that used 'eqvts[symmetric]'
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2100diff
changeset | 515 | apply (simp add: supp_eqvt[symmetric] atom_eqvt[symmetric]) | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 516 | apply (rule conjI) | 
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 517 | apply (rule supp_perm_eq) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 518 | apply (simp add: eqvts) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 519 | apply (subst supp_finite_atom_set) | 
| 2104 
2205b572bc9b
Declare alpha_gen_eqvt as eqvt and change the proofs that used 'eqvts[symmetric]'
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2100diff
changeset | 520 | apply (simp add: supp_eqvt[symmetric] atom_eqvt[symmetric] Diff_eqvt[symmetric] finite_eqvt[symmetric] finite_supp True_eqvt) | 
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 521 | apply (simp add: eqvts) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 522 | apply (subst supp_perm_eq) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 523 | apply (subst supp_finite_atom_set) | 
| 2104 
2205b572bc9b
Declare alpha_gen_eqvt as eqvt and change the proofs that used 'eqvts[symmetric]'
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2100diff
changeset | 524 | apply (simp add: supp_eqvt[symmetric] atom_eqvt[symmetric] Diff_eqvt[symmetric] finite_eqvt[symmetric] finite_supp True_eqvt) | 
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 525 | apply assumption | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 526 | apply (simp add: fresh_star_minus_perm) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 527 | apply (rule a30) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 528 | apply simp | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 529 | apply(rotate_tac 1) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 530 | apply(erule_tac x="(pa + p)" in allE) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 531 | apply simp | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 532 | apply (simp add: eqvts eqvts_raw) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 533 | apply (rule at_set_avoiding2_atom) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 534 | apply (simp add: finite_supp) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 535 | apply (simp add: finite_supp) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 536 | apply (simp add: fresh_def) | 
| 1658 | 537 | apply (simp only: supp_abs eqvts) | 
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 538 | apply blast | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 539 | |
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 540 | |
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 541 | (* GOAL5 a copy-and-paste *) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 542 | apply(subgoal_tac "\<exists>pa. ((pa \<bullet> (atom (p \<bullet> var))) \<sharp> g \<and> | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 543 |                        supp (Abs (p \<bullet> {atom var}) (p \<bullet> trm)) \<sharp>* pa)")
 | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 544 | apply clarify | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 545 | apply (simp only: perm) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 546 | apply(rule_tac t="Lam (p \<bullet> var) (p \<bullet> ty) (p \<bullet> trm)" | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 547 | and s="Lam (pa \<bullet> p \<bullet> var) (p \<bullet> ty) (pa \<bullet> p \<bullet> trm)" in subst) | 
| 2044 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 548 | apply (simp add: eq_iff) | 
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 549 | apply (rule_tac x="-pa" in exI) | 
| 1658 | 550 | apply (simp add: alphas eqvts eqvts_raw supp_abs fv_supp) | 
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 551 |     apply (rule_tac t="supp (pa \<bullet> p \<bullet> trm) - {atom (pa \<bullet> p \<bullet> var)}"
 | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 552 |                 and s="pa \<bullet> (p \<bullet> supp trm - {p \<bullet> atom var})" in subst)
 | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 553 | apply (simp add: eqvts) | 
| 2104 
2205b572bc9b
Declare alpha_gen_eqvt as eqvt and change the proofs that used 'eqvts[symmetric]'
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2100diff
changeset | 554 | apply (simp add: supp_eqvt[symmetric] atom_eqvt[symmetric]) | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 555 | apply (rule conjI) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 556 | apply (rule supp_perm_eq) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 557 | apply (simp add: eqvts) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 558 | apply (subst supp_finite_atom_set) | 
| 2104 
2205b572bc9b
Declare alpha_gen_eqvt as eqvt and change the proofs that used 'eqvts[symmetric]'
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2100diff
changeset | 559 | apply (simp add: supp_eqvt[symmetric] atom_eqvt[symmetric] Diff_eqvt[symmetric] finite_eqvt[symmetric] finite_supp True_eqvt) | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 560 | apply (simp add: eqvts) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 561 | apply (subst supp_perm_eq) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 562 | apply (subst supp_finite_atom_set) | 
| 2104 
2205b572bc9b
Declare alpha_gen_eqvt as eqvt and change the proofs that used 'eqvts[symmetric]'
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2100diff
changeset | 563 | apply (simp add: supp_eqvt[symmetric] atom_eqvt[symmetric] Diff_eqvt[symmetric] finite_eqvt[symmetric] finite_supp True_eqvt) | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 564 | apply assumption | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 565 | apply (simp add: fresh_star_minus_perm) | 
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 566 | apply (rule a32) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 567 | apply simp | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 568 | apply(rotate_tac 1) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 569 | apply(erule_tac x="(pa + p)" in allE) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 570 | apply simp | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 571 | apply (simp add: eqvts eqvts_raw) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 572 | apply (rule at_set_avoiding2_atom) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 573 | apply (simp add: finite_supp) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 574 | apply (simp add: finite_supp) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 575 | apply (simp add: fresh_def) | 
| 1658 | 576 | apply (simp only: supp_abs eqvts) | 
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 577 | apply blast | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 578 | |
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 579 | |
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 580 | (* GOAL6 a copy-and-paste *) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 581 | apply(subgoal_tac "\<exists>pa. ((pa \<bullet> (atom (p \<bullet> var))) \<sharp> g \<and> | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 582 |                        supp (Abs (p \<bullet> {atom var}) (p \<bullet> trm2)) \<sharp>* pa)")
 | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 583 | apply clarify | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 584 | apply (simp only: perm) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 585 | apply(rule_tac t="Let (p \<bullet> var) (p \<bullet> ty) (p \<bullet> trm1) (p \<bullet> trm2)" | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 586 | and s="Let (pa \<bullet> p \<bullet> var) (p \<bullet> ty) (p \<bullet> trm1) (pa \<bullet> p \<bullet> trm2)" in subst) | 
| 2044 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 587 | apply (simp add: eq_iff) | 
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 588 | apply (rule_tac x="-pa" in exI) | 
| 1658 | 589 | apply (simp add: alphas eqvts eqvts_raw supp_abs fv_supp) | 
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 590 |     apply (rule_tac t="supp (pa \<bullet> p \<bullet> trm2) - {atom (pa \<bullet> p \<bullet> var)}"
 | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 591 |                 and s="pa \<bullet> (p \<bullet> supp trm2 - {p \<bullet> atom var})" in subst)
 | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 592 | apply (simp add: eqvts) | 
| 2104 
2205b572bc9b
Declare alpha_gen_eqvt as eqvt and change the proofs that used 'eqvts[symmetric]'
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2100diff
changeset | 593 | apply (simp add: supp_eqvt[symmetric] atom_eqvt[symmetric]) | 
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 594 | apply (rule conjI) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 595 | apply (rule supp_perm_eq) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 596 | apply (simp add: eqvts) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 597 | apply (subst supp_finite_atom_set) | 
| 2104 
2205b572bc9b
Declare alpha_gen_eqvt as eqvt and change the proofs that used 'eqvts[symmetric]'
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2100diff
changeset | 598 | apply (simp add: supp_eqvt[symmetric] atom_eqvt[symmetric] Diff_eqvt[symmetric] finite_eqvt[symmetric] finite_supp True_eqvt) | 
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 599 | apply (simp add: eqvts) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 600 | apply (subst supp_perm_eq) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 601 | apply (subst supp_finite_atom_set) | 
| 2104 
2205b572bc9b
Declare alpha_gen_eqvt as eqvt and change the proofs that used 'eqvts[symmetric]'
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2100diff
changeset | 602 | apply (simp add: supp_eqvt[symmetric] atom_eqvt[symmetric] Diff_eqvt[symmetric] finite_eqvt[symmetric] finite_supp True_eqvt) | 
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 603 | apply assumption | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 604 | apply (simp add: fresh_star_minus_perm) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 605 | apply (rule a34) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 606 | apply simp | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 607 | apply simp | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 608 | apply(rotate_tac 2) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 609 | apply(erule_tac x="(pa + p)" in allE) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 610 | apply simp | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 611 | apply (simp add: eqvts eqvts_raw) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 612 | apply (rule at_set_avoiding2_atom) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 613 | apply (simp add: finite_supp) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 614 | apply (simp add: finite_supp) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 615 | apply (simp add: fresh_def) | 
| 1658 | 616 | apply (simp only: supp_abs eqvts) | 
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 617 | apply blast | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 618 | |
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 619 | (* MAIN ACons Goal *) | 
| 1695 | 620 | apply(subgoal_tac "\<exists>pa. ((pa \<bullet> (set (bv (p \<bullet> pat)))) \<sharp>* h \<and> | 
| 621 | supp (Abs_lst (p \<bullet> (bv pat)) (p \<bullet> trm)) \<sharp>* pa)") | |
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 622 | apply clarify | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 623 | apply (simp only: perm eqvts) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 624 | apply (subst ACons_subst) | 
| 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 625 | apply assumption | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 626 | apply (rule a38) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 627 | apply simp | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 628 | apply(rotate_tac 1) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 629 | apply(erule_tac x="(pa + p)" in allE) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 630 | apply simp | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 631 | apply simp | 
| 1647 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 632 | apply (simp add: perm_bv2[symmetric]) | 
| 1635 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 633 | apply (simp add: eqvts eqvts_raw) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 634 | apply (rule at_set_avoiding2) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 635 | apply (simp add: fin_bv) | 
| 
8b4595cb5524
Further in the strong induction proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1634diff
changeset | 636 | apply (simp add: finite_supp) | 
| 1658 | 637 | apply (simp add: supp_abs) | 
| 1646 
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1645diff
changeset | 638 | apply (simp add: finite_supp) | 
| 1658 | 639 | apply (simp add: fresh_star_def fresh_def supp_abs eqvts) | 
| 1647 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 640 | done | 
| 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 641 | then have b: "P1 a (0 \<bullet> tkind)" and "P2 b (0 \<bullet> ckind)" "P3 c (0 \<bullet> ty)" and "P4 d (0 \<bullet> ty_lst)" and "P5 e (0 \<bullet> co)" and "P6 f (0 \<bullet> co_lst)" and "P7 g (0 \<bullet> trm)" and "P8 h (0 \<bullet> assoc_lst)" by (blast+) | 
| 2142 
c39d4fe31100
moved the exporting part into the parser (this is still a hack); re-added CoreHaskell again to the examples - there seems to be a problem with the variable name pat
 Christian Urban <urbanc@in.tum.de> parents: 
2120diff
changeset | 642 | moreover have "P9 i (permute_bv 0 (0 \<bullet> pt))" and "P10 j (permute_bv_vs 0 (0 \<bullet> vars))" and "P11 k (permute_bv_tvs 0 (0 \<bullet> tvars))" and "P12 l (permute_bv_cvs 0 (0 \<bullet> cvars))" using a1 a2 a3 a4 by (blast+) | 
| 1647 
032649a694d2
Only ACons_subst left to show.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1646diff
changeset | 643 | ultimately show ?thesis by (simp_all add: permute_bv_zero1 permute_bv_zero2) | 
| 1631 
e94bfef17bb8
stating the strong induction; further.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1630diff
changeset | 644 | qed | 
| 1615 
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1609diff
changeset | 645 | |
| 1868 
26c0c35641cb
equivariance for alpha_raw in CoreHaskell is automatically derived
 Christian Urban <urbanc@in.tum.de> parents: 
1867diff
changeset | 646 | section {* test about equivariance for alpha *}
 | 
| 
26c0c35641cb
equivariance for alpha_raw in CoreHaskell is automatically derived
 Christian Urban <urbanc@in.tum.de> parents: 
1867diff
changeset | 647 | |
| 2044 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 648 | (* this should not be an equivariance rule *) | 
| 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 649 | (* for the moment, we force it to be *) | 
| 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 650 | |
| 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 651 | (*declare permute_pure[eqvt]*) | 
| 2104 
2205b572bc9b
Declare alpha_gen_eqvt as eqvt and change the proofs that used 'eqvts[symmetric]'
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2100diff
changeset | 652 | (*setup {* Context.theory_map (Nominal_ThmDecls.add_thm @{thm "permute_pure"}) *} *)
 | 
| 2044 
695c1ed61879
moved CoreHaskell to NewParser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1958diff
changeset | 653 | |
| 1868 
26c0c35641cb
equivariance for alpha_raw in CoreHaskell is automatically derived
 Christian Urban <urbanc@in.tum.de> parents: 
1867diff
changeset | 654 | thm eqvts | 
| 
26c0c35641cb
equivariance for alpha_raw in CoreHaskell is automatically derived
 Christian Urban <urbanc@in.tum.de> parents: 
1867diff
changeset | 655 | thm eqvts_raw | 
| 
26c0c35641cb
equivariance for alpha_raw in CoreHaskell is automatically derived
 Christian Urban <urbanc@in.tum.de> parents: 
1867diff
changeset | 656 | |
| 
26c0c35641cb
equivariance for alpha_raw in CoreHaskell is automatically derived
 Christian Urban <urbanc@in.tum.de> parents: 
1867diff
changeset | 657 | |
| 1601 | 658 | end | 
| 659 |