author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Tue, 23 Mar 2010 09:34:32 +0100 | |
changeset 1604 | 5ab97f43ec24 |
parent 1595 | Nominal/LFex.thy@aeed597d2043 |
permissions | -rw-r--r-- |
1604
5ab97f43ec24
More modification needed for compilation
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1595
diff
changeset
|
1 |
theory ExLF |
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 |
|
1348
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
8 |
nominal_datatype kind = |
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
9 |
Type |
1348
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
10 |
| 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
|
11 |
and ty = |
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
12 |
TConst "ident" |
1348
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
13 |
| TApp "ty" "trm" |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
14 |
| 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
|
15 |
and trm = |
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" |
1348
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
18 |
| App "trm" "trm" |
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
19 |
| 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
|
20 |
|
1553
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
21 |
thm kind_ty_trm.supp |
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
|
22 |
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
23 |
end |
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
24 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
25 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
26 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
27 |