author | Christian Urban <urbanc@in.tum.de> |
Thu, 18 Mar 2010 23:38:01 +0100 | |
changeset 1527 | e1c74b864b1b |
parent 1515 | 76fa21f27f22 |
child 1547 | 57f7af5d7564 |
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 |
1348
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
2 |
imports "Parser" |
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 |
|
1454
7c8cd6eae8e2
FV_bn generated for recursive functions as well, and used in main fv for bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1444
diff
changeset
|
8 |
ML {* val _ = cheat_fv_rsp := false *} |
7c8cd6eae8e2
FV_bn generated for recursive functions as well, and used in main fv for bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1444
diff
changeset
|
9 |
ML {* val _ = cheat_const_rsp := false *} |
7c8cd6eae8e2
FV_bn generated for recursive functions as well, and used in main fv for bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1444
diff
changeset
|
10 |
ML {* val _ = cheat_equivp := false *} |
7c8cd6eae8e2
FV_bn generated for recursive functions as well, and used in main fv for bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1444
diff
changeset
|
11 |
ML {* val _ = cheat_fv_eqvt := false *} |
7c8cd6eae8e2
FV_bn generated for recursive functions as well, and used in main fv for bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1444
diff
changeset
|
12 |
ML {* val _ = cheat_alpha_eqvt := false *} |
1444
aca9a6380c3f
LF works with new alpha...?
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1429
diff
changeset
|
13 |
|
1348
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
14 |
nominal_datatype kind = |
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
15 |
Type |
1348
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
16 |
| KPi "ty" n::"name" k::"kind" bind n in k |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
17 |
and ty = |
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
18 |
TConst "ident" |
1348
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
19 |
| TApp "ty" "trm" |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
20 |
| TPi "ty" n::"name" t::"ty" bind n in t |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
21 |
and trm = |
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
22 |
Const "ident" |
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
23 |
| Var "name" |
1348
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
24 |
| App "trm" "trm" |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
25 |
| Lam "ty" n::"name" t::"trm" bind n in t |
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
|
26 |
|
1344
b320da14e63c
Fixed LF for one quantifier over 2 premises.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1310
diff
changeset
|
27 |
lemma ex_out: |
b320da14e63c
Fixed LF for one quantifier over 2 premises.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1310
diff
changeset
|
28 |
"(\<exists>x. Z x \<and> Q) = (Q \<and> (\<exists>x. Z x))" |
b320da14e63c
Fixed LF for one quantifier over 2 premises.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1310
diff
changeset
|
29 |
"(\<exists>x. Q \<and> Z x) = (Q \<and> (\<exists>x. Z x))" |
b320da14e63c
Fixed LF for one quantifier over 2 premises.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1310
diff
changeset
|
30 |
"(\<exists>x. P x \<and> Q \<and> Z x) = (Q \<and> (\<exists>x. P x \<and> Z x))" |
b320da14e63c
Fixed LF for one quantifier over 2 premises.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1310
diff
changeset
|
31 |
"(\<exists>x. Q \<and> P x \<and> Z x) = (Q \<and> (\<exists>x. P x \<and> Z x))" |
b320da14e63c
Fixed LF for one quantifier over 2 premises.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1310
diff
changeset
|
32 |
apply (blast)+ |
b320da14e63c
Fixed LF for one quantifier over 2 premises.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1310
diff
changeset
|
33 |
done |
b320da14e63c
Fixed LF for one quantifier over 2 premises.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1310
diff
changeset
|
34 |
|
1348
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
35 |
lemma Collect_neg_conj: "{x. \<not>(P x \<and> Q x)} = {x. \<not>(P x)} \<union> {x. \<not>(Q x)}" |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
36 |
by (simp add: Collect_imp_eq Collect_neg_eq[symmetric]) |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
37 |
|
1245
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
38 |
lemma supp_eqs: |
1348
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
39 |
"supp Type = {}" |
1498
2ff84b1f551f
Rename bound variables + minor cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1496
diff
changeset
|
40 |
"supp kind = fv_kind kind \<Longrightarrow> supp (KPi ty name kind) = supp ty \<union> supp (Abs {atom name} kind)" |
1348
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
41 |
"supp (TConst i) = {atom i}" |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
42 |
"supp (TApp A M) = supp A \<union> supp M" |
1498
2ff84b1f551f
Rename bound variables + minor cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1496
diff
changeset
|
43 |
"supp ty2 = fv_ty ty2 \<Longrightarrow> supp (TPi ty1 name ty2) = supp ty1 \<union> supp (Abs {atom name} ty2)" |
1348
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
44 |
"supp (Const i) = {atom i}" |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
45 |
"supp (Var x) = {atom x}" |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
46 |
"supp (App M N) = supp M \<union> supp N" |
1498
2ff84b1f551f
Rename bound variables + minor cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1496
diff
changeset
|
47 |
"supp trm = fv_trm trm \<Longrightarrow> supp (Lam ty name trm) = supp ty \<union> supp (Abs {atom name} trm)" |
1510
be911e869fde
Added fv,bn,distinct,perm to the simplifier.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
48 |
apply(simp_all (no_asm) add: supp_def permute_set_eq atom_eqvt) |
1515
76fa21f27f22
Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
49 |
apply(simp_all only: kind_ty_trm.eq_iff Abs_eq_iff alpha_gen) |
1348
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
50 |
apply(simp_all only: ex_out) |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
51 |
apply(simp_all only: eqvts[symmetric]) |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
52 |
apply(simp_all only: Collect_neg_conj) |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
53 |
apply(simp_all only: supp_at_base[simplified supp_def] Un_commute Un_assoc) |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
54 |
apply(simp_all add: Collect_imp_eq Collect_neg_eq[symmetric] Un_commute Un_assoc) |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
55 |
apply(simp_all add: Un_left_commute) |
1245
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
56 |
done |
1002
3f227ed7e3e5
More proofs in the LF example.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
997
diff
changeset
|
57 |
|
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
|
58 |
lemma supp_fv: |
1348
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
59 |
"supp t1 = fv_kind t1 \<and> supp t2 = fv_ty t2 \<and> supp t3 = fv_trm t3" |
1515
76fa21f27f22
Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
60 |
apply(induct rule: kind_ty_trm.induct) |
1510
be911e869fde
Added fv,bn,distinct,perm to the simplifier.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
61 |
apply(simp_all (no_asm) only: supp_eqs) |
1245
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
62 |
apply(simp_all) |
1348
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
63 |
apply(simp_all add: supp_eqs) |
1245
18310dff4e3a
Made the fv-supp proof much more straightforward.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1244
diff
changeset
|
64 |
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
|
65 |
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
|
66 |
|
1498
2ff84b1f551f
Rename bound variables + minor cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1496
diff
changeset
|
67 |
lemma supp_kind_ty_trm: |
1348
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
68 |
"supp Type = {}" |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
69 |
"supp (KPi A x K) = supp A \<union> (supp K - {atom x})" |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
70 |
"supp (TConst i) = {atom i}" |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
71 |
"supp (TApp A M) = supp A \<union> supp M" |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
72 |
"supp (TPi A x B) = supp A \<union> (supp B - {atom x})" |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
73 |
"supp (Const i) = {atom i}" |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
74 |
"supp (Var x) = {atom x}" |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
75 |
"supp (App M N) = supp M \<union> supp N" |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
76 |
"supp (Lam A x M) = supp A \<union> (supp M - {atom x})" |
1510
be911e869fde
Added fv,bn,distinct,perm to the simplifier.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
77 |
apply (simp_all add: supp_fv) |
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
|
78 |
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
79 |
end |
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
80 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
81 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
82 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
83 |