| author | paulson | 
| Fri, 23 Dec 2005 17:34:46 +0100 | |
| changeset 18508 | c5861e128a95 | 
| parent 17926 | 8e12d3a4b890 | 
| child 18678 | dd0c569fa43d | 
| permissions | -rw-r--r-- | 
| 
4129
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOLCF/cont_consts.ML  | 
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 12625 | 3  | 
Author: Tobias Mayr, David von Oheimb, and Markus Wenzel  | 
| 
4129
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
4  | 
|
| 
14682
 
a5072752114c
HOLCF: discontinued special version of 'constdefs';
 
wenzelm 
parents: 
12876 
diff
changeset
 | 
5  | 
HOLCF version of consts: handle continuous function types in mixfix  | 
| 
 
a5072752114c
HOLCF: discontinued special version of 'constdefs';
 
wenzelm 
parents: 
12876 
diff
changeset
 | 
6  | 
syntax.  | 
| 
4129
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
7  | 
*)  | 
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
|
| 12625 | 9  | 
signature CONT_CONSTS =  | 
10  | 
sig  | 
|
11  | 
val add_consts: (bstring * string * mixfix) list -> theory -> theory  | 
|
12  | 
val add_consts_i: (bstring * typ * mixfix) list -> theory -> theory  | 
|
13  | 
end;  | 
|
14  | 
||
15  | 
structure ContConsts: CONT_CONSTS =  | 
|
| 
4129
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
16  | 
struct  | 
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
17  | 
|
| 12625 | 18  | 
|
19  | 
(* misc utils *)  | 
|
| 
4129
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
20  | 
|
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
21  | 
open HOLCFLogic;  | 
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
22  | 
|
| 12625 | 23  | 
fun first (x,_,_) = x;  | 
24  | 
fun second (_,x,_) = x;  | 
|
25  | 
fun third (_,_,x) = x;  | 
|
| 
4129
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
26  | 
fun upd_first f (x,y,z) = (f x, y, z);  | 
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
27  | 
fun upd_second f (x,y,z) = ( x, f y, z);  | 
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
28  | 
fun upd_third f (x,y,z) = ( x, y, f z);  | 
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
29  | 
|
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
30  | 
fun change_arrow 0 T = T  | 
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
31  | 
|   change_arrow n (Type(_,[S,T])) = Type ("fun",[S,change_arrow (n-1) T])
 | 
