| author | wenzelm | 
| Thu, 28 Sep 2000 14:48:05 +0200 | |
| changeset 10108 | 72a719e997b9 | 
| parent 6642 | 732af87c0650 | 
| child 12030 | 46d57d0290a2 | 
| permissions | -rw-r--r-- | 
| 
2445
 
51993fea433f
removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
 
oheimb 
parents: 
2267 
diff
changeset
 | 
1  | 
(* Title: HOLCF/domain/interface.ML  | 
| 
 
51993fea433f
removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
 
oheimb 
parents: 
2267 
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 
 
51993fea433f
removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
 
oheimb 
parents: 
2267 
diff
changeset
 | 
3  | 
Author : David von Oheimb  | 
| 
 
51993fea433f
removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
 
oheimb 
parents: 
2267 
diff
changeset
 | 
4  | 
Copyright 1995, 1996 TU Muenchen  | 
| 
 
51993fea433f
removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
 
oheimb 
parents: 
2267 
diff
changeset
 | 
5  | 
|
| 
 
51993fea433f
removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
 
oheimb 
parents: 
2267 
diff
changeset
 | 
6  | 
parser for domain section  | 
| 1274 | 7  | 
*)  | 
8  | 
||
9  | 
local  | 
|
10  | 
open ThyParse;  | 
|
11  | 
open Domain_Library;  | 
|
12  | 
||
| 4030 | 13  | 
(* --- generation of bindings for axioms and theorems in .thy.ML - *)  | 
| 1274 | 14  | 
|
| 4043 | 15  | 
fun get dname name = "get_thm thy " ^ quote (dname^"."^name);  | 
16  | 
fun gen_vals dname name = "val "^ name ^ " = get_thm" ^  | 
|
| 4709 | 17  | 
(if hd (rev (Symbol.explode name)) = "s" then "s" else "")^  | 
| 4043 | 18  | 
" thy " ^ quote (dname^"."^name)^";";  | 
19  | 
fun gen_vall name l = "val "^name^" = "^ mk_list l^";";  | 
|
20  | 
val rews = "iso_rews @ when_rews @\n\  | 
|
21  | 
\ con_rews @ sel_rews @ dis_rews @ dist_les @ dist_eqs @\n\  | 
|
22  | 
\ copy_rews";  | 
|
| 1274 | 23  | 
|
| 4043 | 24  | 
fun gen_struct num ((dname,_), cons') = let  | 
| 1274 | 25  | 
val axioms1 = ["abs_iso", "rep_iso", "when_def"];  | 
| 
1637
 
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
 
oheimb 
parents: 
1461 
diff
changeset
 | 
26  | 
(* con_defs , sel_defs, dis_defs *)  | 
| 4043 | 27  | 
val axioms2 = ["copy_def", "reach", "take_def", "finite_def"];  | 
28  | 
val theorems = ["iso_rews", "exhaust", "casedist", "when_rews", "con_rews",  | 
|
29  | 
"sel_rews", "dis_rews", "dist_les", "dist_eqs", "inverts",  | 
|
30  | 
"injects", "copy_rews"];  | 
|
| 1274 | 31  | 
in  | 
32  | 
"structure "^dname^" = struct\n"^  | 
|
| 4043 | 33  | 
cat_lines (map (gen_vals dname) axioms1)^"\n"^  | 
34  | 
gen_vall "con_defs" (map (fn (con,_,_) =>  | 
|
35  | 
get dname (strip_esc con^"_def")) cons')^"\n"^  | 
|
36  | 
gen_vall "sel_defs" (flat (map (fn (_,_,args) => map (fn (_,sel,_) =>  | 
|
37  | 
get dname (sel^"_def")) args) cons'))^"\n"^  | 
|
38  | 
gen_vall "dis_defs" (map (fn (con,_,_) =>  | 
|
39  | 
get dname (dis_name_ con^"_def")) cons')^"\n"^  | 
|
40  | 
cat_lines (map (gen_vals dname) axioms2)^"\n"^  | 
|
41  | 
cat_lines (map (gen_vals dname) theorems)^"\n"^  | 
|
42  | 
(if num > 1 then "val rews = " ^rews ^";\n" else "")  | 
|
| 1274 | 43  | 
end;  | 
44  | 
||
| 4008 | 45  | 
fun mk_domain eqs'' = let  | 
46  | 
val num = length eqs'';  | 
|
47  | 
val dtnvs = map fst eqs'';  | 
|
| 1274 | 48  | 
val dnames = map fst dtnvs;  | 
49  | 
val comp_dname = implode (separate "_" dnames);  | 
|
| 2267 | 50  | 
val conss' = ListPair.map  | 
| 4043 | 51  | 
(fn (dnam,cons'') => let fun sel n m = upd_second  | 
52  | 
(fn None => dnam^"_sel_"^(string_of_int n)^  | 
|
53  | 
"_"^(string_of_int m)  | 
|
54  | 
| Some s => s)  | 
|
| 2267 | 55  | 
fun fill_sels n con = upd_third (mapn (sel n) 1) con;  | 
56  | 
in mapn fill_sels 1 cons'' end)  | 
|
| 4008 | 57  | 
(dnames, map snd eqs'');  | 
| 1274 | 58  | 
val eqs' = dtnvs~~conss';  | 
59  | 
||
| 
1637
 
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
 
oheimb 
parents: 
1461 
diff
changeset
 | 
60  | 
(* ----- string for calling add_domain -------------------------------------- *)  | 
| 1284 | 61  | 
|
| 
1637
 
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
 
oheimb 
parents: 
1461 
diff
changeset
 | 
62  | 
val thy_ext_string = let  | 
| 2240 | 63  | 
fun mk_conslist cons' =  | 
| 4030 | 64  | 
mk_list (map (fn (c,syn,ts)=> "\n"^  | 
| 4008 | 65  | 
mk_triple(quote c, syn, mk_list (map (fn (b,s ,tp) =>  | 
66  | 
mk_triple(Bool.toString b, quote s, tp)) ts))) cons');  | 
|
| 4043 | 67  | 
in "val thy = thy |> Domain_Extender.add_domain "  | 
| 4030 | 68  | 
^ mk_pair(quote comp_dname, mk_list(map (fn ((n,vs),cs) => "\n"^  | 
| 4008 | 69  | 
mk_pair (mk_pair (quote n, mk_list vs),  | 
70  | 
mk_conslist cs)) eqs'))  | 
|
| 4030 | 71  | 
^ ";\n"  | 
| 
1637
 
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
 
oheimb 
parents: 
1461 
diff
changeset
 | 
72  | 
end;  | 
| 
 
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
 
oheimb 
parents: 
1461 
diff
changeset
 | 
73  | 
|
| 4122 | 74  | 
(* ----- string for producing the structures -------------------------------- *)  | 
| 
1637
 
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
 
oheimb 
parents: 
1461 
diff
changeset
 | 
75  | 
|
| 
 
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
 
oheimb 
parents: 
1461 
diff
changeset
 | 
76  | 
val val_gen_string = let  | 
| 
 
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
 
oheimb 
parents: 
1461 
diff
changeset
 | 
77  | 
val comp_axioms = [(* copy, *) "reach", "take_def", "finite_def"  | 
| 
 
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
 
oheimb 
parents: 
1461 
diff
changeset
 | 
78  | 
(*, "bisim_def" *)];  | 
| 4043 | 79  | 
val comp_theorems = ["take_rews", "take_lemmas", "finites",  | 
80  | 
"finite_ind", "ind", "coind"];  | 
|
| 
1637
 
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
 
oheimb 
parents: 
1461 
diff
changeset
 | 
81  | 
fun collect sep name = if num = 1 then name else  | 
| 4043 | 82  | 
implode (separate sep (map (fn s => s^"."^name) dnames));  | 
| 1274 | 83  | 
in  | 
| 4030 | 84  | 
implode (separate "end; (* struct *)\n"  | 
| 4043 | 85  | 
(map (gen_struct num) eqs')) ^  | 
86  | 
(if num > 1 then "end; (* struct *)\n\  | 
|
87  | 
\structure "^comp_dname^" = struct\n" ^  | 
|
88  | 
gen_vals comp_dname "copy_def" ^"\n" ^  | 
|
89  | 
cat_lines (map (fn name => gen_vall (name^"s")  | 
|
90  | 
(map (fn dn => dn^"."^name) dnames)) comp_axioms)^"\n"  | 
|
91  | 
else "") ^  | 
|
92  | 
gen_vals comp_dname "bisim_def" ^"\n"^  | 
|
93  | 
cat_lines (map (gen_vals comp_dname) comp_theorems)^"\n"^  | 
|
94  | 
"val rews = "^(if num>1 then collect" @ " "rews"else rews)^  | 
|
| 4030 | 95  | 
" @ take_rews;\n\  | 
96  | 
\end; (* struct *)\n"  | 
|
| 
1637
 
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
 
oheimb 
parents: 
1461 
diff
changeset
 | 
97  | 
end;  | 
| 4043 | 98  | 
in (thy_ext_string^  | 
| 4030 | 99  | 
val_gen_string^  | 
100  | 
"val thy = thy",  | 
|
101  | 
"") end;  | 
|
| 
1637
 
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
 
oheimb 
parents: 
1461 
diff
changeset
 | 
102  | 
|
| 4030 | 103  | 
(* ----- parser for domain declaration equation ----------------------------- *)  | 
| 1274 | 104  | 
|
| 6642 | 105  | 
val name' = name >> unenclose;  | 
106  | 
val type_var' = type_var >> unenclose;  | 
|
| 4008 | 107  | 
fun esc_quote s = let fun esc [] = []  | 
108  | 
			|   esc ("\""::xs) = esc xs
 | 
|
109  | 
			|   esc ("[" ::xs) = "{"::esc xs
 | 
|
110  | 
			|   esc ("]" ::xs) = "}"::esc xs
 | 
|
111  | 
| esc (x ::xs) = x ::esc xs  | 
|
| 4709 | 112  | 
in implode (esc (Symbol.explode s)) end;  | 
| 4122 | 113  | 
val tvar = (type_var' ^^  | 
114  | 
optional ($$ "::" ^^ (sort >> esc_quote)) "") >> quote;  | 
|
| 4008 | 115  | 
fun type_args l = (tvar >> (fn x => [x])  | 
116  | 
                 ||  "(" $$-- !! (list1 tvar --$$ ")")
 | 
|
117  | 
|| empty >> K []) l  | 
|
118  | 
fun con_typ l = (type_args -- name' >> (fn (x,y) => (y,x))) l  | 
|
119  | 
(* here ident may be used instead of name'  | 
|
120  | 
to avoid ambiguity with standard mixfix option *)  | 
|
| 1274 | 121  | 
  val domain_arg  = "(" $$-- (optional ($$ "lazy" >> K true) false)
 | 
| 
2445
 
51993fea433f
removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
 
oheimb 
parents: 
2267 
diff
changeset
 | 
122  | 
-- (optional ((name' >> Some) --$$ "::") None)  | 
| 4122 | 123  | 
-- typ --$$ ")"  | 
124  | 
>> (fn ((lazy,sel),tp) => (lazy,sel,tp))  | 
|
| 4008 | 125  | 
|| typ >> (fn x => (false,None,x))  | 
| 
2445
 
51993fea433f
removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
 
oheimb 
parents: 
2267 
diff
changeset
 | 
126  | 
val domain_cons = name' -- !! (repeat domain_arg)  | 
| 4122 | 127  | 
-- ThyParse.opt_mixfix  | 
| 
1637
 
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
 
oheimb 
parents: 
1461 
diff
changeset
 | 
128  | 
>> (fn ((con,args),syn) => (con,syn,args));  | 
| 1274 | 129  | 
in  | 
| 4043 | 130  | 
val domain_decl = (enum1 "and" (con_typ --$$ "=" -- !!  | 
131  | 
(enum1 "|" domain_cons))) >> mk_domain;  | 
|
| 3622 | 132  | 
val _ = ThySyn.add_syntax  | 
| 4122 | 133  | 
["lazy", "and"]  | 
134  | 
    [("domain", domain_decl)]
 | 
|
| 1274 | 135  | 
|
| 3622 | 136  | 
end; (* local *)  |