6 (* This file contains only the examples that are not supposed to work yet. *) |
6 (* This file contains only the examples that are not supposed to work yet. *) |
7 |
7 |
8 |
8 |
9 atom_decl name |
9 atom_decl name |
10 |
10 |
11 ML {* val _ = recursive := false *} |
|
12 |
|
13 (* example 7 from Peter Sewell's bestiary *) |
|
14 (* dest_Const raised |
|
15 nominal_datatype exp7 = |
|
16 EVar' name |
|
17 | EUnit' |
|
18 | EPair' exp7 exp7 |
|
19 | ELetRec' l::lrbs e::exp7 bind "b7s l" in e, bind "b7s l" in l |
|
20 and lrb = |
|
21 Assign' name exp7 |
|
22 and lrbs = |
|
23 Single' lrb |
|
24 | More' lrb lrbs |
|
25 binder |
|
26 b7 :: "lrb \<Rightarrow> atom set" and |
|
27 b7s :: "lrbs \<Rightarrow> atom set" |
|
28 where |
|
29 "b7 (Assign x e) = {atom x}" |
|
30 | "b7s (Single a) = b7 a" |
|
31 | "b7s (More a as) = (b7 a) \<union> (b7s as)" |
|
32 thm alpha_exp7_raw_alpha_lrb_raw_alpha_lrbs_raw.intros |
|
33 *) |
|
34 |
|
35 (* example 8 from Peter Sewell's bestiary *) |
|
36 (* |
|
37 *** fv_bn: recursive argument, but wrong datatype. |
|
38 *** At command "nominal_datatype". |
|
39 nominal_datatype exp8 = |
|
40 EVar' name |
|
41 | EUnit' |
|
42 | EPair' exp8 exp8 |
|
43 | ELetRec' l::lrbs8 e::exp8 bind "b_lrbs8 l" in e, bind "b_lrbs8 l" in l |
|
44 and fnclause = |
|
45 K' x::name p::pat8 e::exp8 bind "b_pat p" in e |
|
46 and fnclauses = |
|
47 S' fnclause |
|
48 | ORs' fnclause fnclauses |
|
49 and lrb8 = |
|
50 Clause' fnclauses |
|
51 and lrbs8 = |
|
52 Single' lrb8 |
|
53 | More' lrb8 lrbs8 |
|
54 and pat8 = |
|
55 PVar'' name |
|
56 | PUnit'' |
|
57 | PPair'' pat8 pat8 |
|
58 binder |
|
59 b_lrbs8 :: "lrbs8 \<Rightarrow> atom set" and |
|
60 b_pat :: "pat8 \<Rightarrow> atom set" and |
|
61 b_fnclauses :: "fnclauses \<Rightarrow> atom set" and |
|
62 b_fnclause :: "fnclause \<Rightarrow> atom set" and |
|
63 b_lrb8 :: "lrb8 \<Rightarrow> atom set" |
|
64 where |
|
65 "b_lrbs8 (Single' l) = b_lrb8 l" |
|
66 | "b_lrbs8 (More' l ls) = b_lrb8 l \<union> b_lrbs8 ls" |
|
67 | "b_pat (PVar'' x) = {atom x}" |
|
68 | "b_pat (PUnit'') = {}" |
|
69 | "b_pat (PPair'' p1 p2) = b_pat p1 \<union> b_pat p2" |
|
70 | "b_fnclauses (S' fc) = (b_fnclause fc)" |
|
71 | "b_fnclauses (ORs' fc fcs) = (b_fnclause fc) \<union> (b_fnclauses fcs)" |
|
72 | "b_lrb8 (Clause' fcs) = (b_fnclauses fcs)" |
|
73 | "b_fnclause (K' x pat exp8) = {atom x}" |
|
74 thm alpha_exp8_raw_alpha_fnclause_raw_alpha_fnclauses_raw_alpha_lrb8_raw_alpha_lrbs8_raw_alpha_pat8_raw.intros |
|
75 *) |
|
76 (* example 4 from Terms.thy *) |
11 (* example 4 from Terms.thy *) |
77 (* fv_eqvt does not work, we need to repaire defined permute functions |
12 (* fv_eqvt does not work, we need to repaire defined permute functions |
78 defined fv and defined alpha... *) |
13 defined fv and defined alpha... *) |
79 (* lists-datastructure does not work yet |
14 (* lists-datastructure does not work yet |
80 nominal_datatype trm4 = |
15 nominal_datatype trm4 = |