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