author | Christian Urban <urbanc@in.tum.de> |
Tue, 13 Apr 2010 07:40:54 +0200 | |
changeset 1818 | 37480540c1af |
parent 1773 | c0eac04ae3b4 |
child 2082 | 0854af516f14 |
permissions | -rw-r--r-- |
1600 | 1 |
theory ExPS6 |
1773
c0eac04ae3b4
added README and moved examples into separate directory
Christian Urban <urbanc@in.tum.de>
parents:
1656
diff
changeset
|
2 |
imports "../Parser" |
1600 | 3 |
begin |
4 |
||
5 |
(* example 6 from Peter Sewell's bestiary *) |
|
6 |
||
7 |
atom_decl name |
|
8 |
||
1655
9cec4269b7f9
Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1604
diff
changeset
|
9 |
(* The binding structure is too complicated, so equivalence the |
9cec4269b7f9
Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1604
diff
changeset
|
10 |
way we define it is not true *) |
1604
5ab97f43ec24
More modification needed for compilation
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
11 |
|
1600 | 12 |
ML {* val _ = recursive := false *} |
13 |
nominal_datatype exp6 = |
|
14 |
EVar name |
|
15 |
| EPair exp6 exp6 |
|
16 |
| ELetRec x::name p::pat6 e1::exp6 e2::exp6 bind x in e1, bind x in e2, bind "bp6 p" in e1 |
|
17 |
and pat6 = |
|
18 |
PVar' name |
|
19 |
| PUnit' |
|
20 |
| PPair' pat6 pat6 |
|
21 |
binder |
|
22 |
bp6 :: "pat6 \<Rightarrow> atom set" |
|
23 |
where |
|
24 |
"bp6 (PVar' x) = {atom x}" |
|
25 |
| "bp6 (PUnit') = {}" |
|
26 |
| "bp6 (PPair' p1 p2) = bp6 p1 \<union> bp6 p2" |
|
27 |
||
28 |
thm exp6_pat6.fv |
|
29 |
thm exp6_pat6.eq_iff |
|
30 |
thm exp6_pat6.bn |
|
31 |
thm exp6_pat6.perm |
|
32 |
thm exp6_pat6.induct |
|
33 |
thm exp6_pat6.distinct |
|
1604
5ab97f43ec24
More modification needed for compilation
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1600
diff
changeset
|
34 |
thm exp6_pat6.supp |
1600 | 35 |
|
36 |
end |
|
37 |
||
38 |
||
39 |