author | Christian Urban <urbanc@in.tum.de> |
Thu, 26 Aug 2010 02:08:00 +0800 | |
changeset 2436 | 3885dc2669f9 |
parent 2311 | Nominal/Ex/Ex1rec.thy@4da5c5c29009 |
child 2438 | abafea9b39bb |
permissions | -rw-r--r-- |
2436
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
1 |
theory LetRec |
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 |
|
2436
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
5 |
declare [[STEPS = 14]] |
1596
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
6 |
|
2311
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2120
diff
changeset
|
7 |
atom_decl name |
2056 | 8 |
|
2436
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
9 |
nominal_datatype let_rec: |
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
10 |
lam = |
2056 | 11 |
Var "name" |
12 |
| App "lam" "lam" |
|
2436
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
13 |
| Lam x::"name" t::"lam" bind (set) x in t |
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
14 |
| Let_Rec 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
|
15 |
and bp = |
2056 | 16 |
Bp "name" "lam" |
1596
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
17 |
binder |
1946
3395e87d94b8
tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents:
1773
diff
changeset
|
18 |
bi::"bp \<Rightarrow> atom set" |
1596
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
19 |
where |
2056 | 20 |
"bi (Bp x t) = {atom x}" |
1596
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
21 |
|
2436
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
22 |
thm let_rec.distinct |
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
23 |
thm let_rec.induct |
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
24 |
thm let_rec.exhaust |
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
25 |
thm let_rec.fv_defs |
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
26 |
thm let_rec.bn_defs |
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
27 |
thm let_rec.perm_simps |
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
28 |
thm let_rec.eq_iff |
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
29 |
thm let_rec.fv_bn_eqvt |
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
30 |
thm let_rec.size_eqvt |
2082
0854af516f14
cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents:
2067
diff
changeset
|
31 |
|
2067
ace7775cbd04
Experiments with equivariance.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2056
diff
changeset
|
32 |
|
1596
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
33 |
end |
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
34 |
|
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
35 |
|
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
36 |