author | Christian Urban <urbanc@in.tum.de> |
Sun, 29 Aug 2010 01:45:07 +0800 | |
changeset 2451 | d2e929f51fa9 |
parent 2440 | 0a36825b16c1 |
child 2454 | 9ffee4eb1ae1 |
permissions | -rw-r--r-- |
2083
9568f9f31822
tuned file names for examples
Christian Urban <urbanc@in.tum.de>
parents:
2082
diff
changeset
|
1 |
theory LF |
2059
c615095827e9
Move LF to NewParser. Just works.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1773
diff
changeset
|
2 |
imports "../NewParser" |
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 |
|
2440
0a36825b16c1
"isabelle make test" makes all major examples....they work up to supp theorems (excluding)
Christian Urban <urbanc@in.tum.de>
parents:
2436
diff
changeset
|
5 |
declare [[STEPS = 100]] |
2311
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2120
diff
changeset
|
6 |
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
7 |
atom_decl name |
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
8 |
atom_decl ident |
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
9 |
|
1348
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
10 |
nominal_datatype kind = |
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
11 |
Type |
2311
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2120
diff
changeset
|
12 |
| KPi "ty" n::"name" k::"kind" bind n in k |
1348
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
13 |
and ty = |
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
14 |
TConst "ident" |
1348
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
15 |
| TApp "ty" "trm" |
2436
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
16 |
| TPi "ty" n::"name" ty::"ty" bind n in ty |
1348
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
17 |
and trm = |
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
18 |
Const "ident" |
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
19 |
| Var "name" |
1348
2e2a3cd58f64
Ported LF to the parser interface.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1344
diff
changeset
|
20 |
| App "trm" "trm" |
2311
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2120
diff
changeset
|
21 |
| 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
|
22 |
|
2436
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
23 |
(*thm kind_ty_trm.supp*) |
2082
0854af516f14
cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents:
2059
diff
changeset
|
24 |
|
0854af516f14
cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents:
2059
diff
changeset
|
25 |
|
985
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
26 |
end |
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
27 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
28 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
29 |
|
ef8a2b0b237a
Ported existing part of LF to new permutations and alphas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
30 |