author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Mon, 10 May 2010 15:45:04 +0200 | |
changeset 2093 | 751d1349329b |
parent 2082 | 0854af516f14 |
child 2094 | 56b849d348ae |
permissions | -rw-r--r-- |
1596
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
1 |
theory Ex1rec |
2056 | 2 |
imports "../NewParser" |
1596
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
3 |
begin |
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
4 |
|
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
5 |
atom_decl name |
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
6 |
|
2056 | 7 |
ML {* val _ = cheat_supp_eq := true *} |
8 |
||
1946
3395e87d94b8
tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents:
1773
diff
changeset
|
9 |
nominal_datatype lam = |
2056 | 10 |
Var "name" |
11 |
| App "lam" "lam" |
|
12 |
| Lam x::"name" t::"lam" bind_set x in t |
|
13 |
| Let bp::"bp" t::"lam" bind_set "bi bp" in bp t |
|
1946
3395e87d94b8
tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents:
1773
diff
changeset
|
14 |
and bp = |
2056 | 15 |
Bp "name" "lam" |
1596
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
16 |
binder |
1946
3395e87d94b8
tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents:
1773
diff
changeset
|
17 |
bi::"bp \<Rightarrow> atom set" |
1596
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
18 |
where |
2056 | 19 |
"bi (Bp x t) = {atom x}" |
1596
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
20 |
|
1946
3395e87d94b8
tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents:
1773
diff
changeset
|
21 |
thm lam_bp.fv |
3395e87d94b8
tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents:
1773
diff
changeset
|
22 |
thm lam_bp.eq_iff[no_vars] |
3395e87d94b8
tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents:
1773
diff
changeset
|
23 |
thm lam_bp.bn |
3395e87d94b8
tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents:
1773
diff
changeset
|
24 |
thm lam_bp.perm |
3395e87d94b8
tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents:
1773
diff
changeset
|
25 |
thm lam_bp.induct |
3395e87d94b8
tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents:
1773
diff
changeset
|
26 |
thm lam_bp.inducts |
3395e87d94b8
tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents:
1773
diff
changeset
|
27 |
thm lam_bp.distinct |
2056 | 28 |
thm lam_bp.supp |
1946
3395e87d94b8
tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents:
1773
diff
changeset
|
29 |
ML {* Sign.of_sort @{theory} (@{typ lam}, @{sort fs}) *} |
2056 | 30 |
thm lam_bp.fv[simplified lam_bp.supp(1-2)] |
1596
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
31 |
|
2082
0854af516f14
cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents:
2067
diff
changeset
|
32 |
declare permute_lam_raw_permute_bp_raw.simps[eqvt] |
2067
ace7775cbd04
Experiments with equivariance.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2056
diff
changeset
|
33 |
declare alpha_gen_eqvt[eqvt] |
2082
0854af516f14
cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents:
2067
diff
changeset
|
34 |
|
2067
ace7775cbd04
Experiments with equivariance.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2056
diff
changeset
|
35 |
equivariance alpha_lam_raw |
2082
0854af516f14
cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents:
2067
diff
changeset
|
36 |
|
2067
ace7775cbd04
Experiments with equivariance.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2056
diff
changeset
|
37 |
|
1596
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
38 |
end |
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
39 |
|
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
40 |
|
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
41 |