author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Tue, 23 Mar 2010 16:28:29 +0100 | |
changeset 1615 | 0ea578c6dae3 |
parent 1609 | c9bc3b61046c |
child 1621 | a40dbea68d0b |
permissions | -rw-r--r-- |
1601 | 1 |
theory ExCoreHaskell |
2 |
imports "Parser" |
|
3 |
begin |
|
4 |
||
5 |
(* core haskell *) |
|
6 |
||
7 |
ML {* val _ = recursive := false *} |
|
8 |
||
9 |
atom_decl var |
|
10 |
atom_decl tvar |
|
11 |
||
1606
75403378068b
Initial list unfoldings in Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1601
diff
changeset
|
12 |
(* there are types, coercion types and regular types list-data-structure *) |
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
13 |
|
1601 | 14 |
nominal_datatype tkind = |
15 |
KStar |
|
16 |
| KFun "tkind" "tkind" |
|
17 |
and ckind = |
|
18 |
CKEq "ty" "ty" |
|
19 |
and ty = |
|
20 |
TVar "tvar" |
|
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
21 |
| TC "char" |
1601 | 22 |
| TApp "ty" "ty" |
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
23 |
| TFun "char" "ty_lst" |
1601 | 24 |
| TAll tv::"tvar" "tkind" T::"ty" bind tv in T |
25 |
| TEq "ty" "ty" "ty" |
|
1606
75403378068b
Initial list unfoldings in Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1601
diff
changeset
|
26 |
and ty_lst = |
75403378068b
Initial list unfoldings in Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1601
diff
changeset
|
27 |
TsNil |
75403378068b
Initial list unfoldings in Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1601
diff
changeset
|
28 |
| TsCons "ty" "ty_lst" |
1601 | 29 |
and co = |
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
30 |
CC "char" |
1601 | 31 |
| CApp "co" "co" |
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
32 |
| CFun "char" "co_lst" |
1601 | 33 |
| CAll tv::"tvar" "ckind" C::"co" bind tv in C |
34 |
| CEq "co" "co" "co" |
|
35 |
| CSym "co" |
|
36 |
| CCir "co" "co" |
|
37 |
| CLeft "co" |
|
38 |
| CRight "co" |
|
39 |
| CSim "co" |
|
40 |
| CRightc "co" |
|
41 |
| CLeftc "co" |
|
42 |
| CCoe "co" "co" |
|
1606
75403378068b
Initial list unfoldings in Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1601
diff
changeset
|
43 |
and co_lst = |
75403378068b
Initial list unfoldings in Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1601
diff
changeset
|
44 |
CsNil |
75403378068b
Initial list unfoldings in Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1601
diff
changeset
|
45 |
| CsCons "co" "co_lst" |
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
46 |
and trm = |
1601 | 47 |
Var "var" |
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
48 |
| C "char" |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
49 |
| LAMT tv::"tvar" "tkind" t::"trm" bind tv in t |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
50 |
| LAMC tv::"tvar" "ckind" t::"trm" bind tv in t |
1601 | 51 |
| APP "trm" "ty" |
52 |
| Lam v::"var" "ty" t::"trm" bind v in t |
|
53 |
| App "trm" "trm" |
|
54 |
| Let x::"var" "ty" "trm" t::"trm" bind x in t |
|
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
55 |
| Case "trm" "assoc_lst" |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
56 |
| Cast "trm" "ty" --"ty is supposed to be a coercion type only" |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
57 |
and assoc_lst = |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
58 |
ANil |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
59 |
| ACons p::"pat" t::"trm" "assoc_lst" bind "bv p" in t |
1601 | 60 |
and pat = |
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
61 |
K "char" "tvtk_lst" "tvck_lst" "vt_lst" |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
62 |
and tvtk_lst = |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
63 |
TVTKNil |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
64 |
| TVTKCons "tvar" "tkind" "tvtk_lst" |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
65 |
and tvck_lst = |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
66 |
TVCKNil |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
67 |
| TVCKCons "tvar" "ckind" "tvck_lst" |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
68 |
and vt_lst = |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
69 |
VTNil |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
70 |
| VTCons "var" "ty" "vt_lst" |
1601 | 71 |
binder |
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
72 |
bv :: "pat \<Rightarrow> atom set" |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
73 |
(*and bv_vt :: "vt_lst \<Rightarrow> atom set" |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
74 |
and bv_tvtk :: "tvtk_lst \<Rightarrow> atom set" |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
75 |
and bv_tvck :: "tvck_lst \<Rightarrow> atom set"*) |
1601 | 76 |
where |
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
77 |
"bv (K s tvts tvcs vs) = {}" (*(bv_tvtk tvts) \<union> (bv_tvck tvcs) \<union> (bv_vt vs) *) |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
78 |
(*| "bv_vt VTNil = {}" |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
79 |
| "bv_vt (VTCons v k t) = {atom v} \<union> bv_vt t" |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
80 |
| "bv_tvtk TVTKNil = {}" |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
81 |
| "bv_tvtk (TVTKCons v k t) = {atom v} \<union> bv_tvtk t" |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
82 |
| "bv_tvck TVCKNil = {}" |
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1606
diff
changeset
|
83 |
| "bv_tvck (TVCKCons v k t) = {atom v} \<union> bv_tvck t"*) |
1601 | 84 |
|
1615
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
85 |
thm tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_tvtk_lst_tvck_lst_vt_lst.fv |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
86 |
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
87 |
|
1601 | 88 |
end |
89 |
||
90 |
||
91 |