| author | paulson | 
| Mon, 26 May 1997 12:29:55 +0200 | |
| changeset 3333 | 0bbf06e86c06 | 
| parent 3308 | da002cef7090 | 
| child 3345 | 4d888ddd6284 | 
| permissions | -rw-r--r-- | 
| 923 | 1  | 
(* Title: HOL/thy_syntax.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Markus Wenzel and Lawrence C Paulson and Carsten Clasohm  | 
|
4  | 
||
5  | 
Additional theory file sections for HOL.  | 
|
6  | 
||
7  | 
TODO:  | 
|
8  | 
move datatype / primrec stuff to pre_datatype.ML (?)  | 
|
9  | 
*)  | 
|
10  | 
||
11  | 
(*the kind of distinctiveness axioms depends on number of constructors*)  | 
|
| 2930 | 12  | 
val dtK = 7; (* FIXME rename?, move? *)  | 
| 923 | 13  | 
|
14  | 
structure ThySynData: THY_SYN_DATA =  | 
|
15  | 
struct  | 
|
16  | 
||
17  | 
open ThyParse;  | 
|
18  | 
||
19  | 
||
| 1475 | 20  | 
(** typedef **)  | 
| 923 | 21  | 
|
| 1475 | 22  | 
fun mk_typedef_decl (((((opt_name, vs), t), mx), rhs), wt) =  | 
| 923 | 23  | 
let  | 
24  | 
val name' = if_none opt_name t;  | 
|
25  | 
val name = strip_quotes name';  | 
|
26  | 
in  | 
|
27  | 
(cat_lines [name', mk_triple (t, mk_list vs, mx), rhs, wt],  | 
|
28  | 
[name ^ "_def", "Rep_" ^ name, "Rep_" ^ name ^ "_inverse",  | 
|
29  | 
"Abs_" ^ name ^ "_inverse"])  | 
|
30  | 
end;  | 
|
31  | 
||
| 1475 | 32  | 
val typedef_decl =  | 
| 923 | 33  | 
  optional ("(" $$-- name --$$ ")" >> Some) None --
 | 
34  | 
type_args -- name -- opt_infix --$$ "=" -- string -- opt_witness  | 
|
| 1475 | 35  | 
>> mk_typedef_decl;  | 
| 923 | 36  | 
|
37  | 
||
38  | 
||
39  | 
(** (co)inductive **)  | 
|
40  | 
||
41  | 
(*co is either "" or "Co"*)  | 
|
42  | 
fun inductive_decl co =  | 
|
43  | 
let  | 
|
44  | 
fun mk_intr_name (s, _) = (*the "op" cancels any infix status*)  | 
|
45  | 
if Syntax.is_identifier s then "op " ^ s else "_";  | 
|
46  | 
fun mk_params (((recs, ipairs), monos), con_defs) =  | 
|
47  | 
let val big_rec_name = space_implode "_" (map (scan_to_id o trim) recs)  | 
|
48  | 
and srec_tms = mk_list recs  | 
|
49  | 
and sintrs = mk_big_list (map snd ipairs)  | 
|
| 3194 | 50  | 
val intrnl_name = big_rec_name ^ "_Intrnl"  | 
| 923 | 51  | 
in  | 
52  | 
         (";\n\n\
 | 
|
| 3194 | 53  | 
\structure " ^ intrnl_name ^ " =\n\  | 
| 923 | 54  | 
\ struct\n\  | 
55  | 
\ val _ = writeln \"" ^ co ^  | 
|
56  | 
"Inductive definition " ^ big_rec_name ^ "\"\n\  | 
|
| 
1430
 
439e1476a7f8
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
1316 
diff
changeset
 | 
57  | 
\ val rec_tms\t= map (readtm (sign_of thy) Ind_Syntax.termTVar) "  | 
| 923 | 58  | 
^ srec_tms ^ "\n\  | 
59  | 
\ and intr_tms\t= map (readtm (sign_of thy) propT)\n"  | 
|
60  | 
^ sintrs ^ "\n\  | 
|
| 
1430
 
439e1476a7f8
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
1316 
diff
changeset
 | 
61  | 
\ end;\n\n\  | 
| 923 | 62  | 
          \val thy = thy |> " ^ co ^ "Ind.add_fp_def_i \n    (" ^
 | 
| 3194 | 63  | 
intrnl_name ^ ".rec_tms, " ^  | 
64  | 
intrnl_name ^ ".intr_tms)"  | 
|
| 923 | 65  | 
,  | 
66  | 
"structure " ^ big_rec_name ^ " =\n\  | 
|
| 
1430
 
439e1476a7f8
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
1316 
diff
changeset
 | 
67  | 
\ let\n\  | 
| 
 
439e1476a7f8
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
1316 
diff
changeset
 | 
68  | 
\ val _ = writeln \"Proofs for " ^ co ^  | 
| 
 
439e1476a7f8
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
1316 
diff
changeset
 | 
69  | 
"Inductive definition " ^ big_rec_name ^ "\"\n\  | 
| 923 | 70  | 
\ structure Result = " ^ co ^ "Ind_section_Fun\n\  | 
| 3194 | 71  | 
\\t (open " ^ intrnl_name ^ "\n\  | 
| 
1430
 
439e1476a7f8
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
1316 
diff
changeset
 | 
72  | 
\\t val thy\t\t= thy\n\  | 
| 
 
439e1476a7f8
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
1316 
diff
changeset
 | 
73  | 
\\t val monos\t\t= " ^ monos ^ "\n\  | 
| 
 
439e1476a7f8
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
1316 
diff
changeset
 | 
74  | 
\\t val con_defs\t\t= " ^ con_defs ^ ");\n\n\  | 
| 1465 | 75  | 
\ in\n\  | 
| 
1430
 
439e1476a7f8
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
1316 
diff
changeset
 | 
76  | 
\ struct\n\  | 
| 923 | 77  | 
\ val " ^ mk_list (map mk_intr_name ipairs) ^ " = Result.intrs;\n\  | 
78  | 
\ open Result\n\  | 
|
| 
1430
 
439e1476a7f8
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
1316 
diff
changeset
 | 
79  | 
\ end\n\  | 
| 
 
439e1476a7f8
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
1316 
diff
changeset
 | 
80  | 
\ end;\n\n\  | 
| 3194 | 81  | 
\structure " ^ intrnl_name ^ " = struct end;\n\n"  | 
| 923 | 82  | 
)  | 
83  | 
end  | 
|
84  | 
val ipairs = "intrs" $$-- repeat1 (ident -- !! string)  | 
|
85  | 
fun optstring s = optional (s $$-- string) "\"[]\"" >> trim  | 
|
86  | 
in  | 
|
| 1788 | 87  | 
repeat1 name -- ipairs -- optstring "monos" -- optstring "con_defs"  | 
| 923 | 88  | 
>> mk_params  | 
89  | 
end;  | 
|
90  | 
||
91  | 
||
92  | 
||
93  | 
(** datatype **)  | 
|
94  | 
||
95  | 
local  | 
|
96  | 
(* FIXME err -> add_datatype *)  | 
|
97  | 
fun mk_cons cs =  | 
|
98  | 
(case duplicates (map (fst o fst) cs) of  | 
|
99  | 
[] => map (fn ((s, ts), syn) => mk_triple (s, mk_list ts, syn)) cs  | 
|
100  | 
    | dups => error ("Duplicate constructors: " ^ commas_quote dups));
 | 
|
101  | 
||
102  | 
(*generate names of distinctiveness axioms*)  | 
|
103  | 
fun mk_distinct_rules cs tname =  | 
|
104  | 
let  | 
|
105  | 
val uqcs = map (fn ((s, _), _) => strip_quotes s) cs;  | 
|
106  | 
(*combine all constructor names with all others w/o duplicates*)  | 
|
107  | 
fun neg_one c = map (fn c2 => quote (c ^ "_not_" ^ c2));  | 
|
108  | 
fun neg1 [] = []  | 
|
109  | 
| neg1 (c1 :: cs) = neg_one c1 cs @ neg1 cs;  | 
|
110  | 
in  | 
|
111  | 
if length uqcs < dtK then neg1 uqcs  | 
|
112  | 
else quote (tname ^ "_ord_distinct") ::  | 
|
113  | 
map (fn c => quote (tname ^ "_ord_" ^ c)) uqcs  | 
|
114  | 
end;  | 
|
115  | 
||
116  | 
fun mk_rules tname cons pre = " map (get_axiom thy) " ^  | 
|
| 3194 | 117  | 
mk_list (map (fn ((s,_), _) => quote (tname ^ pre ^ strip_quotes s)) cons);  | 
| 923 | 118  | 
|
| 1668 | 119  | 
(*generate string for calling add_datatype and build_record*)  | 
| 923 | 120  | 
fun mk_params ((ts, tname), cons) =  | 
| 
3308
 
da002cef7090
Added overloaded function `size' for all datatypes.
 
nipkow 
parents: 
3194 
diff
changeset
 | 
121  | 
   ("val (thy," ^ tname ^ "_add_primrec," ^ tname ^ "_size_eqns) = Datatype.add_datatype\n"
 | 
| 923 | 122  | 
^ mk_triple (mk_list ts, quote tname, mk_list (mk_cons cons)) ^ " thy\n\  | 
| 
3308
 
da002cef7090
Added overloaded function `size' for all datatypes.
 
nipkow 
parents: 
3194 
diff
changeset
 | 
123  | 
    \val thy = ("^tname^"_add_primrec "^tname^"_size_eqns thy)"
 | 
| 2922 | 124  | 
,  | 
| 923 | 125  | 
"structure " ^ tname ^ " =\n\  | 
126  | 
\struct\n\  | 
|
127  | 
\ val inject = map (get_axiom thy) " ^  | 
|
128  | 
        mk_list (map (fn ((s, _), _) => quote ("inject_" ^ strip_quotes s))
 | 
|
129  | 
(filter_out (null o snd o fst) cons)) ^ ";\n\  | 
|
130  | 
\ val distinct = " ^  | 
|
131  | 
(if length cons < dtK then "let val distinct' = " else "") ^  | 
|
132  | 
"map (get_axiom thy) " ^ mk_list (mk_distinct_rules cons tname) ^  | 
|
133  | 
(if length cons < dtK then  | 
|
134  | 
" in distinct' @ (map (fn t => sym COMP (t RS contrapos))\  | 
|
135  | 
\ distinct') end"  | 
|
136  | 
else "") ^ ";\n\  | 
|
137  | 
\ val induct = get_axiom thy \"" ^ tname ^ "_induct\";\n\  | 
|
138  | 
\ val cases =" ^ mk_rules tname cons "_case_" ^ ";\n\  | 
|
139  | 
\ val recs =" ^ mk_rules tname cons "_rec_" ^ ";\n\  | 
|
140  | 
\ val simps = inject @ distinct @ cases @ recs;\n\  | 
|
141  | 
    \ fun induct_tac a = res_inst_tac [(" ^ quote tname ^ ", a)] induct;\n\
 | 
|
| 
1264
 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 
clasohm 
parents: 
1251 
diff
changeset
 | 
142  | 
\end;\n\  | 
| 1668 | 143  | 
\val dummy = datatypes := Dtype.build_record (thy, " ^  | 
144  | 
mk_pair (quote tname, mk_list (map (fst o fst) cons)) ^  | 
|
145  | 
", " ^ tname ^ ".induct_tac) :: (!datatypes);\n\  | 
|
| 2930 | 146  | 
    \val dummy = Addsimps(" ^ tname ^ ".cases @ " ^ tname ^ ".recs);\n\
 | 
147  | 
\val dummy = AddIffs " ^ tname ^ ".inject;\n\  | 
|
148  | 
\val dummy = " ^  | 
|
149  | 
(if length cons < dtK then "AddIffs " else "Addsimps ") ^  | 
|
| 
3308
 
da002cef7090
Added overloaded function `size' for all datatypes.
 
nipkow 
parents: 
3194 
diff
changeset
 | 
150  | 
tname ^ ".distinct;\n\  | 
| 
 
da002cef7090
Added overloaded function `size' for all datatypes.
 
nipkow 
parents: 
3194 
diff
changeset
 | 
151  | 
\val dummy = Addsimps(map (fn (_,eqn) =>\n\  | 
| 
 
da002cef7090
Added overloaded function `size' for all datatypes.
 
nipkow 
parents: 
3194 
diff
changeset
 | 
152  | 
    \ prove_goalw thy [get_def thy " ^ quote("size_"^tname) ^
 | 
| 
 
da002cef7090
Added overloaded function `size' for all datatypes.
 
nipkow 
parents: 
3194 
diff
changeset
 | 
153  | 
"] eqn (fn _ => [Simp_tac 1]))\n" ^  | 
| 
 
da002cef7090
Added overloaded function `size' for all datatypes.
 
nipkow 
parents: 
3194 
diff
changeset
 | 
154  | 
tname^"_size_eqns)\n"  | 
| 
 
da002cef7090
Added overloaded function `size' for all datatypes.
 
nipkow 
parents: 
3194 
diff
changeset
 | 
155  | 
);  | 
| 923 | 156  | 
|
157  | 
(*parsers*)  | 
|
158  | 
val tvars = type_args >> map (cat "dtVar");  | 
|
| 1316 | 159  | 
|
160  | 
val simple_typ = ident >> (cat "dtTyp" o curry mk_pair "[]" o quote) ||  | 
|
161  | 
type_var >> cat "dtVar";  | 
|
162  | 
||
| 
1251
 
81fc4d8e3eda
added nested types on right hand side of datatype definitions
 
clasohm 
parents: 
977 
diff
changeset
 | 
163  | 
fun complex_typ toks =  | 
| 1316 | 164  | 
    let val typ = simple_typ || "(" $$-- complex_typ --$$ ")";
 | 
165  | 
        val typ2 = complex_typ || "(" $$-- complex_typ --$$ ")";
 | 
|
166  | 
in  | 
|
167  | 
(typ -- repeat (ident>>quote) >>  | 
|
168  | 
(foldl (fn (x,y) => "dtTyp " ^ mk_pair (brackets x, y))) ||  | 
|
169  | 
      "(" $$-- !! (list1 typ2) --$$ ")" -- !! (repeat1 (ident>>quote)) >>
 | 
|
170  | 
(fn (fst, ids) => foldl (fn (x,y) => "dtTyp " ^  | 
|
171  | 
mk_pair (brackets x, y)) (commas fst, ids))) toks  | 
|
172  | 
end;  | 
|
173  | 
||
| 
977
 
5d57287e5e1e
changed syntax of datatype declarations (curried types for constructor
 
clasohm 
parents: 
923 
diff
changeset
 | 
174  | 
  val opt_typs = repeat (simple_typ || ("(" $$-- complex_typ --$$ ")"));
 | 
| 923 | 175  | 
val constructor = name -- opt_typs -- opt_mixfix;  | 
176  | 
in  | 
|
177  | 
val datatype_decl =  | 
|
178  | 
tvars -- ident --$$ "=" -- enum1 "|" constructor >> mk_params;  | 
|
179  | 
end;  | 
|
180  | 
||
181  | 
||
182  | 
||
183  | 
(** primrec **)  | 
|
184  | 
||
| 2922 | 185  | 
(*recursion equations have user-supplied names*)  | 
| 1845 | 186  | 
fun mk_primrec_decl_1 ((fname, tname), axms) =  | 
| 923 | 187  | 
let  | 
| 
1574
 
5a63ab90ee8a
modified primrec so it can be used in MiniML/Type.thy
 
clasohm 
parents: 
1475 
diff
changeset
 | 
188  | 
(*Isolate type name from the structure's identifier it may be stored in*)  | 
| 
 
5a63ab90ee8a
modified primrec so it can be used in MiniML/Type.thy
 
clasohm 
parents: 
1475 
diff
changeset
 | 
189  | 
val tname' = implode (snd (take_suffix (not_equal ".") (explode tname)));  | 
| 
 
5a63ab90ee8a
modified primrec so it can be used in MiniML/Type.thy
 
clasohm 
parents: 
1475 
diff
changeset
 | 
190  | 
|
| 923 | 191  | 
fun mk_prove (name, eqn) =  | 
| 
1264
 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 
clasohm 
parents: 
1251 
diff
changeset
 | 
192  | 
      "val " ^ name ^ " = store_thm (" ^ quote name
 | 
| 
1574
 
5a63ab90ee8a
modified primrec so it can be used in MiniML/Type.thy
 
clasohm 
parents: 
1475 
diff
changeset
 | 
193  | 
^ ", prove_goalw thy [get_def thy "  | 
| 
 
5a63ab90ee8a
modified primrec so it can be used in MiniML/Type.thy
 
clasohm 
parents: 
1475 
diff
changeset
 | 
194  | 
^ (quote (strip_quotes fname ^ "_" ^ tname')) ^ "] " ^ eqn ^ "\n\  | 
| 
1264
 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 
clasohm 
parents: 
1251 
diff
changeset
 | 
195  | 
\ (fn _ => [Simp_tac 1]));";  | 
| 
 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 
clasohm 
parents: 
1251 
diff
changeset
 | 
196  | 
|
| 923 | 197  | 
val axs = mk_list (map (fn (n, a) => mk_pair (quote n, a)) axms);  | 
| 2922 | 198  | 
  in ("|> " ^ tname ^ "_add_primrec " ^ axs
 | 
199  | 
,  | 
|
200  | 
cat_lines (map mk_prove axms)  | 
|
| 
1264
 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 
clasohm 
parents: 
1251 
diff
changeset
 | 
201  | 
^ "\nval dummy = Addsimps " ^ mk_list (map fst axms) ^ ";")  | 
| 
 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 
clasohm 
parents: 
1251 
diff
changeset
 | 
202  | 
end;  | 
| 923 | 203  | 
|
| 2922 | 204  | 
(*recursion equations have no names*)  | 
| 1845 | 205  | 
fun mk_primrec_decl_2 ((fname, tname), axms) =  | 
206  | 
let  | 
|
207  | 
(*Isolate type name from the structure's identifier it may be stored in*)  | 
|
208  | 
val tname' = implode (snd (take_suffix (not_equal ".") (explode tname)));  | 
|
209  | 
||
210  | 
fun mk_prove eqn =  | 
|
211  | 
"prove_goalw thy [get_def thy "  | 
|
212  | 
^ (quote (strip_quotes fname ^ "_" ^ tname')) ^ "] " ^ eqn ^ " \  | 
|
213  | 
\(fn _ => [Simp_tac 1])";  | 
|
214  | 
||
215  | 
    val axs = mk_list (map (fn a => mk_pair ("\"\"", a)) axms);
 | 
|
| 2922 | 216  | 
  in ("|> " ^ tname ^ "_add_primrec " ^ axs
 | 
217  | 
,  | 
|
| 1845 | 218  | 
"val dummy = Addsimps " ^  | 
219  | 
brackets(space_implode ",\n" (map mk_prove axms)) ^ ";")  | 
|
220  | 
end;  | 
|
221  | 
||
| 2922 | 222  | 
(*function name, argument type and either (name,axiom) pairs or just axioms*)  | 
| 923 | 223  | 
val primrec_decl =  | 
| 1845 | 224  | 
(name -- long_id -- repeat1 (ident -- string) >> mk_primrec_decl_1) ||  | 
225  | 
(name -- long_id -- repeat1 string >> mk_primrec_decl_2) ;  | 
|
| 923 | 226  | 
|
227  | 
||
228  | 
||
| 2922 | 229  | 
|
230  | 
(** rec: interface to Slind's TFL **)  | 
|
231  | 
||
232  | 
||
| 3194 | 233  | 
(*fname: name of function being defined; rel: well-founded relation*)  | 
| 2922 | 234  | 
fun mk_rec_decl ((fname, rel), axms) =  | 
235  | 
let val fid = trim fname  | 
|
| 3194 | 236  | 
val intrnl_name = fid ^ "_Intrnl"  | 
| 2922 | 237  | 
in  | 
238  | 
	 (";\n\n\
 | 
|
| 3194 | 239  | 
\val _ = writeln \"Recursive function " ^ fid ^ "\"\n\  | 
240  | 
\val (thy, pats_" ^ intrnl_name ^ ") = Tfl.define thy " ^  | 
|
241  | 
rel ^ "\n" ^ mk_big_list axms ^ ";\n\  | 
|
| 2922 | 242  | 
\val thy = thy"  | 
243  | 
,  | 
|
| 3194 | 244  | 
"structure " ^ fid ^ " =\n\  | 
245  | 
\ struct\n\  | 
|
246  | 
\ val _ = writeln \"Proofs for recursive function " ^ fid ^ "\"\n\  | 
|
247  | 
          \  val {rules, induct, tcs} = \n\
 | 
|
248  | 
\ \t Tfl.simplify_defn(thy, (\"" ^ fid ^  | 
|
249  | 
"\", pats_" ^ intrnl_name ^ "))\n\  | 
|
250  | 
\ end;\n\  | 
|
251  | 
\val pats_" ^ intrnl_name ^ " = ();\n")  | 
|
| 2922 | 252  | 
end;  | 
253  | 
||
254  | 
val rec_decl = (name -- string -- repeat1 string >> mk_rec_decl) ;  | 
|
255  | 
||
256  | 
||
257  | 
||
258  | 
||
259  | 
||
| 923 | 260  | 
(** sections **)  | 
261  | 
||
262  | 
val user_keywords = ["intrs", "monos", "con_defs", "|"];  | 
|
263  | 
||
264  | 
val user_sections =  | 
|
| 1475 | 265  | 
[axm_section "typedef" "|> Typedef.add_typedef" typedef_decl,  | 
| 923 | 266  | 
  ("inductive", inductive_decl ""),
 | 
267  | 
  ("coinductive", inductive_decl "Co"),
 | 
|
268  | 
  ("datatype", datatype_decl),
 | 
|
| 2922 | 269  | 
  ("primrec", primrec_decl),
 | 
270  | 
  ("recdef", rec_decl)];
 | 
|
| 923 | 271  | 
|
272  | 
||
273  | 
end;  | 
|
274  | 
||
275  | 
||
276  | 
structure ThySyn = ThySynFun(ThySynData);  | 
|
277  | 
init_thy_reader ();  | 
|
278  |