author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Wed, 03 Mar 2010 17:47:29 +0100 | |
changeset 1334 | 80441e27dfd6 |
parent 1310 | c668d65fd988 |
child 1344 | b320da14e63c |
permissions | -rw-r--r-- |
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
1 |
theory LFex |
1241
ab1aa1b48547
Define the constants automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1240
diff
changeset
|
2 |
imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs" "Perm" "Fv" "Rsp" |
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
3 |
begin |
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
4 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
5 |
atom_decl name |
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
6 |
atom_decl ident |
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
7 |
|
1234 | 8 |
datatype rkind = |
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
9 |
Type |
1234 | 10 |
| KPi "rty" "name" "rkind" |
11 |
and rty = |
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
12 |
TConst "ident" |
1234 | 13 |
| TApp "rty" "rtrm" |
14 |
| TPi "rty" "name" "rty" |
|
15 |
and rtrm = |
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
16 |
Const "ident" |
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
17 |
| Var "name" |
1234 | 18 |
| App "rtrm" "rtrm" |
19 |
| Lam "rty" "name" "rtrm" |
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
20 |
|
1234 | 21 |
|
1277
6eacf60ce41d
Permutation and FV_Alpha interface change.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1264
diff
changeset
|
22 |
setup {* snd o define_raw_perms (Datatype.the_info @{theory} "LFex.rkind") 3 *} |
1253 | 23 |
print_theorems |
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
24 |
|
1219
c74c8ba46db7
All works in LF but will require renaming.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1218
diff
changeset
|
25 |
local_setup {* |
1277
6eacf60ce41d
Permutation and FV_Alpha interface change.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1264
diff
changeset
|
26 |
snd o define_fv_alpha (Datatype.the_info @{theory} "LFex.rkind") |
1300 | 27 |
[[ [], [(NONE, 1, 2)]], |
28 |
[ [], [], [(NONE, 1, 2)] ], |
|
29 |
[ [], [], [], [(NONE, 1, 2)]]] *} |
|
1219
c74c8ba46db7
All works in LF but will require renaming.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1218
diff
changeset
|
30 |
notation |
1234 | 31 |
alpha_rkind ("_ \<approx>ki _" [100, 100] 100) |
1237
38eb2bd9ad3a
LF renaming part 3 (proper names of alpha equvalences)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1236
diff
changeset
|
32 |
and alpha_rty ("_ \<approx>ty _" [100, 100] 100) |
1234 | 33 |
and alpha_rtrm ("_ \<approx>tr _" [100, 100] 100) |
34 |
thm fv_rkind_fv_rty_fv_rtrm.simps alpha_rkind_alpha_rty_alpha_rtrm.intros |
|
1238
c88159ffa7a3
Final synchronization of names.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1237
diff
changeset
|
35 |
local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha_rkind_alpha_rty_alpha_rtrm_inj}, []), (build_alpha_inj @{thms alpha_rkind_alpha_rty_alpha_rtrm.intros} @{thms rkind.distinct rty.distinct rtrm.distinct rkind.inject rty.inject rtrm.inject} @{thms alpha_rkind.cases alpha_rty.cases alpha_rtrm.cases} ctxt)) ctxt)) *} |
c88159ffa7a3
Final synchronization of names.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1237
diff
changeset
|
36 |
thm alpha_rkind_alpha_rty_alpha_rtrm_inj |
1219
c74c8ba46db7
All works in LF but will require renaming.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1218
diff
changeset
|
37 |
|
992
74e9a580448c
equivariance of rfv and alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
991
diff
changeset
|
38 |
lemma rfv_eqvt[eqvt]: |
1236
ca3c69545a78
LF renaming part 2 (proper fv functions)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1234
diff
changeset
|
39 |
"((pi\<bullet>fv_rkind t1) = fv_rkind (pi\<bullet>t1))" |
ca3c69545a78
LF renaming part 2 (proper fv functions)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1234
diff
changeset
|
40 |
"((pi\<bullet>fv_rty t2) = fv_rty (pi\<bullet>t2))" |
ca3c69545a78
LF renaming part 2 (proper fv functions)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1234
diff
changeset
|
41 |
"((pi\<bullet>fv_rtrm t3) = fv_rtrm (pi\<bullet>t3))" |
1234 | 42 |
apply(induct t1 and t2 and t3 rule: rkind_rty_rtrm.inducts) |
1239
ae73578feb64
Use the infrastructure in LF. Much shorter :).
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1238
diff
changeset
|
43 |
apply(simp_all add: union_eqvt Diff_eqvt) |
993
5c0d9a507bcb
Fixed the induction problem + some more proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
992
diff
changeset
|
44 |
apply(simp_all add: permute_set_eq atom_eqvt) |
992
74e9a580448c
equivariance of rfv and alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
991
diff
changeset
|
45 |
done |
74e9a580448c
equivariance of rfv and alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
991
diff
changeset
|
46 |
|
74e9a580448c
equivariance of rfv and alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
991
diff
changeset
|
47 |
lemma alpha_eqvt: |
1310 | 48 |
"(t1 \<approx>ki s1 \<longrightarrow> (p \<bullet> t1) \<approx>ki (p \<bullet> s1)) \<and> |
49 |
(t2 \<approx>ty s2 \<longrightarrow> (p \<bullet> t2) \<approx>ty (p \<bullet> s2)) \<and> |
|
50 |
(t3 \<approx>tr s3 \<longrightarrow> (p \<bullet> t3) \<approx>tr (p \<bullet> s3))" |
|
51 |
apply(rule alpha_rkind_alpha_rty_alpha_rtrm.induct) |
|
1237
38eb2bd9ad3a
LF renaming part 3 (proper names of alpha equvalences)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1236
diff
changeset
|
52 |
apply (simp_all add: alpha_rkind_alpha_rty_alpha_rtrm_inj) |
1264
1dedc0b76f32
Code for proving eqvt, still in Terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1259
diff
changeset
|
53 |
apply (erule alpha_gen_compose_eqvt) |
1dedc0b76f32
Code for proving eqvt, still in Terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1259
diff
changeset
|
54 |
apply (simp_all add: rfv_eqvt eqvts atom_eqvt) |
1dedc0b76f32
Code for proving eqvt, still in Terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1259
diff
changeset
|
55 |
apply (erule alpha_gen_compose_eqvt) |
1dedc0b76f32
Code for proving eqvt, still in Terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1259
diff
changeset
|
56 |
apply (simp_all add: rfv_eqvt eqvts atom_eqvt) |
1dedc0b76f32
Code for proving eqvt, still in Terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1259
diff
changeset
|
57 |
apply (erule alpha_gen_compose_eqvt) |
1dedc0b76f32
Code for proving eqvt, still in Terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1259
diff
changeset
|
58 |
apply (simp_all add: rfv_eqvt eqvts atom_eqvt) |
992
74e9a580448c
equivariance of rfv and alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
991
diff
changeset
|
59 |
done |
74e9a580448c
equivariance of rfv and alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
991
diff
changeset
|
60 |
|
1239
ae73578feb64
Use the infrastructure in LF. Much shorter :).
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1238
diff
changeset
|
61 |
local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha_equivps}, []), |
ae73578feb64
Use the infrastructure in LF. Much shorter :).
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1238
diff
changeset
|
62 |
(build_equivps [@{term alpha_rkind}, @{term alpha_rty}, @{term alpha_rtrm}] |
ae73578feb64
Use the infrastructure in LF. Much shorter :).
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1238
diff
changeset
|
63 |
@{thm rkind_rty_rtrm.induct} @{thm alpha_rkind_alpha_rty_alpha_rtrm.induct} |
ae73578feb64
Use the infrastructure in LF. Much shorter :).
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1238
diff
changeset
|
64 |
@{thms rkind.inject rty.inject rtrm.inject} @{thms alpha_rkind_alpha_rty_alpha_rtrm_inj} |
ae73578feb64
Use the infrastructure in LF. Much shorter :).
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1238
diff
changeset
|
65 |
@{thms rkind.distinct rty.distinct rtrm.distinct} |
ae73578feb64
Use the infrastructure in LF. Much shorter :).
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1238
diff
changeset
|
66 |
@{thms alpha_rkind.cases alpha_rty.cases alpha_rtrm.cases} |
ae73578feb64
Use the infrastructure in LF. Much shorter :).
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1238
diff
changeset
|
67 |
@{thms alpha_eqvt} ctxt)) ctxt)) *} |
ae73578feb64
Use the infrastructure in LF. Much shorter :).
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1238
diff
changeset
|
68 |
thm alpha_equivps |
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
69 |
|
1241
ab1aa1b48547
Define the constants automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1240
diff
changeset
|
70 |
local_setup {* define_quotient_type |
ab1aa1b48547
Define the constants automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1240
diff
changeset
|
71 |
[(([], @{binding kind}, NoSyn), (@{typ rkind}, @{term alpha_rkind})), |
ab1aa1b48547
Define the constants automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1240
diff
changeset
|
72 |
(([], @{binding ty}, NoSyn), (@{typ rty}, @{term alpha_rty} )), |
ab1aa1b48547
Define the constants automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1240
diff
changeset
|
73 |
(([], @{binding trm}, NoSyn), (@{typ rtrm}, @{term alpha_rtrm} ))] |
ab1aa1b48547
Define the constants automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1240
diff
changeset
|
74 |
(ALLGOALS (resolve_tac @{thms alpha_equivps})) |
ab1aa1b48547
Define the constants automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1240
diff
changeset
|
75 |
*} |
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
76 |
|
1241
ab1aa1b48547
Define the constants automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1240
diff
changeset
|
77 |
local_setup {* |
ab1aa1b48547
Define the constants automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1240
diff
changeset
|
78 |
(fn ctxt => ctxt |
ab1aa1b48547
Define the constants automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1240
diff
changeset
|
79 |
|> snd o (Quotient_Def.quotient_lift_const ("TYP", @{term Type})) |
ab1aa1b48547
Define the constants automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1240
diff
changeset
|
80 |
|> snd o (Quotient_Def.quotient_lift_const ("KPI", @{term KPi})) |
ab1aa1b48547
Define the constants automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1240
diff
changeset
|
81 |
|> snd o (Quotient_Def.quotient_lift_const ("TCONST", @{term TConst})) |
ab1aa1b48547
Define the constants automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1240
diff
changeset
|
82 |
|> snd o (Quotient_Def.quotient_lift_const ("TAPP", @{term TApp})) |
ab1aa1b48547
Define the constants automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1240
diff
changeset
|
83 |
|> snd o (Quotient_Def.quotient_lift_const ("TPI", @{term TPi})) |
ab1aa1b48547
Define the constants automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1240
diff
changeset
|
84 |
|> snd o (Quotient_Def.quotient_lift_const ("CONS", @{term Const})) |
ab1aa1b48547
Define the constants automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1240
diff
changeset
|
85 |
|> snd o (Quotient_Def.quotient_lift_const ("VAR", @{term Var})) |
ab1aa1b48547
Define the constants automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1240
diff
changeset
|
86 |
|> snd o (Quotient_Def.quotient_lift_const ("APP", @{term App})) |
ab1aa1b48547
Define the constants automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1240
diff
changeset
|
87 |
|> snd o (Quotient_Def.quotient_lift_const ("LAM", @{term Lam})) |
ab1aa1b48547
Define the constants automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1240
diff
changeset
|
88 |
|> snd o (Quotient_Def.quotient_lift_const ("fv_kind", @{term fv_rkind})) |
ab1aa1b48547
Define the constants automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1240
diff
changeset
|
89 |
|> snd o (Quotient_Def.quotient_lift_const ("fv_ty", @{term fv_rty})) |
ab1aa1b48547
Define the constants automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1240
diff
changeset
|
90 |
|> snd o (Quotient_Def.quotient_lift_const ("fv_trm", @{term fv_rtrm}))) *} |
ab1aa1b48547
Define the constants automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1240
diff
changeset
|
91 |
print_theorems |
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
92 |
|
1278
8814494fe4da
Change in signature of prove_const_rsp for general lifting.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
93 |
local_setup {* snd o prove_const_rsp @{binding rfv_rsp} [@{term fv_rkind}, @{term fv_rty}, @{term fv_rtrm}] |
1242
76de3440887c
Generate fv_rsp automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1241
diff
changeset
|
94 |
(fn _ => fvbv_rsp_tac @{thm alpha_rkind_alpha_rty_alpha_rtrm.induct} @{thms fv_rkind_fv_rty_fv_rtrm.simps} 1) *} |
1279
d53b7f24450b
RSP of perms can be shown in one go.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1278
diff
changeset
|
95 |
local_setup {* snd o prove_const_rsp Binding.empty [@{term "permute :: perm \<Rightarrow> rkind \<Rightarrow> rkind"}, @{term "permute :: perm \<Rightarrow> rty \<Rightarrow> rty"}, @{term "permute :: perm \<Rightarrow> rtrm \<Rightarrow> rtrm"}] |
1243
14cf3d2b0e16
Respects of permute and constructors.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1242
diff
changeset
|
96 |
(fn _ => asm_simp_tac (HOL_ss addsimps @{thms alpha_eqvt}) 1) *} |
14cf3d2b0e16
Respects of permute and constructors.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1242
diff
changeset
|
97 |
ML {* fun const_rsp_tac _ = constr_rsp_tac @{thms alpha_rkind_alpha_rty_alpha_rtrm_inj} |
14cf3d2b0e16
Respects of permute and constructors.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1242
diff
changeset
|
98 |
@{thms rfv_rsp} @{thms alpha_equivps} 1 *} |
1278
8814494fe4da
Change in signature of prove_const_rsp for general lifting.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
99 |
local_setup {* snd o prove_const_rsp Binding.empty [@{term TConst}] const_rsp_tac *} |
8814494fe4da
Change in signature of prove_const_rsp for general lifting.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
100 |
local_setup {* snd o prove_const_rsp Binding.empty [@{term TApp}] const_rsp_tac *} |
8814494fe4da
Change in signature of prove_const_rsp for general lifting.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
101 |
local_setup {* snd o prove_const_rsp Binding.empty [@{term Var}] const_rsp_tac *} |
8814494fe4da
Change in signature of prove_const_rsp for general lifting.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
102 |
local_setup {* snd o prove_const_rsp Binding.empty [@{term App}] const_rsp_tac *} |
8814494fe4da
Change in signature of prove_const_rsp for general lifting.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
103 |
local_setup {* snd o prove_const_rsp Binding.empty [@{term Const}] const_rsp_tac *} |
8814494fe4da
Change in signature of prove_const_rsp for general lifting.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
104 |
local_setup {* snd o prove_const_rsp Binding.empty [@{term KPi}] const_rsp_tac *} |
8814494fe4da
Change in signature of prove_const_rsp for general lifting.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
105 |
local_setup {* snd o prove_const_rsp Binding.empty [@{term TPi}] const_rsp_tac *} |
8814494fe4da
Change in signature of prove_const_rsp for general lifting.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
106 |
local_setup {* snd o prove_const_rsp Binding.empty [@{term Lam}] const_rsp_tac *} |
992
74e9a580448c
equivariance of rfv and alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
991
diff
changeset
|
107 |
|
1240
9348581d7d92
Rename also the lifted types to non-capital.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1239
diff
changeset
|
108 |
lemmas kind_ty_trm_induct = rkind_rty_rtrm.induct[quot_lifted] |
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
109 |
|
1234 | 110 |
thm rkind_rty_rtrm.inducts |
1240
9348581d7d92
Rename also the lifted types to non-capital.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1239
diff
changeset
|
111 |
lemmas kind_ty_trm_inducts = rkind_rty_rtrm.inducts[quot_lifted] |
994
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
993
diff
changeset
|
112 |
|
1253 | 113 |
setup {* define_lifted_perms ["LFex.kind", "LFex.ty", "LFex.trm"] |
114 |
[("permute_kind", @{term "permute :: perm \<Rightarrow> rkind \<Rightarrow> rkind"}), |
|
115 |
("permute_ty", @{term "permute :: perm \<Rightarrow> rty \<Rightarrow> rty"}), |
|
116 |
("permute_trm", @{term "permute :: perm \<Rightarrow> rtrm \<Rightarrow> rtrm"})] |
|
117 |
@{thms permute_rkind_permute_rty_permute_rtrm_zero permute_rkind_permute_rty_permute_rtrm_append} *} |
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
118 |
|
1246
845bf16eee18
Regularize finite support proof for trm1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1245
diff
changeset
|
119 |
(* |
845bf16eee18
Regularize finite support proof for trm1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1245
diff
changeset
|
120 |
Lifts, but slow and not needed?. |
845bf16eee18
Regularize finite support proof for trm1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1245
diff
changeset
|
121 |
lemmas alpha_kind_alpha_ty_alpha_trm_induct = alpha_rkind_alpha_rty_alpha_rtrm.induct[unfolded alpha_gen, quot_lifted, folded alpha_gen] |
845bf16eee18
Regularize finite support proof for trm1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1245
diff
changeset
|
122 |
*) |
994
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
993
diff
changeset
|
123 |
|
1250
d1ab27c10049
With permute_rsp we can lift the instance proofs :).
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1246
diff
changeset
|
124 |
lemmas permute_ktt[simp] = permute_rkind_permute_rty_permute_rtrm.simps[quot_lifted] |
d1ab27c10049
With permute_rsp we can lift the instance proofs :).
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1246
diff
changeset
|
125 |
|
1246
845bf16eee18
Regularize finite support proof for trm1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1245
diff
changeset
|
126 |
lemmas kind_ty_trm_inj = alpha_rkind_alpha_rty_alpha_rtrm_inj[unfolded alpha_gen, quot_lifted, folded alpha_gen] |
994
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
993
diff
changeset
|
127 |
|
1236
ca3c69545a78
LF renaming part 2 (proper fv functions)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1234
diff
changeset
|
128 |
lemmas fv_kind_ty_trm = fv_rkind_fv_rty_fv_rtrm.simps[quot_lifted] |
994
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
993
diff
changeset
|
129 |
|
1082
f8029d8c88a9
Fixed the context import/export and simplified LFex.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1045
diff
changeset
|
130 |
lemmas fv_eqvt = rfv_eqvt[quot_lifted] |
994
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
993
diff
changeset
|
131 |
|
1244
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
132 |
lemma supports: |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
133 |
"{} supports TYP" |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
134 |
"(supp (atom i)) supports (TCONST i)" |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
135 |
"(supp A \<union> supp M) supports (TAPP A M)" |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
136 |
"(supp (atom i)) supports (CONS i)" |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
137 |
"(supp (atom x)) supports (VAR x)" |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
138 |
"(supp M \<union> supp N) supports (APP M N)" |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
139 |
"(supp ty \<union> supp (atom na) \<union> supp ki) supports (KPI ty na ki)" |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
140 |
"(supp ty \<union> supp (atom na) \<union> supp ty2) supports (TPI ty na ty2)" |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
141 |
"(supp ty \<union> supp (atom na) \<union> supp trm) supports (LAM ty na trm)" |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
142 |
apply(simp_all add: supports_def fresh_def[symmetric] swap_fresh_fresh) |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
143 |
apply(rule_tac [!] allI)+ |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
144 |
apply(rule_tac [!] impI) |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
145 |
apply(tactic {* ALLGOALS (REPEAT o etac conjE) *}) |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
146 |
apply(simp_all add: fresh_atom) |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
147 |
done |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
148 |
|
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
149 |
lemma kind_ty_trm_fs: |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
150 |
"finite (supp (x\<Colon>kind))" |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
151 |
"finite (supp (y\<Colon>ty))" |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
152 |
"finite (supp (z\<Colon>trm))" |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
153 |
apply(induct x and y and z rule: kind_ty_trm_inducts) |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
154 |
apply(tactic {* ALLGOALS (rtac @{thm supports_finite} THEN' resolve_tac @{thms supports}) *}) |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
155 |
apply(simp_all add: supp_atom) |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
156 |
done |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
157 |
|
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
158 |
instance kind and ty and trm :: fs |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
159 |
apply(default) |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
160 |
apply(simp_all only: kind_ty_trm_fs) |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
161 |
done |
605a0ebe87da
Regularize the proofs about finite support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1243
diff
changeset
|
162 |
|
1245
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
163 |
lemma supp_eqs: |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
164 |
"supp TYP = {}" |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
165 |
"supp rkind = fv_kind rkind \<Longrightarrow> supp (KPI rty name rkind) = supp rty \<union> supp (Abs {atom name} rkind)" |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
166 |
"supp (TCONST i) = {atom i}" |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
167 |
"supp (TAPP A M) = supp A \<union> supp M" |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
168 |
"supp rty2 = fv_ty rty2 \<Longrightarrow> supp (TPI rty1 name rty2) = supp rty1 \<union> supp (Abs {atom name} rty2)" |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
169 |
"supp (CONS i) = {atom i}" |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
170 |
"supp (VAR x) = {atom x}" |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
171 |
"supp (APP M N) = supp M \<union> supp N" |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
172 |
"supp rtrm = fv_trm rtrm \<Longrightarrow> supp (LAM rty name rtrm) = supp rty \<union> supp (Abs {atom name} rtrm)" |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
173 |
apply(simp_all (no_asm) add: supp_def) |
1246
845bf16eee18
Regularize finite support proof for trm1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1245
diff
changeset
|
174 |
apply(simp_all only: kind_ty_trm_inj Abs_eq_iff alpha_gen) |
1245
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
175 |
apply(simp_all only: insert_eqvt empty_eqvt atom_eqvt supp_eqvt[symmetric] fv_eqvt[symmetric]) |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
176 |
apply(simp_all add: Collect_imp_eq Collect_neg_eq[symmetric] Set.Un_commute) |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
177 |
apply(simp_all add: supp_at_base[simplified supp_def]) |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
178 |
done |
1002
3f227ed7e3e5
More proofs in the LF example.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
997
diff
changeset
|
179 |
|
994
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
993
diff
changeset
|
180 |
lemma supp_fv: |
1245
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
181 |
"supp t1 = fv_kind t1" |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
182 |
"supp t2 = fv_ty t2" |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
183 |
"supp t3 = fv_trm t3" |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
184 |
apply(induct t1 and t2 and t3 rule: kind_ty_trm_inducts) |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
185 |
apply(simp_all (no_asm) only: supp_eqs fv_kind_ty_trm) |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
186 |
apply(simp_all) |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
187 |
apply(subst supp_eqs) |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
188 |
apply(simp_all add: supp_Abs) |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
189 |
apply(subst supp_eqs) |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
190 |
apply(simp_all add: supp_Abs) |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
191 |
apply(subst supp_eqs) |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
192 |
apply(simp_all add: supp_Abs) |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
193 |
done |
994
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
993
diff
changeset
|
194 |
|
1234 | 195 |
lemma supp_rkind_rty_rtrm: |
1245
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
196 |
"supp TYP = {}" |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
197 |
"supp (KPI A x K) = supp A \<union> (supp K - {atom x})" |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
198 |
"supp (TCONST i) = {atom i}" |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
199 |
"supp (TAPP A M) = supp A \<union> supp M" |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
200 |
"supp (TPI A x B) = supp A \<union> (supp B - {atom x})" |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
201 |
"supp (CONS i) = {atom i}" |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
202 |
"supp (VAR x) = {atom x}" |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
203 |
"supp (APP M N) = supp M \<union> supp N" |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
204 |
"supp (LAM A x M) = supp A \<union> (supp M - {atom x})" |
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
205 |
by (simp_all only: supp_fv fv_kind_ty_trm) |
994
333c24bd595d
More in the LF example in the new nominal way, all is clear until support.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
993
diff
changeset
|
206 |
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
207 |
end |
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
208 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
209 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
210 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
211 |