1601
+ − 1
theory ExCoreHaskell
+ − 2
imports "Parser"
+ − 3
begin
+ − 4
+ − 5
(* core haskell *)
+ − 6
1624
+ − 7
ML {* val _ = recursive := false *}
1656
+ − 8
1601
+ − 9
atom_decl var
+ − 10
atom_decl tvar
+ − 11
1606
+ − 12
(* there are types, coercion types and regular types list-data-structure *)
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 13
1601
+ − 14
nominal_datatype tkind =
+ − 15
KStar
+ − 16
| KFun "tkind" "tkind"
+ − 17
and ckind =
+ − 18
CKEq "ty" "ty"
+ − 19
and ty =
+ − 20
TVar "tvar"
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 21
| TC "char"
1601
+ − 22
| TApp "ty" "ty"
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 23
| TFun "char" "ty_lst"
1601
+ − 24
| TAll tv::"tvar" "tkind" T::"ty" bind tv in T
+ − 25
| TEq "ty" "ty" "ty"
1606
+ − 26
and ty_lst =
+ − 27
TsNil
+ − 28
| TsCons "ty" "ty_lst"
1601
+ − 29
and co =
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 30
CC "char"
1601
+ − 31
| CApp "co" "co"
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 32
| CFun "char" "co_lst"
1601
+ − 33
| CAll tv::"tvar" "ckind" C::"co" bind tv in C
+ − 34
| CEq "co" "co" "co"
+ − 35
| CSym "co"
+ − 36
| CCir "co" "co"
+ − 37
| CLeft "co"
+ − 38
| CRight "co"
+ − 39
| CSim "co"
+ − 40
| CRightc "co"
+ − 41
| CLeftc "co"
+ − 42
| CCoe "co" "co"
1606
+ − 43
and co_lst =
+ − 44
CsNil
+ − 45
| 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>
diff
changeset
+ − 46
and trm =
1601
+ − 47
Var "var"
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 48
| C "char"
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 49
| LAMT tv::"tvar" "tkind" t::"trm" bind tv in t
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 50
| LAMC tv::"tvar" "ckind" t::"trm" bind tv in t
1601
+ − 51
| APP "trm" "ty"
+ − 52
| Lam v::"var" "ty" t::"trm" bind v in t
+ − 53
| App "trm" "trm"
+ − 54
| Let x::"var" "ty" "trm" t::"trm" bind x in t
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 55
| Case "trm" "assoc_lst"
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 56
| 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>
diff
changeset
+ − 57
and assoc_lst =
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 58
ANil
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 59
| ACons p::"pat" t::"trm" "assoc_lst" bind "bv p" in t
1601
+ − 60
and pat =
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 61
K "char" "tvtk_lst" "tvck_lst" "vt_lst"
1621
+ − 62
and vt_lst =
+ − 63
VTNil
+ − 64
| VTCons "var" "ty" "vt_lst"
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 65
and tvtk_lst =
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 66
TVTKNil
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 67
| TVTKCons "tvar" "tkind" "tvtk_lst"
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 68
and tvck_lst =
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 69
TVCKNil
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 70
| TVCKCons "tvar" "ckind" "tvck_lst"
1601
+ − 71
binder
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 72
bv :: "pat \<Rightarrow> atom set"
1621
+ − 73
and bv_vt :: "vt_lst \<Rightarrow> atom set"
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 74
and bv_tvtk :: "tvtk_lst \<Rightarrow> atom set"
1621
+ − 75
and bv_tvck :: "tvck_lst \<Rightarrow> atom set"
1601
+ − 76
where
1621
+ − 77
"bv (K s tvts tvcs vs) = (bv_tvtk tvts) \<union> (bv_tvck tvcs) \<union> (bv_vt vs)"
+ − 78
| "bv_vt VTNil = {}"
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 79
| "bv_vt (VTCons v k t) = {atom v} \<union> bv_vt t"
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 80
| "bv_tvtk TVTKNil = {}"
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 81
| "bv_tvtk (TVTKCons v k t) = {atom v} \<union> bv_tvtk t"
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 82
| "bv_tvck TVCKNil = {}"
1621
+ − 83
| "bv_tvck (TVCKCons v k t) = {atom v} \<union> bv_tvck t"
+ − 84
1626
+ − 85
lemmas fv_supp=tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vt_lst_tvtk_lst_tvck_lst.supp(1-9,11,13,15)
+ − 86
lemmas supp=tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vt_lst_tvtk_lst_tvck_lst.fv[simplified fv_supp]
1634
+ − 87
lemmas perm=tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vt_lst_tvtk_lst_tvck_lst.perm
1635
+ − 88
lemmas eq_iff=tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vt_lst_tvtk_lst_tvck_lst.eq_iff
+ − 89
lemmas inducts=tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vt_lst_tvtk_lst_tvck_lst.inducts
1634
+ − 90
1647
+ − 91
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_vt_lst_raw_alpha_tvtk_lst_raw_alpha_tvck_lst_raw_alpha_bv_raw_alpha_bv_vt_raw_alpha_bv_tvtk_raw_alpha_bv_tvck_raw.inducts
+ − 92
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_vt_lst_raw_alpha_tvtk_lst_raw_alpha_tvck_lst_raw_alpha_bv_raw_alpha_bv_vt_raw_alpha_bv_tvtk_raw_alpha_bv_tvck_raw.intros
+ − 93
1634
+ − 94
lemma fresh_star_minus_perm: "as \<sharp>* - p = as \<sharp>* (p :: perm)"
+ − 95
unfolding fresh_star_def Ball_def
+ − 96
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>
diff
changeset
+ − 97
1647
+ − 98
primrec permute_bv_vt_raw
+ − 99
where "permute_bv_vt_raw p VTNil_raw = VTNil_raw"
+ − 100
| "permute_bv_vt_raw p (VTCons_raw v t l) = VTCons_raw (p \<bullet> v) t (permute_bv_vt_raw p l)"
+ − 101
primrec permute_bv_tvck_raw
+ − 102
where "permute_bv_tvck_raw p TVCKNil_raw = TVCKNil_raw"
+ − 103
| "permute_bv_tvck_raw p (TVCKCons_raw v t l) = TVCKCons_raw (p \<bullet> v) t (permute_bv_tvck_raw p l)"
+ − 104
primrec permute_bv_tvtk_raw
+ − 105
where "permute_bv_tvtk_raw p TVTKNil_raw = TVTKNil_raw"
+ − 106
| "permute_bv_tvtk_raw p (TVTKCons_raw v t l) = TVTKCons_raw (p \<bullet> v) t (permute_bv_tvtk_raw p l)"
+ − 107
primrec permute_bv_raw
+ − 108
where "permute_bv_raw p (K_raw c l1 l2 l3) =
+ − 109
K_raw c (permute_bv_tvtk_raw p l1) (permute_bv_tvck_raw p l2) (permute_bv_vt_raw p l3)"
+ − 110
+ − 111
quotient_definition "permute_bv_vt :: perm \<Rightarrow> vt_lst \<Rightarrow> vt_lst"
+ − 112
is "permute_bv_vt_raw"
+ − 113
quotient_definition "permute_bv_tvck :: perm \<Rightarrow> tvck_lst \<Rightarrow> tvck_lst"
+ − 114
is "permute_bv_tvck_raw"
+ − 115
quotient_definition "permute_bv_tvtk :: perm \<Rightarrow> tvtk_lst \<Rightarrow> tvtk_lst"
+ − 116
is "permute_bv_tvtk_raw"
+ − 117
quotient_definition "permute_bv :: perm \<Rightarrow> pat \<Rightarrow> pat"
+ − 118
is "permute_bv_raw"
+ − 119
+ − 120
lemma rsp_pre:
+ − 121
"alpha_tvtk_lst_raw d a \<Longrightarrow> alpha_tvtk_lst_raw (permute_bv_tvtk_raw x d) (permute_bv_tvtk_raw x a)"
+ − 122
"alpha_tvck_lst_raw e b \<Longrightarrow> alpha_tvck_lst_raw (permute_bv_tvck_raw x e) (permute_bv_tvck_raw x b)"
+ − 123
"alpha_vt_lst_raw f c \<Longrightarrow> alpha_vt_lst_raw (permute_bv_vt_raw x f) (permute_bv_vt_raw x c)"
+ − 124
apply (erule_tac [!] alpha_inducts)
+ − 125
apply simp_all
+ − 126
apply (rule_tac [!] alpha_intros)
+ − 127
apply simp_all
+ − 128
done
1646
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 129
1647
+ − 130
lemma [quot_respect]:
+ − 131
"(op = ===> alpha_pat_raw ===> alpha_pat_raw) permute_bv_raw permute_bv_raw"
+ − 132
"(op = ===> alpha_tvtk_lst_raw ===> alpha_tvtk_lst_raw) permute_bv_tvtk_raw permute_bv_tvtk_raw"
+ − 133
"(op = ===> alpha_tvck_lst_raw ===> alpha_tvck_lst_raw) permute_bv_tvck_raw permute_bv_tvck_raw"
+ − 134
"(op = ===> alpha_vt_lst_raw ===> alpha_vt_lst_raw) permute_bv_vt_raw permute_bv_vt_raw"
+ − 135
apply (simp_all add: rsp_pre)
+ − 136
apply clarify
+ − 137
apply (erule_tac alpha_inducts)
+ − 138
apply (simp_all)
+ − 139
apply (rule alpha_intros)
+ − 140
apply (simp_all add: rsp_pre)
+ − 141
done
+ − 142
+ − 143
lemmas permute_bv[simp] = permute_bv_raw.simps[quot_lifted]
+ − 144
permute_bv_vt_raw.simps[quot_lifted]
+ − 145
permute_bv_tvck_raw.simps[quot_lifted]
+ − 146
permute_bv_tvtk_raw.simps[quot_lifted]
1646
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 147
1647
+ − 148
lemma perm_bv1:
+ − 149
"p \<bullet> bv_tvck b = bv_tvck (permute_bv_tvck p b)"
+ − 150
"p \<bullet> bv_tvtk c = bv_tvtk (permute_bv_tvtk p c)"
+ − 151
"p \<bullet> bv_vt d = bv_vt (permute_bv_vt p d)"
+ − 152
apply(induct b rule: inducts(12))
+ − 153
apply(rule TrueI)
+ − 154
apply(simp_all add:permute_bv eqvts)
+ − 155
apply(induct c rule: inducts(11))
+ − 156
apply(rule TrueI)
+ − 157
apply(simp_all add:permute_bv eqvts)
+ − 158
apply(induct d rule: inducts(10))
+ − 159
apply(rule TrueI)
+ − 160
apply(simp_all add:permute_bv eqvts)
+ − 161
done
+ − 162
+ − 163
lemma perm_bv2:
+ − 164
"p \<bullet> bv l = bv (permute_bv p l)"
+ − 165
apply(induct l rule: inducts(9))
+ − 166
apply(rule TrueI)
+ − 167
apply(simp_all add:permute_bv)
+ − 168
apply(simp add: perm_bv1[symmetric])
+ − 169
apply(simp add: eqvts)
+ − 170
done
1646
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 171
1648
+ − 172
lemma alpha_perm_bn1:
+ − 173
" alpha_bv_tvtk tvtk_lst (permute_bv_tvtk q tvtk_lst)"
+ − 174
"alpha_bv_tvck tvck_lst (permute_bv_tvck q tvck_lst)"
+ − 175
"alpha_bv_vt vt_lst (permute_bv_vt q vt_lst)"
+ − 176
apply(induct tvtk_lst rule: inducts(11))
+ − 177
apply(simp_all add:permute_bv eqvts eq_iff)
+ − 178
apply(induct tvck_lst rule: inducts(12))
+ − 179
apply(simp_all add:permute_bv eqvts eq_iff)
+ − 180
apply(induct vt_lst rule: inducts(10))
+ − 181
apply(simp_all add:permute_bv eqvts eq_iff)
+ − 182
done
+ − 183
+ − 184
lemma alpha_perm_bn:
+ − 185
"alpha_bv pat (permute_bv q pat)"
+ − 186
apply(induct pat rule: inducts(9))
+ − 187
apply(simp_all add:permute_bv eqvts eq_iff alpha_perm_bn1)
+ − 188
done
+ − 189
1646
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 190
lemma ACons_subst:
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 191
"supp (Abs (bv pat) trm) \<sharp>* q \<Longrightarrow> (ACons pat trm al) = ACons (permute_bv q pat) (q \<bullet> trm) al"
1648
+ − 192
apply (simp only: eq_iff)
+ − 193
apply (rule_tac x="q" in exI)
+ − 194
apply (simp add: alphas)
+ − 195
apply (simp add: perm_bv2[symmetric])
+ − 196
apply (simp add: eqvts[symmetric])
+ − 197
apply (simp add: supp_Abs)
+ − 198
apply (simp add: fv_supp)
+ − 199
apply (simp add: alpha_perm_bn)
+ − 200
apply (rule supp_perm_eq[symmetric])
+ − 201
apply (subst supp_finite_atom_set)
+ − 202
apply (rule finite_Diff)
+ − 203
apply (simp add: finite_supp)
+ − 204
apply (assumption)
+ − 205
done
1646
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 206
1647
+ − 207
lemma permute_bv_zero1:
+ − 208
"permute_bv_tvck 0 b = b"
+ − 209
"permute_bv_tvtk 0 c = c"
+ − 210
"permute_bv_vt 0 d = d"
+ − 211
apply(induct b rule: inducts(12))
+ − 212
apply(rule TrueI)
+ − 213
apply(simp_all add:permute_bv eqvts)
+ − 214
apply(induct c rule: inducts(11))
+ − 215
apply(rule TrueI)
+ − 216
apply(simp_all add:permute_bv eqvts)
+ − 217
apply(induct d rule: inducts(10))
+ − 218
apply(rule TrueI)
+ − 219
apply(simp_all add:permute_bv eqvts)
+ − 220
done
1646
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 221
1647
+ − 222
lemma permute_bv_zero2:
+ − 223
"permute_bv 0 a = a"
+ − 224
apply(induct a rule: inducts(9))
+ − 225
apply(rule TrueI)
+ − 226
apply(simp_all add:permute_bv eqvts permute_bv_zero1)
+ − 227
done
1646
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 228
1635
+ − 229
lemma fv_alpha1: "fv_bv_tvtk x \<sharp>* pa \<Longrightarrow> alpha_bv_tvtk (pa \<bullet> x) x"
+ − 230
apply (induct x rule: inducts(11))
+ − 231
apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *})
+ − 232
apply (simp_all add: eq_iff fresh_star_union)
+ − 233
apply (subst supp_perm_eq)
+ − 234
apply (simp_all add: fv_supp)
+ − 235
done
+ − 236
+ − 237
lemma fv_alpha2: "fv_bv_tvck x \<sharp>* pa \<Longrightarrow> alpha_bv_tvck (pa \<bullet> x) x"
+ − 238
apply (induct x rule: inducts(12))
+ − 239
apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *})
+ − 240
apply (simp_all add: eq_iff fresh_star_union)
+ − 241
apply (subst supp_perm_eq)
+ − 242
apply (simp_all add: fv_supp)
+ − 243
done
+ − 244
+ − 245
lemma fv_alpha3: "fv_bv_vt x \<sharp>* pa \<Longrightarrow> alpha_bv_vt (pa \<bullet> x) x"
+ − 246
apply (induct x rule: inducts(10))
+ − 247
apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *})
+ − 248
apply (simp_all add: eq_iff fresh_star_union)
+ − 249
apply (subst supp_perm_eq)
+ − 250
apply (simp_all add: fv_supp)
+ − 251
done
+ − 252
+ − 253
lemma fv_alpha: "fv_bv x \<sharp>* pa \<Longrightarrow> alpha_bv (pa \<bullet> x) x"
+ − 254
apply (induct x rule: inducts(9))
+ − 255
apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *})
+ − 256
apply (simp_all add: eq_iff fresh_star_union)
+ − 257
apply (simp add: fv_alpha1 fv_alpha2 fv_alpha3)
+ − 258
apply (simp add: eqvts)
+ − 259
done
+ − 260
+ − 261
lemma fin1: "finite (fv_bv_tvtk x)"
+ − 262
apply (induct x rule: inducts(11))
+ − 263
apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *})
+ − 264
apply (simp_all add: fv_supp finite_supp)
+ − 265
done
+ − 266
+ − 267
lemma fin2: "finite (fv_bv_tvck x)"
+ − 268
apply (induct x rule: inducts(12))
+ − 269
apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *})
+ − 270
apply (simp_all add: fv_supp finite_supp)
+ − 271
done
+ − 272
+ − 273
lemma fin3: "finite (fv_bv_vt x)"
+ − 274
apply (induct x rule: inducts(10))
+ − 275
apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *})
+ − 276
apply (simp_all add: fv_supp finite_supp)
+ − 277
done
+ − 278
+ − 279
lemma fin_fv_bv: "finite (fv_bv x)"
+ − 280
apply (induct x rule: inducts(9))
+ − 281
apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *})
+ − 282
apply (simp add: fin1 fin2 fin3)
+ − 283
done
+ − 284
+ − 285
lemma finb1: "finite (bv_tvtk x)"
+ − 286
apply (induct x rule: inducts(11))
+ − 287
apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *})
+ − 288
apply (simp_all add: fv_supp finite_supp)
+ − 289
done
+ − 290
+ − 291
lemma finb2: "finite (bv_tvck x)"
+ − 292
apply (induct x rule: inducts(12))
+ − 293
apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *})
+ − 294
apply (simp_all add: fv_supp finite_supp)
+ − 295
done
+ − 296
+ − 297
lemma finb3: "finite (bv_vt x)"
+ − 298
apply (induct x rule: inducts(10))
+ − 299
apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *})
+ − 300
apply (simp_all add: fv_supp finite_supp)
+ − 301
done
+ − 302
+ − 303
lemma fin_bv: "finite (bv x)"
+ − 304
apply (induct x rule: inducts(9))
+ − 305
apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *})
+ − 306
apply (simp add: finb1 finb2 finb3)
+ − 307
done
+ − 308
+ − 309
lemma "bv x \<sharp>* fv_bv x"
+ − 310
apply (induct x rule: inducts(9))
+ − 311
apply (tactic {* ALLGOALS (TRY o rtac @{thm TrueI}) *})
+ − 312
apply (simp)
+ − 313
oops
+ − 314
1630
+ − 315
lemma
1631
+ − 316
assumes a01: "\<And>b. P1 b KStar"
+ − 317
and a02: "\<And>tkind1 tkind2 b. \<lbrakk>\<And>c. P1 c tkind1; \<And>c. P1 c tkind2\<rbrakk> \<Longrightarrow> P1 b (KFun tkind1 tkind2)"
+ − 318
and a03: "\<And>ty1 ty2 b. \<lbrakk>\<And>c. P3 c ty1; \<And>c. P3 c ty2\<rbrakk> \<Longrightarrow> P2 b (CKEq ty1 ty2)"
+ − 319
and a04: "\<And>tvar b. P3 b (TVar tvar)"
+ − 320
and a05: "\<And>char b. P3 b (TC char)"
+ − 321
and a06: "\<And>ty1 ty2 b. \<lbrakk>\<And>c. P3 c ty1; \<And>c. P3 c ty2\<rbrakk> \<Longrightarrow> P3 b (TApp ty1 ty2)"
+ − 322
and a07: "\<And>char ty_lst b. \<lbrakk>\<And>c. P4 c ty_lst\<rbrakk> \<Longrightarrow> P3 b (TFun char ty_lst)"
1632
+ − 323
and a08: "\<And>tvar tkind ty b. \<lbrakk>\<And>c. P1 c tkind; \<And>c. P3 c ty; atom tvar \<sharp> b\<rbrakk>
+ − 324
\<Longrightarrow> P3 b (TAll tvar tkind ty)"
1631
+ − 325
and a09: "\<And>ty1 ty2 ty3 b. \<lbrakk>\<And>c. P3 c ty1; \<And>c. P3 c ty2; \<And>c. P3 c ty3\<rbrakk> \<Longrightarrow> P3 b (TEq ty1 ty2 ty3)"
+ − 326
and a10: "\<And>b. P4 b TsNil"
+ − 327
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)"
+ − 328
and a12: "\<And>char b. P5 b (CC char)"
+ − 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)"
+ − 330
and a14: "\<And>char co_lst b. \<lbrakk>\<And>c. P6 c co_lst\<rbrakk> \<Longrightarrow> P5 b (CFun char co_lst)"
1632
+ − 331
and a15: "\<And>tvar ckind co b. \<lbrakk>\<And>c. P2 c ckind; \<And>c. P5 c co; atom tvar \<sharp> b\<rbrakk>
+ − 332
\<Longrightarrow> P5 b (CAll tvar ckind co)"
1631
+ − 333
and a16: "\<And>co1 co2 co3 b. \<lbrakk>\<And>c. P5 c co1; \<And>c. P5 c co2; \<And>c. P5 c co3\<rbrakk> \<Longrightarrow> P5 b (CEq co1 co2 co3)"
+ − 334
and a17: "\<And>co b. \<lbrakk>\<And>c. P5 c co\<rbrakk> \<Longrightarrow> P5 b (CSym co)"
+ − 335
and a18: "\<And>co1 co2 b. \<lbrakk>\<And>c. P5 c co1; \<And>c. P5 c co2\<rbrakk> \<Longrightarrow> P5 b (CCir co1 co2)"
+ − 336
and a19: "\<And>co b. \<lbrakk>\<And>c. P5 c co\<rbrakk> \<Longrightarrow> P5 b (CLeft co)"
+ − 337
and a20: "\<And>co b. \<lbrakk>\<And>c. P5 c co\<rbrakk> \<Longrightarrow> P5 b (CRight co)"
+ − 338
and a21: "\<And>co b. \<lbrakk>\<And>c. P5 c co\<rbrakk> \<Longrightarrow> P5 b (CSim co)"
+ − 339
and a22: "\<And>co b. \<lbrakk>\<And>c. P5 c co\<rbrakk> \<Longrightarrow> P5 b (CRightc co)"
+ − 340
and a23: "\<And>co b. \<lbrakk>\<And>c. P5 c co\<rbrakk> \<Longrightarrow> P5 b (CLeftc co)"
+ − 341
and a24: "\<And>co1 co2 b. \<lbrakk>\<And>c. P5 c co1; \<And>c. P5 c co2\<rbrakk> \<Longrightarrow> P5 b (CCoe co1 co2)"
+ − 342
and a25: "\<And>b. P6 b CsNil"
+ − 343
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)"
+ − 344
and a27: "\<And>var b. P7 b (Var var)"
+ − 345
and a28: "\<And>char b. P7 b (C char)"
1632
+ − 346
and a29: "\<And>tvar tkind trm b. \<lbrakk>\<And>c. P1 c tkind; \<And>c. P7 c trm; atom tvar \<sharp> b\<rbrakk>
+ − 347
\<Longrightarrow> P7 b (LAMT tvar tkind trm)"
+ − 348
and a30: "\<And>tvar ckind trm b. \<lbrakk>\<And>c. P2 c ckind; \<And>c. P7 c trm; atom tvar \<sharp> b\<rbrakk>
+ − 349
\<Longrightarrow> P7 b (LAMC tvar ckind trm)"
1631
+ − 350
and a31: "\<And>trm ty b. \<lbrakk>\<And>c. P7 c trm; \<And>c. P3 c ty\<rbrakk> \<Longrightarrow> P7 b (APP trm ty)"
1632
+ − 351
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
+ − 352
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
+ − 353
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
+ − 354
\<Longrightarrow> P7 b (Let var ty trm1 trm2)"
+ − 355
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)"
+ − 356
and a36: "\<And>trm ty b. \<lbrakk>\<And>c. P7 c trm; \<And>c. P3 c ty\<rbrakk> \<Longrightarrow> P7 b (Cast trm ty)"
+ − 357
and a37: "\<And>b. P8 b ANil"
1632
+ − 358
and a38: "\<And>pat trm assoc_lst b. \<lbrakk>\<And>c. P9 c pat; \<And>c. P7 c trm; \<And>c. P8 c assoc_lst; bv(pat) \<sharp>* b\<rbrakk>
1631
+ − 359
\<Longrightarrow> P8 b (ACons pat trm assoc_lst)"
+ − 360
and a39: "\<And>char tvtk_lst tvck_lst vt_lst b. \<lbrakk>\<And>c. P11 c tvtk_lst; \<And>c. P12 c tvck_lst; \<And>c. P10 c vt_lst\<rbrakk>
+ − 361
\<Longrightarrow> P9 b (K char tvtk_lst tvck_lst vt_lst)"
+ − 362
and a40: "\<And>b. P10 b VTNil"
+ − 363
and a41: "\<And>var ty vt_lst b. \<lbrakk>\<And>c. P3 c ty; \<And>c. P10 c vt_lst\<rbrakk> \<Longrightarrow> P10 b (VTCons var ty vt_lst)"
+ − 364
and a42: "\<And>b. P11 b TVTKNil"
+ − 365
and a43: "\<And>tvar tkind tvtk_lst b. \<lbrakk>\<And>c. P1 c tkind; \<And>c. P11 c tvtk_lst\<rbrakk>
+ − 366
\<Longrightarrow> P11 b (TVTKCons tvar tkind tvtk_lst)"
+ − 367
and a44: "\<And>b. P12 b TVCKNil"
+ − 368
and a45: "\<And>tvar ckind tvck_lst b. \<lbrakk>\<And>c. P2 c ckind; \<And>c. P12 c tvck_lst\<rbrakk>
+ − 369
\<Longrightarrow> P12 b (TVCKCons tvar ckind tvck_lst)"
1632
+ − 370
shows "P1 (a :: 'a :: pt) tkind \<and>
1631
+ − 371
P2 (b :: 'b :: pt) ckind \<and>
1634
+ − 372
P3 (c :: 'c :: {pt,fs}) ty \<and>
1631
+ − 373
P4 (d :: 'd :: pt) ty_lst \<and>
1635
+ − 374
P5 (e :: 'e :: {pt,fs}) co \<and>
1631
+ − 375
P6 (f :: 'f :: pt) co_lst \<and>
1645
+ − 376
P7 (g :: 'g :: {pt,fs}) trm \<and>
1635
+ − 377
P8 (h :: 'h :: {pt,fs}) assoc_lst \<and>
1631
+ − 378
P9 (i :: 'i :: pt) pat \<and>
+ − 379
P10 (j :: 'j :: pt) vt_lst \<and>
+ − 380
P11 (k :: 'k :: pt) tvtk_lst \<and>
+ − 381
P12 (l :: 'l :: pt) tvck_lst"
+ − 382
proof -
1647
+ − 383
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> pat)))" and a2:"(\<forall>p q j. P10 j (permute_bv_vt q (p \<bullet> vt_lst)))" and a3:"(\<forall>p q k. P11 k ( permute_bv_tvtk q (p \<bullet> tvtk_lst)))" and a4:"(\<forall>p q l. P12 l (permute_bv_tvck q (p \<bullet> tvck_lst)))"
+ − 384
apply (induct rule: tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vt_lst_tvtk_lst_tvck_lst.inducts)
1632
+ − 385
apply (tactic {* ALLGOALS (REPEAT o rtac allI) *})
+ − 386
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
+ − 387
+ − 388
(* GOAL1 *)
+ − 389
apply(subgoal_tac "\<exists>pa. ((pa \<bullet> (atom (p \<bullet> tvar))) \<sharp> c \<and>
+ − 390
supp (Abs (p \<bullet> {atom tvar}) (p \<bullet> ty)) \<sharp>* pa)")
+ − 391
apply clarify
+ − 392
apply (simp only: perm)
+ − 393
apply(rule_tac t="TAll (p \<bullet> tvar) (p \<bullet> tkind) (p \<bullet> ty)"
+ − 394
and s="TAll (pa \<bullet> p \<bullet> tvar) (p \<bullet> tkind) (pa \<bullet> p \<bullet> ty)" in subst)
+ − 395
apply (simp only: eq_iff)
+ − 396
apply (rule_tac x="-pa" in exI)
+ − 397
apply (simp add: alphas eqvts eqvts_raw supp_Abs fv_supp)
+ − 398
apply (rule_tac t="supp (pa \<bullet> p \<bullet> ty) - {atom (pa \<bullet> p \<bullet> tvar)}"
+ − 399
and s="pa \<bullet> (p \<bullet> supp ty - {p \<bullet> atom tvar})" in subst)
+ − 400
apply (simp add: eqvts)
+ − 401
apply (simp add: eqvts[symmetric])
+ − 402
apply (rule conjI)
+ − 403
apply (rule supp_perm_eq)
+ − 404
apply (simp add: eqvts)
+ − 405
apply (subst supp_finite_atom_set)
+ − 406
apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp)
+ − 407
apply (simp add: eqvts)
+ − 408
apply (simp add: eqvts)
+ − 409
apply (subst supp_perm_eq)
+ − 410
apply (subst supp_finite_atom_set)
+ − 411
apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp)
+ − 412
apply (simp add: eqvts)
+ − 413
apply assumption
+ − 414
apply (simp add: fresh_star_minus_perm)
+ − 415
apply (rule a08)
+ − 416
apply simp
+ − 417
apply(rotate_tac 1)
+ − 418
apply(erule_tac x="(pa + p)" in allE)
+ − 419
apply simp
+ − 420
apply (simp add: eqvts eqvts_raw)
1634
+ − 421
apply (rule at_set_avoiding2_atom)
+ − 422
apply (simp add: finite_supp)
+ − 423
apply (simp add: finite_supp)
+ − 424
apply (simp add: fresh_def)
+ − 425
apply (simp only: supp_Abs eqvts)
+ − 426
apply blast
1635
+ − 427
+ − 428
(* GOAL2 *)
+ − 429
apply(subgoal_tac "\<exists>pa. ((pa \<bullet> (atom (p \<bullet> tvar))) \<sharp> e \<and>
+ − 430
supp (Abs (p \<bullet> {atom tvar}) (p \<bullet> co)) \<sharp>* pa)")
1634
+ − 431
apply clarify
+ − 432
apply (simp only: perm)
1635
+ − 433
apply(rule_tac t="CAll (p \<bullet> tvar) (p \<bullet> ckind) (p \<bullet> co)"
+ − 434
and s="CAll (pa \<bullet> p \<bullet> tvar) (p \<bullet> ckind) (pa \<bullet> p \<bullet> co)" in subst)
+ − 435
apply (simp only: eq_iff)
1634
+ − 436
apply (rule_tac x="-pa" in exI)
1635
+ − 437
apply (simp add: alphas eqvts eqvts_raw supp_Abs fv_supp)
+ − 438
apply (rule_tac t="supp (pa \<bullet> p \<bullet> co) - {atom (pa \<bullet> p \<bullet> tvar)}"
+ − 439
and s="pa \<bullet> (p \<bullet> supp co - {p \<bullet> atom tvar})" in subst)
+ − 440
apply (simp add: eqvts)
+ − 441
apply (simp add: eqvts[symmetric])
1634
+ − 442
apply (rule conjI)
+ − 443
apply (rule supp_perm_eq)
+ − 444
apply (simp add: eqvts)
+ − 445
apply (subst supp_finite_atom_set)
+ − 446
apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp)
+ − 447
apply (simp add: eqvts)
+ − 448
apply (simp add: eqvts)
+ − 449
apply (subst supp_perm_eq)
+ − 450
apply (subst supp_finite_atom_set)
+ − 451
apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp)
+ − 452
apply (simp add: eqvts)
+ − 453
apply assumption
+ − 454
apply (simp add: fresh_star_minus_perm)
1635
+ − 455
apply (rule a15)
+ − 456
apply simp
+ − 457
apply(rotate_tac 1)
+ − 458
apply(erule_tac x="(pa + p)" in allE)
+ − 459
apply simp
+ − 460
apply (simp add: eqvts eqvts_raw)
+ − 461
apply (rule at_set_avoiding2_atom)
+ − 462
apply (simp add: finite_supp)
+ − 463
apply (simp add: finite_supp)
+ − 464
apply (simp add: fresh_def)
+ − 465
apply (simp only: supp_Abs eqvts)
+ − 466
apply blast
+ − 467
1645
+ − 468
+ − 469
(* GOAL3 a copy-and-paste of Goal2 with consts and variable names changed *)
+ − 470
apply(subgoal_tac "\<exists>pa. ((pa \<bullet> (atom (p \<bullet> tvar))) \<sharp> g \<and>
+ − 471
supp (Abs (p \<bullet> {atom tvar}) (p \<bullet> trm)) \<sharp>* pa)")
+ − 472
apply clarify
+ − 473
apply (simp only: perm)
+ − 474
apply(rule_tac t="LAMT (p \<bullet> tvar) (p \<bullet> tkind) (p \<bullet> trm)"
+ − 475
and s="LAMT (pa \<bullet> p \<bullet> tvar) (p \<bullet> tkind) (pa \<bullet> p \<bullet> trm)" in subst)
+ − 476
apply (simp only: eq_iff)
+ − 477
apply (rule_tac x="-pa" in exI)
+ − 478
apply (simp add: alphas eqvts eqvts_raw supp_Abs fv_supp)
+ − 479
apply (rule_tac t="supp (pa \<bullet> p \<bullet> trm) - {atom (pa \<bullet> p \<bullet> tvar)}"
+ − 480
and s="pa \<bullet> (p \<bullet> supp trm - {p \<bullet> atom tvar})" in subst)
+ − 481
apply (simp add: eqvts)
+ − 482
apply (simp add: eqvts[symmetric])
+ − 483
apply (rule conjI)
+ − 484
apply (rule supp_perm_eq)
+ − 485
apply (simp add: eqvts)
+ − 486
apply (subst supp_finite_atom_set)
+ − 487
apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp)
+ − 488
apply (simp add: eqvts)
+ − 489
apply (simp add: eqvts)
+ − 490
apply (subst supp_perm_eq)
+ − 491
apply (subst supp_finite_atom_set)
+ − 492
apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp)
+ − 493
apply (simp add: eqvts)
+ − 494
apply assumption
+ − 495
apply (simp add: fresh_star_minus_perm)
+ − 496
apply (rule a29)
+ − 497
apply simp
+ − 498
apply(rotate_tac 1)
+ − 499
apply(erule_tac x="(pa + p)" in allE)
+ − 500
apply simp
+ − 501
apply (simp add: eqvts eqvts_raw)
+ − 502
apply (rule at_set_avoiding2_atom)
+ − 503
apply (simp add: finite_supp)
+ − 504
apply (simp add: finite_supp)
+ − 505
apply (simp add: fresh_def)
+ − 506
apply (simp only: supp_Abs eqvts)
+ − 507
apply blast
+ − 508
1646
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 509
(* GOAL4 a copy-and-paste *)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 510
apply(subgoal_tac "\<exists>pa. ((pa \<bullet> (atom (p \<bullet> tvar))) \<sharp> g \<and>
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 511
supp (Abs (p \<bullet> {atom tvar}) (p \<bullet> trm)) \<sharp>* pa)")
1635
+ − 512
apply clarify
+ − 513
apply (simp only: perm)
1646
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 514
apply(rule_tac t="LAMC (p \<bullet> tvar) (p \<bullet> ckind) (p \<bullet> trm)"
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 515
and s="LAMC (pa \<bullet> p \<bullet> tvar) (p \<bullet> ckind) (pa \<bullet> p \<bullet> trm)" in subst)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 516
apply (simp only: eq_iff)
1635
+ − 517
apply (rule_tac x="-pa" in exI)
+ − 518
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>
diff
changeset
+ − 519
apply (rule_tac t="supp (pa \<bullet> p \<bullet> trm) - {atom (pa \<bullet> p \<bullet> tvar)}"
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 520
and s="pa \<bullet> (p \<bullet> supp trm - {p \<bullet> atom tvar})" in subst)
1635
+ − 521
apply (simp add: eqvts)
1646
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 522
apply (simp add: eqvts[symmetric])
1635
+ − 523
apply (rule conjI)
1646
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 524
apply (rule supp_perm_eq)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
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>
diff
changeset
+ − 526
apply (subst supp_finite_atom_set)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 527
apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 528
apply (simp add: eqvts)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 529
apply (simp add: eqvts)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 530
apply (subst supp_perm_eq)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 531
apply (subst supp_finite_atom_set)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 532
apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 533
apply (simp add: eqvts)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 534
apply assumption
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 535
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>
diff
changeset
+ − 536
apply (rule a30)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 537
apply simp
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 538
apply(rotate_tac 1)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 539
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>
diff
changeset
+ − 540
apply simp
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 541
apply (simp add: eqvts eqvts_raw)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 542
apply (rule at_set_avoiding2_atom)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 543
apply (simp add: finite_supp)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 544
apply (simp add: finite_supp)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 545
apply (simp add: fresh_def)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 546
apply (simp only: supp_Abs eqvts)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 547
apply blast
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 548
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 549
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 550
(* GOAL5 a copy-and-paste *)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 551
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>
diff
changeset
+ − 552
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>
diff
changeset
+ − 553
apply clarify
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 554
apply (simp only: perm)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 555
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>
diff
changeset
+ − 556
and s="Lam (pa \<bullet> p \<bullet> var) (p \<bullet> ty) (pa \<bullet> p \<bullet> trm)" in subst)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 557
apply (simp only: eq_iff)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 558
apply (rule_tac x="-pa" in exI)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 559
apply (simp add: alphas eqvts eqvts_raw supp_Abs fv_supp)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 560
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>
diff
changeset
+ − 561
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>
diff
changeset
+ − 562
apply (simp add: eqvts)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 563
apply (simp add: eqvts[symmetric])
1635
+ − 564
apply (rule conjI)
+ − 565
apply (rule supp_perm_eq)
+ − 566
apply (simp add: eqvts)
+ − 567
apply (subst supp_finite_atom_set)
+ − 568
apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp)
+ − 569
apply (simp add: eqvts)
+ − 570
apply (simp add: eqvts)
+ − 571
apply (subst supp_perm_eq)
+ − 572
apply (subst supp_finite_atom_set)
+ − 573
apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp)
+ − 574
apply (simp add: eqvts)
+ − 575
apply assumption
+ − 576
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>
diff
changeset
+ − 577
apply (rule a32)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 578
apply simp
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 579
apply(rotate_tac 1)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 580
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>
diff
changeset
+ − 581
apply simp
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 582
apply (simp add: eqvts eqvts_raw)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 583
apply (rule at_set_avoiding2_atom)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 584
apply (simp add: finite_supp)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 585
apply (simp add: finite_supp)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 586
apply (simp add: fresh_def)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 587
apply (simp only: supp_Abs eqvts)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 588
apply blast
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 589
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 590
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 591
(* GOAL6 a copy-and-paste *)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 592
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>
diff
changeset
+ − 593
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>
diff
changeset
+ − 594
apply clarify
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 595
apply (simp only: perm)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 596
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>
diff
changeset
+ − 597
and s="Let (pa \<bullet> p \<bullet> var) (p \<bullet> ty) (p \<bullet> trm1) (pa \<bullet> p \<bullet> trm2)" in subst)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 598
apply (simp only: eq_iff)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 599
apply (rule_tac x="-pa" in exI)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 600
apply (simp add: alphas eqvts eqvts_raw supp_Abs fv_supp)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 601
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>
diff
changeset
+ − 602
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>
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>
diff
changeset
+ − 604
apply (simp add: eqvts[symmetric])
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 605
apply (rule conjI)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 606
apply (rule supp_perm_eq)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 607
apply (simp add: eqvts)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 608
apply (subst supp_finite_atom_set)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 609
apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 610
apply (simp add: eqvts)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 611
apply (simp add: eqvts)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 612
apply (subst supp_perm_eq)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 613
apply (subst supp_finite_atom_set)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 614
apply (simp add: eqvts[symmetric] finite_eqvt[symmetric] finite_supp)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 615
apply (simp add: eqvts)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 616
apply assumption
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 617
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>
diff
changeset
+ − 618
apply (rule a34)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 619
apply simp
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 620
apply simp
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 621
apply(rotate_tac 2)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 622
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>
diff
changeset
+ − 623
apply simp
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 624
apply (simp add: eqvts eqvts_raw)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 625
apply (rule at_set_avoiding2_atom)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 626
apply (simp add: finite_supp)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 627
apply (simp add: finite_supp)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 628
apply (simp add: fresh_def)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 629
apply (simp only: supp_Abs eqvts)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 630
apply blast
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 631
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 632
(* MAIN ACons Goal *)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 633
apply(subgoal_tac "\<exists>pa. ((pa \<bullet> (bv (p \<bullet> pat))) \<sharp>* h \<and>
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 634
supp (Abs (p \<bullet> (bv pat)) (p \<bullet> trm)) \<sharp>* pa)")
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 635
apply clarify
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 636
apply (simp only: perm eqvts)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 637
apply (subst ACons_subst)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 638
apply assumption
1635
+ − 639
apply (rule a38)
+ − 640
apply simp
+ − 641
apply(rotate_tac 1)
+ − 642
apply(erule_tac x="(pa + p)" in allE)
+ − 643
apply simp
+ − 644
apply simp
1647
+ − 645
apply (simp add: perm_bv2[symmetric])
1635
+ − 646
apply (simp add: eqvts eqvts_raw)
+ − 647
apply (rule at_set_avoiding2)
+ − 648
apply (simp add: fin_bv)
+ − 649
apply (simp add: finite_supp)
1646
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 650
apply (simp add: supp_Abs)
1635
+ − 651
apply (rule finite_Diff)
1646
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 652
apply (simp add: finite_supp)
733bac87d5bf
Solved all boring subgoals, and looking at properly defning permute_bv
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 653
apply (simp add: fresh_star_def fresh_def supp_Abs eqvts)
1647
+ − 654
done
+ − 655
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+)
+ − 656
moreover have "P9 i (permute_bv 0 (0 \<bullet> pat))" and "P10 j (permute_bv_vt 0 (0 \<bullet> vt_lst))" and "P11 k (permute_bv_tvtk 0 (0 \<bullet> tvtk_lst))" and "P12 l (permute_bv_tvck 0 (0 \<bullet> tvck_lst))" using a1 a2 a3 a4 by (blast+)
+ − 657
ultimately show ?thesis by (simp_all add: permute_bv_zero1 permute_bv_zero2)
1631
+ − 658
qed
1615
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 659
1601
+ − 660
end
+ − 661