| 12625 | 32  | 
| change_arrow _ _ = sys_error "cont_consts: change_arrow";  | 
| 
4129
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
33  | 
|
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
34  | 
fun trans_rules name2 name1 n mx = let  | 
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
35  | 
fun argnames _ 0 = []  | 
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
36  | 
| argnames c n = chr c::argnames (c+1) (n-1);  | 
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
37  | 
val vnames = argnames (ord "A") n;  | 
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
38  | 
val extra_parse_rule = Syntax.ParseRule (Constant name2, Constant name1);  | 
| 5700 | 39  | 
in [Syntax.ParsePrintRule (Syntax.mk_appl (Constant name2) (map Variable vnames),  | 
| 15570 | 40  | 
Library.foldl (fn (t,arg) => (Syntax.mk_appl (Constant "Rep_CFun")  | 
| 12625 | 41  | 
[t,Variable arg]))  | 
42  | 
(Constant name1,vnames))]  | 
|
| 11651 | 43  | 
@(case mx of InfixName _ => [extra_parse_rule]  | 
44  | 
| InfixlName _ => [extra_parse_rule]  | 
|
| 12625 | 45  | 
| InfixrName _ => [extra_parse_rule] | _ => []) end;  | 
| 
4129
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
46  | 
|
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
47  | 
|
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
48  | 
(* transforming infix/mixfix declarations of constants with type ...->...  | 
| 12625 | 49  | 
a declaration of such a constant is transformed to a normal declaration with  | 
| 
4129
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
50  | 
an internal name, the same type, and nofix. Additionally, a purely syntactic  | 
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
51  | 
declaration with the original name, type ...=>..., and the original mixfix  | 
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
52  | 
is generated and connected to the other declaration via some translation.  | 
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
53  | 
*)  | 
| 12625 | 54  | 
fun fix_mixfix (syn , T, mx as Infix p ) =  | 
55  | 
(Syntax.const_name syn mx, T, InfixName (syn, p))  | 
|
56  | 
| fix_mixfix (syn , T, mx as Infixl p ) =  | 
|
57  | 
(Syntax.const_name syn mx, T, InfixlName (syn, p))  | 
|
58  | 
| fix_mixfix (syn , T, mx as Infixr p ) =  | 
|
59  | 
(Syntax.const_name syn mx, T, InfixrName (syn, p))  | 
|
| 
4129
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
60  | 
| fix_mixfix decl = decl;  | 
| 12625 | 61  | 
fun transform decl = let  | 
62  | 
val (c, T, mx) = fix_mixfix decl;  | 
|
63  | 
val c2 = "_cont_" ^ c;  | 
|
64  | 
val n = Syntax.mixfix_args mx  | 
|
65  | 
in ((c , T,NoSyn),  | 
|
66  | 
(c2,change_arrow n T,mx ),  | 
|
67  | 
trans_rules c2 c n mx) end;  | 
|
| 
4129
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
68  | 
|
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
69  | 
fun cfun_arity (Type(n,[_,T])) = if n = cfun_arrow then 1+cfun_arity T else 0  | 
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
70  | 
| cfun_arity _ = 0;  | 
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
71  | 
|
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
72  | 
fun is_contconst (_,_,NoSyn ) = false  | 
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
73  | 
| is_contconst (_,_,Binder _) = false  | 
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
74  | 
| is_contconst (c,T,mx ) = cfun_arity T >= Syntax.mixfix_args mx  | 
| 12625 | 75  | 
                         handle ERROR => error ("in mixfix annotation for " ^
 | 
76  | 
quote (Syntax.const_name c mx));  | 
|
77  | 
||
| 
4129
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
78  | 
|
| 12625 | 79  | 
(* add_consts(_i) *)  | 
| 
4129
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
80  | 
|
| 12625 | 81  | 
fun gen_add_consts prep_typ raw_decls thy =  | 
82  | 
let  | 
|
| 16842 | 83  | 
val decls = map (upd_second (Thm.typ_of o prep_typ thy)) raw_decls;  | 
| 16621 | 84  | 
val (contconst_decls, normal_decls) = List.partition is_contconst decls;  | 
| 
4129
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
85  | 
val transformed_decls = map transform contconst_decls;  | 
| 12625 | 86  | 
in  | 
87  | 
thy  | 
|
88  | 
|> Theory.add_consts_i normal_decls  | 
|
89  | 
|> Theory.add_consts_i (map first transformed_decls)  | 
|
90  | 
|> Theory.add_syntax_i (map second transformed_decls)  | 
|
| 15570 | 91  | 
|> Theory.add_trrules_i (List.concat (map third transformed_decls))  | 
| 12625 | 92  | 
end;  | 
93  | 
||
94  | 
val add_consts = gen_add_consts Thm.read_ctyp;  | 
|
95  | 
val add_consts_i = gen_add_consts Thm.ctyp_of;  | 
|
96  | 
||
97  | 
||
98  | 
(* outer syntax *)  | 
|
99  | 
||
| 17057 | 100  | 
local structure P = OuterParse and K = OuterKeyword in  | 
| 12625 | 101  | 
|
102  | 
val constsP =  | 
|
103  | 
OuterSyntax.command "consts" "declare constants (HOLCF)" K.thy_decl  | 
|
| 
12876
 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 
wenzelm 
parents: 
12625 
diff
changeset
 | 
104  | 
(Scan.repeat1 P.const >> (Toplevel.theory o add_consts));  | 
| 12625 | 105  | 
|
| 
14682
 
a5072752114c
HOLCF: discontinued special version of 'constdefs';
 
wenzelm 
parents: 
12876 
diff
changeset
 | 
106  | 
val _ = OuterSyntax.add_parsers [constsP];  | 
| 12625 | 107  | 
|
| 
4129
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
108  | 
end;  | 
| 
 
2fd816aa6206
HOLCFLogic.ML, contconsts.ML renamed to holcf_logic.ML, cont_consts.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
109  | 
|
| 12625 | 110  | 
end;  |