| author | blanchet | 
| Tue, 14 Sep 2010 23:38:20 +0200 | |
| changeset 39376 | ca81b7ae543c | 
| parent 39205 | 13c6e91efcb6 | 
| child 39397 | 9b0a8d72edc8 | 
| permissions | -rw-r--r-- | 
| 37744 | 1  | 
(* Title: Tools/Code/code_thingol.ML  | 
| 24219 | 2  | 
Author: Florian Haftmann, TU Muenchen  | 
3  | 
||
4  | 
Intermediate language ("Thin-gol") representing executable code.
 | 
|
| 24918 | 5  | 
Representation and translation.  | 
| 24219 | 6  | 
*)  | 
7  | 
||
8  | 
infix 8 `%%;  | 
|
9  | 
infix 4 `$;  | 
|
10  | 
infix 4 `$$;  | 
|
| 31724 | 11  | 
infixr 3 `|=>;  | 
12  | 
infixr 3 `|==>;  | 
|
| 24219 | 13  | 
|
14  | 
signature BASIC_CODE_THINGOL =  | 
|
15  | 
sig  | 
|
16  | 
type vname = string;  | 
|
17  | 
datatype dict =  | 
|
18  | 
DictConst of string * dict list list  | 
|
19  | 
| DictVar of string list * (vname * (int * int));  | 
|
20  | 
datatype itype =  | 
|
21  | 
`%% of string * itype list  | 
|
22  | 
| ITyVar of vname;  | 
|
| 31049 | 23  | 
type const = string * ((itype list * dict list list) * itype list (*types of arguments*))  | 
| 24219 | 24  | 
datatype iterm =  | 
| 24591 | 25  | 
IConst of const  | 
| 31889 | 26  | 
| IVar of vname option  | 
| 24219 | 27  | 
| `$ of iterm * iterm  | 
| 31888 | 28  | 
| `|=> of (vname option * itype) * iterm  | 
| 24219 | 29  | 
| ICase of ((iterm * itype) * (iterm * iterm) list) * iterm;  | 
30  | 
(*((term, type), [(selector pattern, body term )]), primitive term)*)  | 
|
31  | 
val `$$ : iterm * iterm list -> iterm;  | 
|
| 31888 | 32  | 
val `|==> : (vname option * itype) list * iterm -> iterm;  | 
| 24219 | 33  | 
type typscheme = (vname * sort) list * itype;  | 
34  | 
end;  | 
|
35  | 
||
36  | 
signature CODE_THINGOL =  | 
|
37  | 
sig  | 
|
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
38  | 
include BASIC_CODE_THINGOL  | 
| 34084 | 39  | 
val fun_tyco: string  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
40  | 
  val unfoldl: ('a -> ('a * 'b) option) -> 'a -> 'a * 'b list
 | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
41  | 
  val unfoldr: ('a -> ('b * 'a) option) -> 'a -> 'b list * 'a
 | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
42  | 
val unfold_fun: itype -> itype list * itype  | 
| 37640 | 43  | 
val unfold_fun_n: int -> itype -> itype list * itype  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
44  | 
val unfold_app: iterm -> iterm * iterm list  | 
| 31888 | 45  | 
val unfold_abs: iterm -> (vname option * itype) list * iterm  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
46  | 
val split_let: iterm -> (((iterm * itype) * iterm) * iterm) option  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
47  | 
val unfold_let: iterm -> ((iterm * itype) * iterm) list * iterm  | 
| 31889 | 48  | 
val split_pat_abs: iterm -> ((iterm * itype) * iterm) option  | 
49  | 
val unfold_pat_abs: iterm -> (iterm * itype) list * iterm  | 
|
| 31049 | 50  | 
val unfold_const_app: iterm -> (const * iterm list) option  | 
| 32909 | 51  | 
val is_IVar: iterm -> bool  | 
| 31049 | 52  | 
val eta_expand: int -> const * iterm list -> iterm  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
53  | 
val contains_dictvar: iterm -> bool  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
54  | 
val locally_monomorphic: iterm -> bool  | 
| 
31935
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
55  | 
val add_constnames: iterm -> string list -> string list  | 
| 32917 | 56  | 
val add_tyconames: iterm -> string list -> string list  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
57  | 
val fold_varnames: (string -> 'a -> 'a) -> iterm -> 'a -> 'a  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
58  | 
|
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
59  | 
type naming  | 
| 28706 | 60  | 
val empty_naming: naming  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
61  | 
val lookup_class: naming -> class -> string option  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
62  | 
val lookup_classrel: naming -> class * class -> string option  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
63  | 
val lookup_tyco: naming -> string -> string option  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
64  | 
val lookup_instance: naming -> class * string -> string option  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
65  | 
val lookup_const: naming -> string -> string option  | 
| 31054 | 66  | 
val ensure_declared_const: theory -> string -> naming -> string * naming  | 
| 24219 | 67  | 
|
| 24918 | 68  | 
datatype stmt =  | 
| 27024 | 69  | 
NoStmt  | 
| 37437 | 70  | 
| Fun of string * ((typscheme * ((iterm list * iterm) * (thm option * bool)) list) * thm option)  | 
| 
37448
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
71  | 
| Datatype of string * ((vname * sort) list *  | 
| 
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
72  | 
((string * vname list (*type argument wrt. canonical order*)) * itype list) list)  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
73  | 
| Datatypecons of string * string  | 
| 
37447
 
ad3e04f289b6
transitive superclasses were also only a misunderstanding
 
haftmann 
parents: 
37446 
diff
changeset
 | 
74  | 
| Class of class * (vname * ((class * string) list * (string * itype) list))  | 
| 24219 | 75  | 
| Classrel of class * class  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
76  | 
| Classparam of string * class  | 
| 
37448
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
77  | 
| Classinst of (class * (string * (vname * sort) list) (*class and arity*))  | 
| 
37445
 
e372fa3c7239
dropped obscure type argument weakening mapping -- was only a misunderstanding
 
haftmann 
parents: 
37440 
diff
changeset
 | 
78  | 
* ((class * (string * (string * dict list list))) list (*super instances*)  | 
| 
37448
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
79  | 
* (((string * const) * (thm * bool)) list (*class parameter instances*)  | 
| 
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
80  | 
* ((string * const) * (thm * bool)) list (*super class parameter instances*)))  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
81  | 
type program = stmt Graph.T  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
82  | 
val empty_funs: program -> string list  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
83  | 
val map_terms_bottom_up: (iterm -> iterm) -> iterm -> iterm  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
84  | 
val map_terms_stmt: (iterm -> iterm) -> stmt -> stmt  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
85  | 
val is_cons: program -> string -> bool  | 
| 
37440
 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 
haftmann 
parents: 
37437 
diff
changeset
 | 
86  | 
val is_case: stmt -> bool  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
87  | 
val contr_classparam_typs: program -> string -> itype option list  | 
| 32895 | 88  | 
val labelled_name: theory -> program -> string -> string  | 
89  | 
val group_stmts: theory -> program  | 
|
90  | 
-> ((string * stmt) list * (string * stmt) list  | 
|
91  | 
* ((string * stmt) list * (string * stmt) list)) list  | 
|
| 24219 | 92  | 
|
| 31036 | 93  | 
val read_const_exprs: theory -> string list -> string list * string list  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
94  | 
val consts_program: theory -> bool -> string list -> string list * (naming * program)  | 
| 
38669
 
9ff76d0f0610
refined and unified naming convention for dynamic code evaluation techniques
 
haftmann 
parents: 
37871 
diff
changeset
 | 
95  | 
val dynamic_eval_conv: theory  | 
| 
38672
 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 
haftmann 
parents: 
38669 
diff
changeset
 | 
96  | 
-> (naming -> program -> ((string * sort) list * typscheme) * iterm -> string list -> conv)  | 
| 
 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 
haftmann 
parents: 
38669 
diff
changeset
 | 
97  | 
-> conv  | 
| 
38669
 
9ff76d0f0610
refined and unified naming convention for dynamic code evaluation techniques
 
haftmann 
parents: 
37871 
diff
changeset
 | 
98  | 
val dynamic_eval_value: theory -> ((term -> term) -> 'a -> 'a)  | 
| 
31063
 
88aaab83b6fc
dropped explicit suppport for frees in evaluation conversion stack
 
haftmann 
parents: 
31054 
diff
changeset
 | 
99  | 
-> (naming -> program -> ((string * sort) list * typscheme) * iterm -> string list -> 'a)  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
100  | 
-> term -> 'a  | 
| 
38672
 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 
haftmann 
parents: 
38669 
diff
changeset
 | 
101  | 
val static_eval_conv: theory -> string list  | 
| 
 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 
haftmann 
parents: 
38669 
diff
changeset
 | 
102  | 
-> (naming -> program -> ((string * sort) list * typscheme) * iterm -> string list -> conv)  | 
| 
 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 
haftmann 
parents: 
38669 
diff
changeset
 | 
103  | 
-> conv  | 
| 
 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 
haftmann 
parents: 
38669 
diff
changeset
 | 
104  | 
val static_eval_conv_simple: theory -> string list  | 
| 
 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 
haftmann 
parents: 
38669 
diff
changeset
 | 
105  | 
-> (program -> conv) -> conv  | 
| 24219 | 106  | 
end;  | 
107  | 
||
| 28054 | 108  | 
structure Code_Thingol: CODE_THINGOL =  | 
| 24219 | 109  | 
struct  | 
110  | 
||
111  | 
(** auxiliary **)  | 
|
112  | 
||
113  | 
fun unfoldl dest x =  | 
|
114  | 
case dest x  | 
|
115  | 
of NONE => (x, [])  | 
|
116  | 
| SOME (x1, x2) =>  | 
|
117  | 
let val (x', xs') = unfoldl dest x1 in (x', xs' @ [x2]) end;  | 
|
118  | 
||
119  | 
fun unfoldr dest x =  | 
|
120  | 
case dest x  | 
|
121  | 
of NONE => ([], x)  | 
|
122  | 
| SOME (x1, x2) =>  | 
|
123  | 
let val (xs', x') = unfoldr dest x2 in (x1::xs', x') end;  | 
|
124  | 
||
125  | 
||
| 
29962
 
bd4dc7fa742d
tuned comments, stripped ID, deleted superfluous code
 
haftmann 
parents: 
29952 
diff
changeset
 | 
126  | 
(** language core - types, terms **)  | 
| 24219 | 127  | 
|
128  | 
type vname = string;  | 
|
129  | 
||
130  | 
datatype dict =  | 
|
131  | 
DictConst of string * dict list list  | 
|
132  | 
| DictVar of string list * (vname * (int * int));  | 
|
133  | 
||
134  | 
datatype itype =  | 
|
135  | 
`%% of string * itype list  | 
|
136  | 
| ITyVar of vname;  | 
|
137  | 
||
| 31049 | 138  | 
type const = string * ((itype list * dict list list) * itype list (*types of arguments*))  | 
| 24591 | 139  | 
|
| 24219 | 140  | 
datatype iterm =  | 
| 24591 | 141  | 
IConst of const  | 
| 31889 | 142  | 
| IVar of vname option  | 
| 24219 | 143  | 
| `$ of iterm * iterm  | 
| 31888 | 144  | 
| `|=> of (vname option * itype) * iterm  | 
| 24219 | 145  | 
| ICase of ((iterm * itype) * (iterm * iterm) list) * iterm;  | 
146  | 
(*see also signature*)  | 
|
147  | 
||
| 32909 | 148  | 
fun is_IVar (IVar _) = true  | 
149  | 
| is_IVar _ = false;  | 
|
150  | 
||
| 24219 | 151  | 
val op `$$ = Library.foldl (op `$);  | 
| 31724 | 152  | 
val op `|==> = Library.foldr (op `|=>);  | 
| 24219 | 153  | 
|
154  | 
val unfold_app = unfoldl  | 
|
155  | 
(fn op `$ t => SOME t  | 
|
156  | 
| _ => NONE);  | 
|
157  | 
||
| 31874 | 158  | 
val unfold_abs = unfoldr  | 
159  | 
(fn op `|=> t => SOME t  | 
|
| 24219 | 160  | 
| _ => NONE);  | 
161  | 
||
162  | 
val split_let =  | 
|
163  | 
(fn ICase (((td, ty), [(p, t)]), _) => SOME (((p, ty), td), t)  | 
|
164  | 
| _ => NONE);  | 
|
165  | 
||
166  | 
val unfold_let = unfoldr split_let;  | 
|
167  | 
||
168  | 
fun unfold_const_app t =  | 
|
169  | 
case unfold_app t  | 
|
170  | 
of (IConst c, ts) => SOME (c, ts)  | 
|
171  | 
| _ => NONE;  | 
|
172  | 
||
| 32917 | 173  | 
fun fold_constexprs f =  | 
174  | 
let  | 
|
175  | 
fun fold' (IConst c) = f c  | 
|
176  | 
| fold' (IVar _) = I  | 
|
177  | 
| fold' (t1 `$ t2) = fold' t1 #> fold' t2  | 
|
178  | 
| fold' (_ `|=> t) = fold' t  | 
|
179  | 
| fold' (ICase (((t, _), ds), _)) = fold' t  | 
|
180  | 
#> fold (fn (pat, body) => fold' pat #> fold' body) ds  | 
|
181  | 
in fold' end;  | 
|
182  | 
||
183  | 
val add_constnames = fold_constexprs (fn (c, _) => insert (op =) c);  | 
|
184  | 
||
185  | 
fun add_tycos (tyco `%% tys) = insert (op =) tyco #> fold add_tycos tys  | 
|
186  | 
| add_tycos (ITyVar _) = I;  | 
|
187  | 
||
188  | 
val add_tyconames = fold_constexprs (fn (_, ((tys, _), _)) => fold add_tycos tys);  | 
|
| 24219 | 189  | 
|
190  | 
fun fold_varnames f =  | 
|
191  | 
let  | 
|
| 
31935
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
192  | 
fun fold_aux add f =  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
193  | 
let  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
194  | 
fun fold_term _ (IConst _) = I  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
195  | 
| fold_term vs (IVar (SOME v)) = if member (op =) vs v then I else f v  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
196  | 
| fold_term _ (IVar NONE) = I  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
197  | 
| fold_term vs (t1 `$ t2) = fold_term vs t1 #> fold_term vs t2  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
198  | 
| fold_term vs ((SOME v, _) `|=> t) = fold_term (insert (op =) v vs) t  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
199  | 
| fold_term vs ((NONE, _) `|=> t) = fold_term vs t  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
200  | 
| fold_term vs (ICase (((t, _), ds), _)) = fold_term vs t #> fold (fold_case vs) ds  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
201  | 
and fold_case vs (p, t) = fold_term (add p vs) t;  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
202  | 
in fold_term [] end;  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
203  | 
fun add t = fold_aux add (insert (op =)) t;  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
204  | 
in fold_aux add f end;  | 
| 24219 | 205  | 
|
| 
31935
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
206  | 
fun exists_var t v = fold_varnames (fn w => fn b => v = w orelse b) t false;  | 
| 31874 | 207  | 
|
| 31889 | 208  | 
fun split_pat_abs ((NONE, ty) `|=> t) = SOME ((IVar NONE, ty), t)  | 
| 31888 | 209  | 
| split_pat_abs ((SOME v, ty) `|=> t) = SOME (case t  | 
| 31889 | 210  | 
of ICase (((IVar (SOME w), _), [(p, t')]), _) =>  | 
| 31888 | 211  | 
if v = w andalso (exists_var p v orelse not (exists_var t' v))  | 
| 31889 | 212  | 
then ((p, ty), t')  | 
213  | 
else ((IVar (SOME v), ty), t)  | 
|
214  | 
| _ => ((IVar (SOME v), ty), t))  | 
|
| 31888 | 215  | 
| split_pat_abs _ = NONE;  | 
| 31874 | 216  | 
|
217  | 
val unfold_pat_abs = unfoldr split_pat_abs;  | 
|
| 24219 | 218  | 
|
| 
31890
 
e943b039f0ac
an intermediate step towards a refined translation of cases
 
haftmann 
parents: 
31889 
diff
changeset
 | 
219  | 
fun unfold_abs_eta [] t = ([], t)  | 
| 
 
e943b039f0ac
an intermediate step towards a refined translation of cases
 
haftmann 
parents: 
31889 
diff
changeset
 | 
220  | 
| unfold_abs_eta (_ :: tys) (v_ty `|=> t) =  | 
| 
 
e943b039f0ac
an intermediate step towards a refined translation of cases
 
haftmann 
parents: 
31889 
diff
changeset
 | 
221  | 
let  | 
| 
 
e943b039f0ac
an intermediate step towards a refined translation of cases
 
haftmann 
parents: 
31889 
diff
changeset
 | 
222  | 
val (vs_tys, t') = unfold_abs_eta tys t;  | 
| 
 
e943b039f0ac
an intermediate step towards a refined translation of cases
 
haftmann 
parents: 
31889 
diff
changeset
 | 
223  | 
in (v_ty :: vs_tys, t') end  | 
| 31892 | 224  | 
| unfold_abs_eta tys t =  | 
| 
31890
 
e943b039f0ac
an intermediate step towards a refined translation of cases
 
haftmann 
parents: 
31889 
diff
changeset
 | 
225  | 
let  | 
| 
 
e943b039f0ac
an intermediate step towards a refined translation of cases
 
haftmann 
parents: 
31889 
diff
changeset
 | 
226  | 
val ctxt = fold_varnames Name.declare t Name.context;  | 
| 
 
e943b039f0ac
an intermediate step towards a refined translation of cases
 
haftmann 
parents: 
31889 
diff
changeset
 | 
227  | 
val vs_tys = (map o apfst) SOME (Name.names ctxt "a" tys);  | 
| 
 
e943b039f0ac
an intermediate step towards a refined translation of cases
 
haftmann 
parents: 
31889 
diff
changeset
 | 
228  | 
in (vs_tys, t `$$ map (IVar o fst) vs_tys) end;  | 
| 
 
e943b039f0ac
an intermediate step towards a refined translation of cases
 
haftmann 
parents: 
31889 
diff
changeset
 | 
229  | 
|
| 37841 | 230  | 
fun eta_expand k (c as (name, (_, tys)), ts) =  | 
| 24219 | 231  | 
let  | 
232  | 
val j = length ts;  | 
|
233  | 
val l = k - j;  | 
|
| 37841 | 234  | 
val _ = if l > length tys  | 
235  | 
      then error ("Impossible eta-expansion for constant " ^ quote name) else ();
 | 
|
| 24219 | 236  | 
val ctxt = (fold o fold_varnames) Name.declare ts Name.context;  | 
| 31889 | 237  | 
val vs_tys = (map o apfst) SOME  | 
| 33957 | 238  | 
(Name.names ctxt "a" ((take l o drop j) tys));  | 
| 31889 | 239  | 
in vs_tys `|==> IConst c `$$ ts @ map (IVar o fst) vs_tys end;  | 
| 24219 | 240  | 
|
| 
24662
 
f79f6061525c
more precise treatment of free dictionary parameters for evaluation
 
haftmann 
parents: 
24591 
diff
changeset
 | 
241  | 
fun contains_dictvar t =  | 
| 
 
f79f6061525c
more precise treatment of free dictionary parameters for evaluation
 
haftmann 
parents: 
24591 
diff
changeset
 | 
242  | 
let  | 
| 
31935
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
243  | 
fun cont_dict (DictConst (_, dss)) = (exists o exists) cont_dict dss  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
244  | 
| cont_dict (DictVar _) = true;  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
245  | 
fun cont_term (IConst (_, ((_, dss), _))) = (exists o exists) cont_dict dss  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
246  | 
| cont_term (IVar _) = false  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
247  | 
| cont_term (t1 `$ t2) = cont_term t1 orelse cont_term t2  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
248  | 
| cont_term (_ `|=> t) = cont_term t  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
249  | 
| cont_term (ICase (_, t)) = cont_term t;  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
250  | 
in cont_term t end;  | 
| 
24662
 
f79f6061525c
more precise treatment of free dictionary parameters for evaluation
 
haftmann 
parents: 
24591 
diff
changeset
 | 
251  | 
|
| 25621 | 252  | 
fun locally_monomorphic (IConst _) = false  | 
253  | 
| locally_monomorphic (IVar _) = true  | 
|
254  | 
| locally_monomorphic (t `$ _) = locally_monomorphic t  | 
|
| 31724 | 255  | 
| locally_monomorphic (_ `|=> t) = locally_monomorphic t  | 
| 25621 | 256  | 
| locally_monomorphic (ICase ((_, ds), _)) = exists (locally_monomorphic o snd) ds;  | 
257  | 
||
258  | 
||
| 28688 | 259  | 
(** namings **)  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
260  | 
|
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
261  | 
(* policies *)  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
262  | 
|
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
263  | 
local  | 
| 33172 | 264  | 
fun thyname_of_class thy = #theory_name o Name_Space.the_entry (Sign.class_space thy);  | 
265  | 
fun thyname_of_instance thy inst = case AxClass.thynames_of_arity thy inst  | 
|
266  | 
   of [] => error ("No such instance: " ^ quote (snd inst ^ " :: " ^ fst inst))
 | 
|
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
267  | 
| thyname :: _ => thyname;  | 
| 28706 | 268  | 
fun thyname_of_const thy c = case AxClass.class_of_param thy c  | 
269  | 
of SOME class => thyname_of_class thy class  | 
|
| 
35299
 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 
haftmann 
parents: 
35226 
diff
changeset
 | 
270  | 
| NONE => (case Code.get_type_of_constr_or_abstr thy c  | 
| 35226 | 271  | 
of SOME (tyco, _) => Codegen.thyname_of_type thy tyco  | 
| 33172 | 272  | 
| NONE => Codegen.thyname_of_const thy c);  | 
| 33420 | 273  | 
fun purify_base "==>" = "follows"  | 
| 31036 | 274  | 
| purify_base s = Name.desymbolize false s;  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
275  | 
fun namify thy get_basename get_thyname name =  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
276  | 
let  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
277  | 
val prefix = get_thyname thy name;  | 
| 31036 | 278  | 
val base = (purify_base o get_basename) name;  | 
| 
30364
 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
279  | 
in Long_Name.append prefix base end;  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
280  | 
in  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
281  | 
|
| 
30364
 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
282  | 
fun namify_class thy = namify thy Long_Name.base_name thyname_of_class;  | 
| 
37384
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
283  | 
fun namify_classrel thy = namify thy (fn (sub_class, super_class) =>  | 
| 37431 | 284  | 
Long_Name.base_name super_class ^ "_" ^ Long_Name.base_name sub_class)  | 
| 33172 | 285  | 
(fn thy => thyname_of_class thy o fst);  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
286  | 
(*order fits nicely with composed projections*)  | 
| 28688 | 287  | 
fun namify_tyco thy "fun" = "Pure.fun"  | 
| 33172 | 288  | 
| namify_tyco thy tyco = namify thy Long_Name.base_name Codegen.thyname_of_type tyco;  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
289  | 
fun namify_instance thy = namify thy (fn (class, tyco) =>  | 
| 
30364
 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
290  | 
Long_Name.base_name class ^ "_" ^ Long_Name.base_name tyco) thyname_of_instance;  | 
| 
 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
291  | 
fun namify_const thy = namify thy Long_Name.base_name thyname_of_const;  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
292  | 
|
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
293  | 
end; (* local *)  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
294  | 
|
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
295  | 
|
| 28688 | 296  | 
(* data *)  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
297  | 
|
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
298  | 
datatype naming = Naming of {
 | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
299  | 
class: class Symtab.table * Name.context,  | 
| 
30648
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30364 
diff
changeset
 | 
300  | 
classrel: string Symreltab.table * Name.context,  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
301  | 
tyco: string Symtab.table * Name.context,  | 
| 
30648
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30364 
diff
changeset
 | 
302  | 
instance: string Symreltab.table * Name.context,  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
303  | 
const: string Symtab.table * Name.context  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
304  | 
}  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
305  | 
|
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
306  | 
fun dest_Naming (Naming naming) = naming;  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
307  | 
|
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
308  | 
val empty_naming = Naming {
 | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
309  | 
class = (Symtab.empty, Name.context),  | 
| 
30648
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30364 
diff
changeset
 | 
310  | 
classrel = (Symreltab.empty, Name.context),  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
311  | 
tyco = (Symtab.empty, Name.context),  | 
| 
30648
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30364 
diff
changeset
 | 
312  | 
instance = (Symreltab.empty, Name.context),  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
313  | 
const = (Symtab.empty, Name.context)  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
314  | 
};  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
315  | 
|
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
316  | 
local  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
317  | 
fun mk_naming (class, classrel, tyco, instance, const) =  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
318  | 
    Naming { class = class, classrel = classrel,
 | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
319  | 
tyco = tyco, instance = instance, const = const };  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
320  | 
  fun map_naming f (Naming { class, classrel, tyco, instance, const }) =
 | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
321  | 
mk_naming (f (class, classrel, tyco, instance, const));  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
322  | 
in  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
323  | 
fun map_class f = map_naming  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
324  | 
(fn (class, classrel, tyco, inst, const) =>  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
325  | 
(f class, classrel, tyco, inst, const));  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
326  | 
fun map_classrel f = map_naming  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
327  | 
(fn (class, classrel, tyco, inst, const) =>  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
328  | 
(class, f classrel, tyco, inst, const));  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
329  | 
fun map_tyco f = map_naming  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
330  | 
(fn (class, classrel, tyco, inst, const) =>  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
331  | 
(class, classrel, f tyco, inst, const));  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
332  | 
fun map_instance f = map_naming  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
333  | 
(fn (class, classrel, tyco, inst, const) =>  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
334  | 
(class, classrel, tyco, f inst, const));  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
335  | 
fun map_const f = map_naming  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
336  | 
(fn (class, classrel, tyco, inst, const) =>  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
337  | 
(class, classrel, tyco, inst, f const));  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
338  | 
end; (*local*)  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
339  | 
|
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
340  | 
fun add_variant update (thing, name) (tab, used) =  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
341  | 
let  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
342  | 
val (name', used') = yield_singleton Name.variants name used;  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
343  | 
val tab' = update (thing, name') tab;  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
344  | 
in (tab', used') end;  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
345  | 
|
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
346  | 
fun declare thy mapp lookup update namify thing =  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
347  | 
mapp (add_variant update (thing, namify thy thing))  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
348  | 
#> `(fn naming => the (lookup naming thing));  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
349  | 
|
| 28688 | 350  | 
|
351  | 
(* lookup and declare *)  | 
|
352  | 
||
353  | 
local  | 
|
354  | 
||
355  | 
val suffix_class = "class";  | 
|
356  | 
val suffix_classrel = "classrel"  | 
|
357  | 
val suffix_tyco = "tyco";  | 
|
358  | 
val suffix_instance = "inst";  | 
|
359  | 
val suffix_const = "const";  | 
|
360  | 
||
361  | 
fun add_suffix nsp NONE = NONE  | 
|
| 
30364
 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
362  | 
| add_suffix nsp (SOME name) = SOME (Long_Name.append name nsp);  | 
| 28688 | 363  | 
|
364  | 
in  | 
|
365  | 
||
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
366  | 
val lookup_class = add_suffix suffix_class  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
367  | 
oo Symtab.lookup o fst o #class o dest_Naming;  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
368  | 
val lookup_classrel = add_suffix suffix_classrel  | 
| 
30648
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30364 
diff
changeset
 | 
369  | 
oo Symreltab.lookup o fst o #classrel o dest_Naming;  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
370  | 
val lookup_tyco = add_suffix suffix_tyco  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
371  | 
oo Symtab.lookup o fst o #tyco o dest_Naming;  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
372  | 
val lookup_instance = add_suffix suffix_instance  | 
| 
30648
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30364 
diff
changeset
 | 
373  | 
oo Symreltab.lookup o fst o #instance o dest_Naming;  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
374  | 
val lookup_const = add_suffix suffix_const  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
375  | 
oo Symtab.lookup o fst o #const o dest_Naming;  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
376  | 
|
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
377  | 
fun declare_class thy = declare thy map_class  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
378  | 
lookup_class Symtab.update_new namify_class;  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
379  | 
fun declare_classrel thy = declare thy map_classrel  | 
| 
30648
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30364 
diff
changeset
 | 
380  | 
lookup_classrel Symreltab.update_new namify_classrel;  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
381  | 
fun declare_tyco thy = declare thy map_tyco  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
382  | 
lookup_tyco Symtab.update_new namify_tyco;  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
383  | 
fun declare_instance thy = declare thy map_instance  | 
| 
30648
 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 
haftmann 
parents: 
30364 
diff
changeset
 | 
384  | 
lookup_instance Symreltab.update_new namify_instance;  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
385  | 
fun declare_const thy = declare thy map_const  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
386  | 
lookup_const Symtab.update_new namify_const;  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
387  | 
|
| 31054 | 388  | 
fun ensure_declared_const thy const naming =  | 
389  | 
case lookup_const naming const  | 
|
390  | 
of SOME const' => (const', naming)  | 
|
391  | 
| NONE => declare_const thy const naming;  | 
|
392  | 
||
| 
37384
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
393  | 
val fun_tyco = Long_Name.append (namify_tyco Pure.thy "fun") suffix_tyco  | 
| 
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
394  | 
(*depends on add_suffix*);  | 
| 34084 | 395  | 
|
| 28688 | 396  | 
val unfold_fun = unfoldr  | 
| 34084 | 397  | 
(fn tyco `%% [ty1, ty2] => if tyco = fun_tyco then SOME (ty1, ty2) else NONE  | 
| 
37384
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
398  | 
| _ => NONE);  | 
| 28688 | 399  | 
|
| 37640 | 400  | 
fun unfold_fun_n n ty =  | 
401  | 
let  | 
|
402  | 
val (tys1, ty1) = unfold_fun ty;  | 
|
403  | 
val (tys3, tys2) = chop n tys1;  | 
|
404  | 
val ty3 = Library.foldr (fn (ty1, ty2) => fun_tyco `%% [ty1, ty2]) (tys2, ty1);  | 
|
405  | 
in (tys3, ty3) end;  | 
|
406  | 
||
| 28688 | 407  | 
end; (* local *)  | 
408  | 
||
| 24219 | 409  | 
|
| 27103 | 410  | 
(** statements, abstract programs **)  | 
| 24219 | 411  | 
|
412  | 
type typscheme = (vname * sort) list * itype;  | 
|
| 
37447
 
ad3e04f289b6
transitive superclasses were also only a misunderstanding
 
haftmann 
parents: 
37446 
diff
changeset
 | 
413  | 
datatype stmt =  | 
| 27024 | 414  | 
NoStmt  | 
| 37437 | 415  | 
| Fun of string * ((typscheme * ((iterm list * iterm) * (thm option * bool)) list) * thm option)  | 
| 
37448
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
416  | 
| Datatype of string * ((vname * sort) list * ((string * vname list) * itype list) list)  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
417  | 
| Datatypecons of string * string  | 
| 
37447
 
ad3e04f289b6
transitive superclasses were also only a misunderstanding
 
haftmann 
parents: 
37446 
diff
changeset
 | 
418  | 
| Class of class * (vname * ((class * string) list * (string * itype) list))  | 
| 24219 | 419  | 
| Classrel of class * class  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
420  | 
| Classparam of string * class  | 
| 24219 | 421  | 
| Classinst of (class * (string * (vname * sort) list))  | 
| 
37445
 
e372fa3c7239
dropped obscure type argument weakening mapping -- was only a misunderstanding
 
haftmann 
parents: 
37440 
diff
changeset
 | 
422  | 
* ((class * (string * (string * dict list list))) list  | 
| 
37448
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
423  | 
* (((string * const) * (thm * bool)) list  | 
| 
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
424  | 
* ((string * const) * (thm * bool)) list))  | 
| 
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
425  | 
(*see also signature*);  | 
| 24219 | 426  | 
|
| 27103 | 427  | 
type program = stmt Graph.T;  | 
| 24219 | 428  | 
|
| 27103 | 429  | 
fun empty_funs program =  | 
| 37437 | 430  | 
Graph.fold (fn (name, (Fun (c, ((_, []), _)), _)) => cons c  | 
| 27103 | 431  | 
| _ => I) program [];  | 
| 24219 | 432  | 
|
| 27711 | 433  | 
fun map_terms_bottom_up f (t as IConst _) = f t  | 
434  | 
| map_terms_bottom_up f (t as IVar _) = f t  | 
|
435  | 
| map_terms_bottom_up f (t1 `$ t2) = f  | 
|
436  | 
(map_terms_bottom_up f t1 `$ map_terms_bottom_up f t2)  | 
|
| 31724 | 437  | 
| map_terms_bottom_up f ((v, ty) `|=> t) = f  | 
438  | 
((v, ty) `|=> map_terms_bottom_up f t)  | 
|
| 27711 | 439  | 
| map_terms_bottom_up f (ICase (((t, ty), ps), t0)) = f  | 
440  | 
(ICase (((map_terms_bottom_up f t, ty), (map o pairself)  | 
|
441  | 
(map_terms_bottom_up f) ps), map_terms_bottom_up f t0));  | 
|
442  | 
||
| 
37448
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
443  | 
fun map_classparam_instances_as_term f =  | 
| 
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
444  | 
(map o apfst o apsnd) (fn const => case f (IConst const) of IConst const' => const')  | 
| 
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
445  | 
|
| 27711 | 446  | 
fun map_terms_stmt f NoStmt = NoStmt  | 
| 37437 | 447  | 
| map_terms_stmt f (Fun (c, ((tysm, eqs), case_cong))) = Fun (c, ((tysm, (map o apfst)  | 
448  | 
(fn (ts, t) => (map f ts, f t)) eqs), case_cong))  | 
|
| 27711 | 449  | 
| map_terms_stmt f (stmt as Datatype _) = stmt  | 
450  | 
| map_terms_stmt f (stmt as Datatypecons _) = stmt  | 
|
451  | 
| map_terms_stmt f (stmt as Class _) = stmt  | 
|
452  | 
| map_terms_stmt f (stmt as Classrel _) = stmt  | 
|
453  | 
| map_terms_stmt f (stmt as Classparam _) = stmt  | 
|
| 
37448
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
454  | 
| map_terms_stmt f (Classinst (arity, (super_instances, classparam_instances))) =  | 
| 
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
455  | 
Classinst (arity, (super_instances, (pairself o map_classparam_instances_as_term) f classparam_instances));  | 
| 27711 | 456  | 
|
| 27103 | 457  | 
fun is_cons program name = case Graph.get_node program name  | 
| 24219 | 458  | 
of Datatypecons _ => true  | 
459  | 
| _ => false;  | 
|
460  | 
||
| 
37440
 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 
haftmann 
parents: 
37437 
diff
changeset
 | 
461  | 
fun is_case (Fun (_, (_, SOME _))) = true  | 
| 
 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 
haftmann 
parents: 
37437 
diff
changeset
 | 
462  | 
| is_case _ = false;  | 
| 
 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 
haftmann 
parents: 
37437 
diff
changeset
 | 
463  | 
|
| 27103 | 464  | 
fun contr_classparam_typs program name = case Graph.get_node program name  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
465  | 
of Classparam (_, class) => let  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
466  | 
val Class (_, (_, (_, params))) = Graph.get_node program class;  | 
| 25621 | 467  | 
val SOME ty = AList.lookup (op =) params name;  | 
468  | 
val (tys, res_ty) = unfold_fun ty;  | 
|
469  | 
fun no_tyvar (_ `%% tys) = forall no_tyvar tys  | 
|
470  | 
| no_tyvar (ITyVar _) = false;  | 
|
471  | 
in if no_tyvar res_ty  | 
|
472  | 
then map (fn ty => if no_tyvar ty then NONE else SOME ty) tys  | 
|
473  | 
else []  | 
|
474  | 
end  | 
|
475  | 
| _ => [];  | 
|
476  | 
||
| 32895 | 477  | 
fun labelled_name thy program name = case Graph.get_node program name  | 
478  | 
of Fun (c, _) => quote (Code.string_of_const thy c)  | 
|
479  | 
| Datatype (tyco, _) => "type " ^ quote (Sign.extern_type thy tyco)  | 
|
480  | 
| Datatypecons (c, _) => quote (Code.string_of_const thy c)  | 
|
481  | 
| Class (class, _) => "class " ^ quote (Sign.extern_class thy class)  | 
|
482  | 
| Classrel (sub, super) => let  | 
|
483  | 
val Class (sub, _) = Graph.get_node program sub  | 
|
484  | 
val Class (super, _) = Graph.get_node program super  | 
|
485  | 
in quote (Sign.extern_class thy sub ^ " < " ^ Sign.extern_class thy super) end  | 
|
486  | 
| Classparam (c, _) => quote (Code.string_of_const thy c)  | 
|
487  | 
| Classinst ((class, (tyco, _)), _) => let  | 
|
488  | 
val Class (class, _) = Graph.get_node program class  | 
|
489  | 
val Datatype (tyco, _) = Graph.get_node program tyco  | 
|
490  | 
in quote (Sign.extern_type thy tyco ^ " :: " ^ Sign.extern_class thy class) end  | 
|
491  | 
||
| 
37440
 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 
haftmann 
parents: 
37437 
diff
changeset
 | 
492  | 
fun linear_stmts program =  | 
| 
 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 
haftmann 
parents: 
37437 
diff
changeset
 | 
493  | 
rev (Graph.strong_conn program)  | 
| 
 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 
haftmann 
parents: 
37437 
diff
changeset
 | 
494  | 
|> map (AList.make (Graph.get_node program));  | 
| 
 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 
haftmann 
parents: 
37437 
diff
changeset
 | 
495  | 
|
| 32895 | 496  | 
fun group_stmts thy program =  | 
497  | 
let  | 
|
498  | 
fun is_fun (_, Fun _) = true | is_fun _ = false;  | 
|
499  | 
fun is_datatypecons (_, Datatypecons _) = true | is_datatypecons _ = false;  | 
|
500  | 
fun is_datatype (_, Datatype _) = true | is_datatype _ = false;  | 
|
501  | 
fun is_class (_, Class _) = true | is_class _ = false;  | 
|
502  | 
fun is_classrel (_, Classrel _) = true | is_classrel _ = false;  | 
|
503  | 
fun is_classparam (_, Classparam _) = true | is_classparam _ = false;  | 
|
504  | 
fun is_classinst (_, Classinst _) = true | is_classinst _ = false;  | 
|
505  | 
fun group stmts =  | 
|
506  | 
if forall (is_datatypecons orf is_datatype) stmts  | 
|
507  | 
then (filter is_datatype stmts, [], ([], []))  | 
|
508  | 
else if forall (is_class orf is_classrel orf is_classparam) stmts  | 
|
509  | 
then ([], filter is_class stmts, ([], []))  | 
|
510  | 
else if forall (is_fun orf is_classinst) stmts  | 
|
511  | 
then ([], [], List.partition is_fun stmts)  | 
|
512  | 
      else error ("Illegal mutual dependencies: " ^
 | 
|
513  | 
(commas o map (labelled_name thy program o fst)) stmts)  | 
|
514  | 
in  | 
|
| 
37440
 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 
haftmann 
parents: 
37437 
diff
changeset
 | 
515  | 
linear_stmts program  | 
| 32895 | 516  | 
|> map group  | 
517  | 
end;  | 
|
518  | 
||
| 24219 | 519  | 
|
| 27103 | 520  | 
(** translation kernel **)  | 
| 24219 | 521  | 
|
| 28724 | 522  | 
(* generic mechanisms *)  | 
523  | 
||
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
524  | 
fun ensure_stmt lookup declare generate thing (dep, (naming, program)) =  | 
| 24219 | 525  | 
let  | 
| 28706 | 526  | 
fun add_dep name = case dep of NONE => I  | 
527  | 
| SOME dep => Graph.add_edge (dep, name);  | 
|
528  | 
val (name, naming') = case lookup naming thing  | 
|
529  | 
of SOME name => (name, naming)  | 
|
530  | 
| NONE => declare thing naming;  | 
|
531  | 
in case try (Graph.get_node program) name  | 
|
532  | 
of SOME stmt => program  | 
|
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
533  | 
|> add_dep name  | 
| 28706 | 534  | 
|> pair naming'  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
535  | 
|> pair dep  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
536  | 
|> pair name  | 
| 28706 | 537  | 
| NONE => program  | 
538  | 
|> Graph.default_node (name, NoStmt)  | 
|
539  | 
|> add_dep name  | 
|
540  | 
|> pair naming'  | 
|
541  | 
|> curry generate (SOME name)  | 
|
542  | 
||> snd  | 
|
543  | 
|-> (fn stmt => (apsnd o Graph.map_node name) (K stmt))  | 
|
544  | 
|> pair dep  | 
|
545  | 
|> pair name  | 
|
| 24219 | 546  | 
end;  | 
547  | 
||
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
548  | 
exception PERMISSIVE of unit;  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
549  | 
|
| 37698 | 550  | 
fun translation_error thy permissive some_thm msg sub_msg =  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
551  | 
if permissive  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
552  | 
then raise PERMISSIVE ()  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
553  | 
else let  | 
| 35226 | 554  | 
val err_thm = case some_thm  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
555  | 
of SOME thm => "\n(in code equation " ^ Display.string_of_thm_global thy thm ^ ")"  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
556  | 
| NONE => "";  | 
| 37698 | 557  | 
in error (msg ^ err_thm ^ ":\n" ^ sub_msg) end;  | 
558  | 
||
559  | 
fun not_wellsorted thy permissive some_thm ty sort e =  | 
|
560  | 
let  | 
|
561  | 
val err_class = Sorts.class_error (Syntax.pp_global thy) e;  | 
|
562  | 
val err_typ = "Type " ^ Syntax.string_of_typ_global thy ty ^ " not of sort "  | 
|
563  | 
^ Syntax.string_of_sort_global thy sort;  | 
|
564  | 
in translation_error thy permissive some_thm "Wellsortedness error" (err_typ ^ "\n" ^ err_class) end;  | 
|
| 26972 | 565  | 
|
| 28724 | 566  | 
|
567  | 
(* translation *)  | 
|
568  | 
||
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
569  | 
fun ensure_tyco thy algbr eqngr permissive tyco =  | 
| 30932 | 570  | 
let  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
571  | 
val (vs, cos) = Code.get_type thy tyco;  | 
| 30932 | 572  | 
val stmt_datatype =  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
573  | 
fold_map (translate_tyvar_sort thy algbr eqngr permissive) vs  | 
| 
37448
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
574  | 
##>> fold_map (fn ((c, vs), tys) =>  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
575  | 
ensure_const thy algbr eqngr permissive c  | 
| 
37448
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
576  | 
##>> pair (map (unprefix "'") vs)  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
577  | 
##>> fold_map (translate_typ thy algbr eqngr permissive) tys) cos  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
578  | 
#>> (fn info => Datatype (tyco, info));  | 
| 30932 | 579  | 
in ensure_stmt lookup_tyco (declare_tyco thy) stmt_datatype tyco end  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
580  | 
and ensure_const thy algbr eqngr permissive c =  | 
| 30932 | 581  | 
let  | 
582  | 
fun stmt_datatypecons tyco =  | 
|
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
583  | 
ensure_tyco thy algbr eqngr permissive tyco  | 
| 30932 | 584  | 
#>> (fn tyco => Datatypecons (c, tyco));  | 
585  | 
fun stmt_classparam class =  | 
|
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
586  | 
ensure_class thy algbr eqngr permissive class  | 
| 30932 | 587  | 
#>> (fn class => Classparam (c, class));  | 
| 
34891
 
99b9a6290446
code certificates as integral part of code generation
 
haftmann 
parents: 
34251 
diff
changeset
 | 
588  | 
fun stmt_fun cert =  | 
| 
32872
 
019201eb7e07
variables in type schemes must be renamed simultaneously with variables in equations
 
haftmann 
parents: 
32795 
diff
changeset
 | 
589  | 
let  | 
| 35226 | 590  | 
val ((vs, ty), eqns) = Code.equations_of_cert thy cert;  | 
| 
37440
 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 
haftmann 
parents: 
37437 
diff
changeset
 | 
591  | 
val some_case_cong = Code.get_case_cong thy c;  | 
| 
32872
 
019201eb7e07
variables in type schemes must be renamed simultaneously with variables in equations
 
haftmann 
parents: 
32795 
diff
changeset
 | 
592  | 
in  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
593  | 
fold_map (translate_tyvar_sort thy algbr eqngr permissive) vs  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
594  | 
##>> translate_typ thy algbr eqngr permissive ty  | 
| 
37440
 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 
haftmann 
parents: 
37437 
diff
changeset
 | 
595  | 
##>> translate_eqns thy algbr eqngr permissive eqns  | 
| 
 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 
haftmann 
parents: 
37437 
diff
changeset
 | 
596  | 
#>> (fn info => Fun (c, (info, some_case_cong)))  | 
| 
32872
 
019201eb7e07
variables in type schemes must be renamed simultaneously with variables in equations
 
haftmann 
parents: 
32795 
diff
changeset
 | 
597  | 
end;  | 
| 
35299
 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 
haftmann 
parents: 
35226 
diff
changeset
 | 
598  | 
val stmt_const = case Code.get_type_of_constr_or_abstr thy c  | 
| 35226 | 599  | 
of SOME (tyco, _) => stmt_datatypecons tyco  | 
| 30932 | 600  | 
| NONE => (case AxClass.class_of_param thy c  | 
601  | 
of SOME class => stmt_classparam class  | 
|
| 
34891
 
99b9a6290446
code certificates as integral part of code generation
 
haftmann 
parents: 
34251 
diff
changeset
 | 
602  | 
| NONE => stmt_fun (Code_Preproc.cert eqngr c))  | 
| 30932 | 603  | 
in ensure_stmt lookup_const (declare_const thy) stmt_const c end  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
604  | 
and ensure_class thy (algbr as (_, algebra)) eqngr permissive class =  | 
| 24918 | 605  | 
let  | 
| 
37384
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
606  | 
val super_classes = (Sorts.minimize_sort algebra o Sorts.super_classes algebra) class;  | 
| 24969 | 607  | 
val cs = #params (AxClass.get_info thy class);  | 
| 24918 | 608  | 
val stmt_class =  | 
| 
37384
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
609  | 
fold_map (fn super_class => ensure_class thy algbr eqngr permissive super_class  | 
| 
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
610  | 
##>> ensure_classrel thy algbr eqngr permissive (class, super_class)) super_classes  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
611  | 
##>> fold_map (fn (c, ty) => ensure_const thy algbr eqngr permissive c  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
612  | 
##>> translate_typ thy algbr eqngr permissive ty) cs  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
613  | 
#>> (fn info => Class (class, (unprefix "'" Name.aT, info)))  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
614  | 
in ensure_stmt lookup_class (declare_class thy) stmt_class class end  | 
| 
37384
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
615  | 
and ensure_classrel thy algbr eqngr permissive (sub_class, super_class) =  | 
| 24918 | 616  | 
let  | 
617  | 
val stmt_classrel =  | 
|
| 
37384
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
618  | 
ensure_class thy algbr eqngr permissive sub_class  | 
| 
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
619  | 
##>> ensure_class thy algbr eqngr permissive super_class  | 
| 24918 | 620  | 
#>> Classrel;  | 
| 
37384
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
621  | 
in ensure_stmt lookup_classrel (declare_classrel thy) stmt_classrel (sub_class, super_class) end  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
622  | 
and ensure_inst thy (algbr as (_, algebra)) eqngr permissive (class, tyco) =  | 
| 24918 | 623  | 
let  | 
| 
37384
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
624  | 
val super_classes = (Sorts.minimize_sort algebra o Sorts.super_classes algebra) class;  | 
| 
37448
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
625  | 
val these_classparams = these o try (#params o AxClass.get_info thy);  | 
| 
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
626  | 
val classparams = these_classparams class;  | 
| 
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
627  | 
val further_classparams = maps these_classparams  | 
| 
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
628  | 
((Sorts.complete_sort algebra o Sorts.super_classes algebra) class);  | 
| 24918 | 629  | 
val vs = Name.names Name.context "'a" (Sorts.mg_domain algebra tyco [class]);  | 
630  | 
val sorts' = Sorts.mg_domain (Sign.classes_of thy) tyco [class];  | 
|
631  | 
val vs' = map2 (fn (v, sort1) => fn sort2 => (v,  | 
|
632  | 
Sorts.inter_sort (Sign.classes_of thy) (sort1, sort2))) vs sorts';  | 
|
633  | 
val arity_typ = Type (tyco, map TFree vs);  | 
|
634  | 
val arity_typ' = Type (tyco, map (fn (v, sort) => TVar ((v, 0), sort)) vs');  | 
|
| 
37384
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
635  | 
fun translate_super_instance super_class =  | 
| 
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
636  | 
ensure_class thy algbr eqngr permissive super_class  | 
| 
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
637  | 
##>> ensure_classrel thy algbr eqngr permissive (class, super_class)  | 
| 
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
638  | 
##>> translate_dicts thy algbr eqngr permissive NONE (arity_typ, [super_class])  | 
| 
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
639  | 
#>> (fn ((super_class, classrel), [DictConst (inst, dss)]) =>  | 
| 
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
640  | 
(super_class, (classrel, (inst, dss))));  | 
| 
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
641  | 
fun translate_classparam_instance (c, ty) =  | 
| 24918 | 642  | 
let  | 
| 
37384
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
643  | 
val raw_const = Const (c, map_type_tfree (K arity_typ') ty);  | 
| 
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
644  | 
val thm = AxClass.unoverload_conv thy (Thm.cterm_of thy raw_const);  | 
| 
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
645  | 
val const = (apsnd Logic.unvarifyT_global o dest_Const o snd  | 
| 24918 | 646  | 
o Logic.dest_equals o Thm.prop_of) thm;  | 
647  | 
in  | 
|
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
648  | 
ensure_const thy algbr eqngr permissive c  | 
| 
37384
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
649  | 
##>> translate_const thy algbr eqngr permissive (SOME thm) (const, NONE)  | 
| 
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
650  | 
#>> (fn (c, IConst const') => ((c, const'), (thm, true)))  | 
| 24918 | 651  | 
end;  | 
652  | 
val stmt_inst =  | 
|
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
653  | 
ensure_class thy algbr eqngr permissive class  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
654  | 
##>> ensure_tyco thy algbr eqngr permissive tyco  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
655  | 
##>> fold_map (translate_tyvar_sort thy algbr eqngr permissive) vs  | 
| 
37384
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
656  | 
##>> fold_map translate_super_instance super_classes  | 
| 
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
657  | 
##>> fold_map translate_classparam_instance classparams  | 
| 
37448
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
658  | 
##>> fold_map translate_classparam_instance further_classparams  | 
| 
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
659  | 
#>> (fn (((((class, tyco), arity_args), super_instances),  | 
| 
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
660  | 
classparam_instances), further_classparam_instances) =>  | 
| 
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
661  | 
Classinst ((class, (tyco, arity_args)), (super_instances,  | 
| 
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
662  | 
(classparam_instances, further_classparam_instances))));  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
663  | 
in ensure_stmt lookup_instance (declare_instance thy) stmt_inst (class, tyco) end  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
664  | 
and translate_typ thy algbr eqngr permissive (TFree (v, _)) =  | 
| 30932 | 665  | 
pair (ITyVar (unprefix "'" v))  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
666  | 
| translate_typ thy algbr eqngr permissive (Type (tyco, tys)) =  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
667  | 
ensure_tyco thy algbr eqngr permissive tyco  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
668  | 
##>> fold_map (translate_typ thy algbr eqngr permissive) tys  | 
| 30932 | 669  | 
#>> (fn (tyco, tys) => tyco `%% tys)  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
670  | 
and translate_term thy algbr eqngr permissive some_thm (Const (c, ty), some_abs) =  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
671  | 
translate_app thy algbr eqngr permissive some_thm (((c, ty), []), some_abs)  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
672  | 
| translate_term thy algbr eqngr permissive some_thm (Free (v, _), some_abs) =  | 
| 31889 | 673  | 
pair (IVar (SOME v))  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
674  | 
| translate_term thy algbr eqngr permissive some_thm (Abs (v, ty, t), some_abs) =  | 
| 24918 | 675  | 
let  | 
| 32273 | 676  | 
val (v', t') = Syntax.variant_abs (Name.desymbolize false v, ty, t);  | 
677  | 
val v'' = if member (op =) (Term.add_free_names t' []) v'  | 
|
678  | 
then SOME v' else NONE  | 
|
| 24918 | 679  | 
in  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
680  | 
translate_typ thy algbr eqngr permissive ty  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
681  | 
##>> translate_term thy algbr eqngr permissive some_thm (t', some_abs)  | 
| 32273 | 682  | 
#>> (fn (ty, t) => (v'', ty) `|=> t)  | 
| 24918 | 683  | 
end  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
684  | 
| translate_term thy algbr eqngr permissive some_thm (t as _ $ _, some_abs) =  | 
| 24918 | 685  | 
case strip_comb t  | 
686  | 
of (Const (c, ty), ts) =>  | 
|
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
687  | 
translate_app thy algbr eqngr permissive some_thm (((c, ty), ts), some_abs)  | 
| 24918 | 688  | 
| (t', ts) =>  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
689  | 
translate_term thy algbr eqngr permissive some_thm (t', some_abs)  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
690  | 
##>> fold_map (translate_term thy algbr eqngr permissive some_thm o rpair NONE) ts  | 
| 24918 | 691  | 
#>> (fn (t, ts) => t `$$ ts)  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
692  | 
and translate_eqn thy algbr eqngr permissive ((args, (rhs, some_abs)), (some_thm, proper)) =  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
693  | 
fold_map (translate_term thy algbr eqngr permissive some_thm) args  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
694  | 
##>> translate_term thy algbr eqngr permissive some_thm (rhs, some_abs)  | 
| 35226 | 695  | 
#>> rpair (some_thm, proper)  | 
| 
37440
 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 
haftmann 
parents: 
37437 
diff
changeset
 | 
696  | 
and translate_eqns thy algbr eqngr permissive eqns prgrm =  | 
| 
 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 
haftmann 
parents: 
37437 
diff
changeset
 | 
697  | 
prgrm |> fold_map (translate_eqn thy algbr eqngr permissive) eqns  | 
| 
 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 
haftmann 
parents: 
37437 
diff
changeset
 | 
698  | 
handle PERMISSIVE () => ([], prgrm)  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
699  | 
and translate_const thy algbr eqngr permissive some_thm ((c, ty), some_abs) =  | 
| 30932 | 700  | 
let  | 
| 37698 | 701  | 
val _ = if (case some_abs of NONE => true | SOME abs => not (c = abs))  | 
| 35226 | 702  | 
andalso Code.is_abstr thy c  | 
| 37698 | 703  | 
then translation_error thy permissive some_thm  | 
704  | 
          "Abstraction violation" ("constant " ^ Code.string_of_const thy c)
 | 
|
705  | 
else ()  | 
|
| 
37448
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
706  | 
val arg_typs = Sign.const_typargs thy (c, ty);  | 
| 32873 | 707  | 
val sorts = Code_Preproc.sortargs eqngr c;  | 
| 
37448
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
708  | 
val function_typs = (fst o Term.strip_type) ty;  | 
| 26972 | 709  | 
in  | 
| 37698 | 710  | 
ensure_const thy algbr eqngr permissive c  | 
| 
37448
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
711  | 
##>> fold_map (translate_typ thy algbr eqngr permissive) arg_typs  | 
| 
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
712  | 
##>> fold_map (translate_dicts thy algbr eqngr permissive some_thm) (arg_typs ~~ sorts)  | 
| 
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
713  | 
##>> fold_map (translate_typ thy algbr eqngr permissive) function_typs  | 
| 
 
3bd4b3809bee
explicit type variable arguments for constructors
 
haftmann 
parents: 
37447 
diff
changeset
 | 
714  | 
#>> (fn (((c, arg_typs), dss), function_typs) => IConst (c, ((arg_typs, dss), function_typs)))  | 
| 26972 | 715  | 
end  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
716  | 
and translate_app_const thy algbr eqngr permissive some_thm ((c_ty, ts), some_abs) =  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
717  | 
translate_const thy algbr eqngr permissive some_thm (c_ty, some_abs)  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
718  | 
##>> fold_map (translate_term thy algbr eqngr permissive some_thm o rpair NONE) ts  | 
| 24918 | 719  | 
#>> (fn (t, ts) => t `$$ ts)  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
720  | 
and translate_case thy algbr eqngr permissive some_thm (num_args, (t_pos, case_pats)) (c_ty, ts) =  | 
| 24918 | 721  | 
let  | 
| 31892 | 722  | 
fun arg_types num_args ty = (fst o chop num_args o fst o strip_type) ty;  | 
723  | 
val tys = arg_types num_args (snd c_ty);  | 
|
| 
29952
 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 
haftmann 
parents: 
29277 
diff
changeset
 | 
724  | 
val ty = nth tys t_pos;  | 
| 31957 | 725  | 
fun mk_constr c t = let val n = Code.args_number thy c  | 
726  | 
in ((c, arg_types n (fastype_of t) ---> ty), n) end;  | 
|
| 31892 | 727  | 
val constrs = if null case_pats then []  | 
728  | 
else map2 mk_constr case_pats (nth_drop t_pos ts);  | 
|
729  | 
fun casify naming constrs ty ts =  | 
|
| 24918 | 730  | 
let  | 
| 
31935
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
731  | 
val undefineds = map_filter (lookup_const naming) (Code.undefineds thy);  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
732  | 
fun collapse_clause vs_map ts body =  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
733  | 
let  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
734  | 
in case body  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
735  | 
of IConst (c, _) => if member (op =) undefineds c  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
736  | 
then []  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
737  | 
else [(ts, body)]  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
738  | 
| ICase (((IVar (SOME v), _), subclauses), _) =>  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
739  | 
if forall (fn (pat', body') => exists_var pat' v  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
740  | 
orelse not (exists_var body' v)) subclauses  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
741  | 
then case AList.lookup (op =) vs_map v  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
742  | 
of SOME i => maps (fn (pat', body') =>  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
743  | 
collapse_clause (AList.delete (op =) v vs_map)  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
744  | 
(nth_map i (K pat') ts) body') subclauses  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
745  | 
| NONE => [(ts, body)]  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
746  | 
else [(ts, body)]  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
747  | 
| _ => [(ts, body)]  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
748  | 
end;  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
749  | 
fun mk_clause mk tys t =  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
750  | 
let  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
751  | 
val (vs, body) = unfold_abs_eta tys t;  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
752  | 
val vs_map = fold_index (fn (i, (SOME v, _)) => cons (v, i) | _ => I) vs [];  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
753  | 
val ts = map (IVar o fst) vs;  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
754  | 
in map mk (collapse_clause vs_map ts body) end;  | 
| 31892 | 755  | 
val t = nth ts t_pos;  | 
756  | 
val ts_clause = nth_drop t_pos ts;  | 
|
| 
31935
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
757  | 
val clauses = if null case_pats  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
758  | 
then mk_clause (fn ([t], body) => (t, body)) [ty] (the_single ts_clause)  | 
| 
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
759  | 
else maps (fn ((constr as IConst (_, (_, tys)), n), t) =>  | 
| 33957 | 760  | 
mk_clause (fn (ts, body) => (constr `$$ ts, body)) (take n tys) t)  | 
| 
31935
 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 
haftmann 
parents: 
31892 
diff
changeset
 | 
761  | 
(constrs ~~ ts_clause);  | 
| 31892 | 762  | 
in ((t, ty), clauses) end;  | 
| 24918 | 763  | 
in  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
764  | 
translate_const thy algbr eqngr permissive some_thm (c_ty, NONE)  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
765  | 
##>> fold_map (fn (constr, n) => translate_const thy algbr eqngr permissive some_thm (constr, NONE)  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
766  | 
#>> rpair n) constrs  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
767  | 
##>> translate_typ thy algbr eqngr permissive ty  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
768  | 
##>> fold_map (translate_term thy algbr eqngr permissive some_thm o rpair NONE) ts  | 
| 31892 | 769  | 
#-> (fn (((t, constrs), ty), ts) =>  | 
770  | 
`(fn (_, (naming, _)) => ICase (casify naming constrs ty ts, t `$$ ts)))  | 
|
| 24918 | 771  | 
end  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
772  | 
and translate_app_case thy algbr eqngr permissive some_thm (case_scheme as (num_args, _)) ((c, ty), ts) =  | 
| 29973 | 773  | 
if length ts < num_args then  | 
774  | 
let  | 
|
775  | 
val k = length ts;  | 
|
| 33957 | 776  | 
val tys = (take (num_args - k) o drop k o fst o strip_type) ty;  | 
| 29973 | 777  | 
val ctxt = (fold o fold_aterms) Term.declare_term_frees ts Name.context;  | 
778  | 
val vs = Name.names ctxt "a" tys;  | 
|
779  | 
in  | 
|
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
780  | 
fold_map (translate_typ thy algbr eqngr permissive) tys  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
781  | 
##>> translate_case thy algbr eqngr permissive some_thm case_scheme ((c, ty), ts @ map Free vs)  | 
| 31888 | 782  | 
#>> (fn (tys, t) => map2 (fn (v, _) => pair (SOME v)) vs tys `|==> t)  | 
| 29973 | 783  | 
end  | 
784  | 
else if length ts > num_args then  | 
|
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
785  | 
translate_case thy algbr eqngr permissive some_thm case_scheme ((c, ty), take num_args ts)  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
786  | 
##>> fold_map (translate_term thy algbr eqngr permissive some_thm o rpair NONE) (drop num_args ts)  | 
| 29973 | 787  | 
#>> (fn (t, ts) => t `$$ ts)  | 
788  | 
else  | 
|
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
789  | 
translate_case thy algbr eqngr permissive some_thm case_scheme ((c, ty), ts)  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
790  | 
and translate_app thy algbr eqngr permissive some_thm (c_ty_ts as ((c, _), _), some_abs) =  | 
| 29973 | 791  | 
case Code.get_case_scheme thy c  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
792  | 
of SOME case_scheme => translate_app_case thy algbr eqngr permissive some_thm case_scheme c_ty_ts  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
793  | 
| NONE => translate_app_const thy algbr eqngr permissive some_thm (c_ty_ts, some_abs)  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
794  | 
and translate_tyvar_sort thy (algbr as (proj_sort, _)) eqngr permissive (v, sort) =  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
795  | 
fold_map (ensure_class thy algbr eqngr permissive) (proj_sort sort)  | 
| 30932 | 796  | 
#>> (fn sort => (unprefix "'" v, sort))  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
797  | 
and translate_dicts thy (algbr as (proj_sort, algebra)) eqngr permissive some_thm (ty, sort) =  | 
| 30932 | 798  | 
let  | 
799  | 
datatype typarg =  | 
|
800  | 
Global of (class * string) * typarg list list  | 
|
801  | 
| Local of (class * class) list * (string * (int * sort));  | 
|
802  | 
fun class_relation (Global ((_, tyco), yss), _) class =  | 
|
803  | 
Global ((class, tyco), yss)  | 
|
| 
37384
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
804  | 
| class_relation (Local (classrels, v), sub_class) super_class =  | 
| 
 
5aba26803073
more consistent naming aroud type classes and instances
 
haftmann 
parents: 
37216 
diff
changeset
 | 
805  | 
Local ((sub_class, super_class) :: classrels, v);  | 
| 
36102
 
a51d1d154c71
of_sort_derivation: pass-through full type information -- following the version by krauss/schropp;
 
wenzelm 
parents: 
35961 
diff
changeset
 | 
806  | 
fun type_constructor (tyco, _) yss class =  | 
| 30932 | 807  | 
Global ((class, tyco), (map o map) fst yss);  | 
808  | 
fun type_variable (TFree (v, sort)) =  | 
|
809  | 
let  | 
|
810  | 
val sort' = proj_sort sort;  | 
|
811  | 
in map_index (fn (n, class) => (Local ([], (v, (n, sort'))), class)) sort' end;  | 
|
| 37698 | 812  | 
val typargs = Sorts.of_sort_derivation algebra  | 
| 
36102
 
a51d1d154c71
of_sort_derivation: pass-through full type information -- following the version by krauss/schropp;
 
wenzelm 
parents: 
35961 
diff
changeset
 | 
813  | 
      {class_relation = K (Sorts.classrel_derivation algebra class_relation),
 | 
| 
35961
 
00e48e1d9afd
Sorts.of_sort_derivation: do not use slow Graph.irreducible_paths here, which not always needed  (SUBTLE CHANGE IN SEMANTICS);
 
wenzelm 
parents: 
35845 
diff
changeset
 | 
814  | 
type_constructor = type_constructor,  | 
| 37698 | 815  | 
type_variable = type_variable} (ty, proj_sort sort)  | 
816  | 
handle Sorts.CLASS_ERROR e => not_wellsorted thy permissive some_thm ty sort e;  | 
|
| 30932 | 817  | 
fun mk_dict (Global (inst, yss)) =  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
818  | 
ensure_inst thy algbr eqngr permissive inst  | 
| 30932 | 819  | 
##>> (fold_map o fold_map) mk_dict yss  | 
820  | 
#>> (fn (inst, dss) => DictConst (inst, dss))  | 
|
| 31962 | 821  | 
| mk_dict (Local (classrels, (v, (n, sort)))) =  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
822  | 
fold_map (ensure_classrel thy algbr eqngr permissive) classrels  | 
| 31962 | 823  | 
#>> (fn classrels => DictVar (classrels, (unprefix "'" v, (n, length sort))))  | 
| 37698 | 824  | 
in fold_map mk_dict typargs end;  | 
| 24918 | 825  | 
|
| 25969 | 826  | 
|
| 27103 | 827  | 
(* store *)  | 
828  | 
||
| 
34173
 
458ced35abb8
reduced code generator cache to the baremost minimum
 
haftmann 
parents: 
34084 
diff
changeset
 | 
829  | 
structure Program = Code_Data  | 
| 27103 | 830  | 
(  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
831  | 
type T = naming * program;  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
832  | 
val empty = (empty_naming, Graph.empty);  | 
| 27103 | 833  | 
);  | 
834  | 
||
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
835  | 
fun cache_generation thy (algebra, eqngr) f name =  | 
| 34251 | 836  | 
Program.change_yield thy (fn naming_program => (NONE, naming_program)  | 
| 32873 | 837  | 
|> f thy algebra eqngr name  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
838  | 
|-> (fn name => fn (_, naming_program) => (name, naming_program)));  | 
| 27103 | 839  | 
|
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
840  | 
fun transient_generation thy (algebra, eqngr) f name =  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
841  | 
(NONE, (empty_naming, Graph.empty))  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
842  | 
|> f thy algebra eqngr name  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
843  | 
|-> (fn name => fn (_, naming_program) => (name, naming_program));  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
844  | 
|
| 27103 | 845  | 
|
846  | 
(* program generation *)  | 
|
847  | 
||
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
848  | 
fun consts_program thy permissive cs =  | 
| 27103 | 849  | 
let  | 
| 
38672
 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 
haftmann 
parents: 
38669 
diff
changeset
 | 
850  | 
fun project_consts cs (naming, program) = (*FIXME only necessary for cache_generation*)  | 
| 27103 | 851  | 
let  | 
852  | 
val cs_all = Graph.all_succs program cs;  | 
|
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
853  | 
in (cs, (naming, Graph.subgraph (member (op =) cs_all) program)) end;  | 
| 32873 | 854  | 
fun generate_consts thy algebra eqngr =  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
855  | 
fold_map (ensure_const thy algebra eqngr permissive);  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
856  | 
val invoke_generation = if permissive  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
857  | 
then transient_generation else cache_generation  | 
| 27103 | 858  | 
in  | 
| 
31125
 
80218ee73167
transferred code generator preprocessor into separate module
 
haftmann 
parents: 
31088 
diff
changeset
 | 
859  | 
invoke_generation thy (Code_Preproc.obtain thy cs []) generate_consts cs  | 
| 27103 | 860  | 
|-> project_consts  | 
861  | 
end;  | 
|
862  | 
||
863  | 
||
864  | 
(* value evaluation *)  | 
|
| 25969 | 865  | 
|
| 32873 | 866  | 
fun ensure_value thy algbr eqngr t =  | 
| 24918 | 867  | 
let  | 
868  | 
val ty = fastype_of t;  | 
|
869  | 
val vs = fold_term_types (K (fold_atyps (insert (eq_fst op =)  | 
|
870  | 
o dest_TFree))) t [];  | 
|
871  | 
val stmt_value =  | 
|
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
872  | 
fold_map (translate_tyvar_sort thy algbr eqngr false) vs  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
873  | 
##>> translate_typ thy algbr eqngr false ty  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
874  | 
##>> translate_term thy algbr eqngr false NONE (Code.subst_signatures thy t, NONE)  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
875  | 
#>> (fn ((vs, ty), t) => Fun  | 
| 37437 | 876  | 
(Term.dummy_patternN, (((vs, ty), [(([], t), (NONE, true))]), NONE)));  | 
| 
31063
 
88aaab83b6fc
dropped explicit suppport for frees in evaluation conversion stack
 
haftmann 
parents: 
31054 
diff
changeset
 | 
877  | 
fun term_value (dep, (naming, program1)) =  | 
| 25969 | 878  | 
let  | 
| 37437 | 879  | 
val Fun (_, ((vs_ty, [(([], t), _)]), _)) =  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
880  | 
Graph.get_node program1 Term.dummy_patternN;  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
881  | 
val deps = Graph.imm_succs program1 Term.dummy_patternN;  | 
| 
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
882  | 
val program2 = Graph.del_nodes [Term.dummy_patternN] program1;  | 
| 27103 | 883  | 
val deps_all = Graph.all_succs program2 deps;  | 
884  | 
val program3 = Graph.subgraph (member (op =) deps_all) program2;  | 
|
| 
31063
 
88aaab83b6fc
dropped explicit suppport for frees in evaluation conversion stack
 
haftmann 
parents: 
31054 
diff
changeset
 | 
885  | 
in (((naming, program3), ((vs_ty, t), deps)), (dep, (naming, program2))) end;  | 
| 26011 | 886  | 
in  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28423 
diff
changeset
 | 
887  | 
ensure_stmt ((K o K) NONE) pair stmt_value Term.dummy_patternN  | 
| 26011 | 888  | 
#> snd  | 
| 
31063
 
88aaab83b6fc
dropped explicit suppport for frees in evaluation conversion stack
 
haftmann 
parents: 
31054 
diff
changeset
 | 
889  | 
#> term_value  | 
| 26011 | 890  | 
end;  | 
| 24219 | 891  | 
|
| 32873 | 892  | 
fun base_evaluator thy evaluator algebra eqngr vs t =  | 
| 
30942
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
893  | 
let  | 
| 
31063
 
88aaab83b6fc
dropped explicit suppport for frees in evaluation conversion stack
 
haftmann 
parents: 
31054 
diff
changeset
 | 
894  | 
val (((naming, program), (((vs', ty'), t'), deps)), _) =  | 
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
895  | 
cache_generation thy (algebra, eqngr) ensure_value t;  | 
| 30947 | 896  | 
val vs'' = map (fn (v, _) => (v, (the o AList.lookup (op =) vs o prefix "'") v)) vs';  | 
| 
31063
 
88aaab83b6fc
dropped explicit suppport for frees in evaluation conversion stack
 
haftmann 
parents: 
31054 
diff
changeset
 | 
897  | 
in evaluator naming program ((vs'', (vs', ty')), t') deps end;  | 
| 
30942
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
898  | 
|
| 
38669
 
9ff76d0f0610
refined and unified naming convention for dynamic code evaluation techniques
 
haftmann 
parents: 
37871 
diff
changeset
 | 
899  | 
fun dynamic_eval_conv thy = Code_Preproc.dynamic_eval_conv thy o base_evaluator thy;  | 
| 
 
9ff76d0f0610
refined and unified naming convention for dynamic code evaluation techniques
 
haftmann 
parents: 
37871 
diff
changeset
 | 
900  | 
fun dynamic_eval_value thy postproc = Code_Preproc.dynamic_eval_value thy postproc o base_evaluator thy;  | 
| 
30942
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
901  | 
|
| 
38672
 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 
haftmann 
parents: 
38669 
diff
changeset
 | 
902  | 
fun static_eval_conv thy consts conv =  | 
| 
 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 
haftmann 
parents: 
38669 
diff
changeset
 | 
903  | 
Code_Preproc.static_eval_conv thy consts (base_evaluator thy conv); (*FIXME avoid re-generation*)  | 
| 
 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 
haftmann 
parents: 
38669 
diff
changeset
 | 
904  | 
|
| 
 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 
haftmann 
parents: 
38669 
diff
changeset
 | 
905  | 
fun static_eval_conv_simple thy consts conv =  | 
| 
 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 
haftmann 
parents: 
38669 
diff
changeset
 | 
906  | 
Code_Preproc.static_eval_conv thy consts (fn algebra => fn eqngr => fn _ => fn _ => fn ct =>  | 
| 
 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 
haftmann 
parents: 
38669 
diff
changeset
 | 
907  | 
conv ((NONE, (empty_naming, Graph.empty)) (*FIXME provide abstraction for this kind of invocation*)  | 
| 
 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 
haftmann 
parents: 
38669 
diff
changeset
 | 
908  | 
|> fold_map (ensure_const thy algebra eqngr false) consts  | 
| 
 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 
haftmann 
parents: 
38669 
diff
changeset
 | 
909  | 
|> (snd o snd o snd)) ct);  | 
| 
 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 
haftmann 
parents: 
38669 
diff
changeset
 | 
910  | 
|
| 
30942
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
911  | 
|
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
912  | 
(** diagnostic commands **)  | 
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
913  | 
|
| 31036 | 914  | 
fun read_const_exprs thy =  | 
915  | 
let  | 
|
| 
36272
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
916  | 
fun consts_of thy' = Symtab.fold (fn (c, (_, NONE)) => cons c | _ => I)  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
917  | 
((snd o #constants o Consts.dest o #consts o Sign.rep_sg) thy') [];  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
918  | 
fun belongs_here thy' c = forall  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
919  | 
(fn thy'' => not (Sign.declared_const thy'' c)) (Theory.parents_of thy');  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
920  | 
fun consts_of_select thy' = filter (belongs_here thy') (consts_of thy');  | 
| 
 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 
haftmann 
parents: 
36210 
diff
changeset
 | 
921  | 
fun read_const_expr "*" = ([], consts_of thy)  | 
| 31036 | 922  | 
| read_const_expr s = if String.isSuffix ".*" s  | 
| 
37871
 
c7ce7685e087
replaced Thy_Info.the_theory by Context.this_theory -- avoid referring to accidental theory loader state;
 
wenzelm 
parents: 
37841 
diff
changeset
 | 
923  | 
then ([], consts_of_select (Context.this_theory thy (unsuffix ".*" s)))  | 
| 31156 | 924  | 
else ([Code.read_const thy s], []);  | 
| 31036 | 925  | 
in pairself flat o split_list o map read_const_expr end;  | 
926  | 
||
| 
30942
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
927  | 
fun code_depgr thy consts =  | 
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
928  | 
let  | 
| 
31125
 
80218ee73167
transferred code generator preprocessor into separate module
 
haftmann 
parents: 
31088 
diff
changeset
 | 
929  | 
val (_, eqngr) = Code_Preproc.obtain thy consts [];  | 
| 
34173
 
458ced35abb8
reduced code generator cache to the baremost minimum
 
haftmann 
parents: 
34084 
diff
changeset
 | 
930  | 
val all_consts = Graph.all_succs eqngr consts;  | 
| 
34891
 
99b9a6290446
code certificates as integral part of code generation
 
haftmann 
parents: 
34251 
diff
changeset
 | 
931  | 
in Graph.subgraph (member (op =) all_consts) eqngr end;  | 
| 
30942
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
932  | 
|
| 
31125
 
80218ee73167
transferred code generator preprocessor into separate module
 
haftmann 
parents: 
31088 
diff
changeset
 | 
933  | 
fun code_thms thy = Pretty.writeln o Code_Preproc.pretty thy o code_depgr thy;  | 
| 
30942
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
934  | 
|
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
935  | 
fun code_deps thy consts =  | 
| 27103 | 936  | 
let  | 
| 
30942
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
937  | 
val eqngr = code_depgr thy consts;  | 
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
938  | 
val constss = Graph.strong_conn eqngr;  | 
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
939  | 
val mapping = Symtab.empty |> fold (fn consts => fold (fn const =>  | 
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
940  | 
Symtab.update (const, consts)) consts) constss;  | 
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
941  | 
fun succs consts = consts  | 
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
942  | 
|> maps (Graph.imm_succs eqngr)  | 
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
943  | 
|> subtract (op =) consts  | 
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
944  | 
|> map (the o Symtab.lookup mapping)  | 
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
945  | 
|> distinct (op =);  | 
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
946  | 
val conn = [] |> fold (fn consts => cons (consts, succs consts)) constss;  | 
| 31156 | 947  | 
fun namify consts = map (Code.string_of_const thy) consts  | 
| 
30942
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
948  | 
|> commas;  | 
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
949  | 
val prgr = map (fn (consts, constss) =>  | 
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
950  | 
      { name = namify consts, ID = namify consts, dir = "", unfold = true,
 | 
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
951  | 
path = "", parents = map namify constss }) conn;  | 
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
952  | 
in Present.display_graph prgr end;  | 
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
953  | 
|
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
954  | 
local  | 
| 27103 | 955  | 
|
| 31036 | 956  | 
fun code_thms_cmd thy = code_thms thy o op @ o read_const_exprs thy;  | 
957  | 
fun code_deps_cmd thy = code_deps thy o op @ o read_const_exprs thy;  | 
|
| 
30942
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
958  | 
|
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
959  | 
in  | 
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
960  | 
|
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
961  | 
val _ =  | 
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36272 
diff
changeset
 | 
962  | 
Outer_Syntax.improper_command "code_thms" "print system of code equations for code" Keyword.diag  | 
| 
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36272 
diff
changeset
 | 
963  | 
(Scan.repeat1 Parse.term_group  | 
| 
30942
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
964  | 
>> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory  | 
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
965  | 
o Toplevel.keep ((fn thy => code_thms_cmd thy cs) o Toplevel.theory_of)));  | 
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
966  | 
|
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
967  | 
val _ =  | 
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36272 
diff
changeset
 | 
968  | 
Outer_Syntax.improper_command "code_deps" "visualize dependencies of code equations for code"  | 
| 
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36272 
diff
changeset
 | 
969  | 
Keyword.diag  | 
| 
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36272 
diff
changeset
 | 
970  | 
(Scan.repeat1 Parse.term_group  | 
| 
30942
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
971  | 
>> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory  | 
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
972  | 
o Toplevel.keep ((fn thy => code_deps_cmd thy cs) o Toplevel.theory_of)));  | 
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
973  | 
|
| 
 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 
haftmann 
parents: 
30932 
diff
changeset
 | 
974  | 
end;  | 
| 27103 | 975  | 
|
| 24219 | 976  | 
end; (*struct*)  | 
977  | 
||
978  | 
||
| 28054 | 979  | 
structure Basic_Code_Thingol: BASIC_CODE_THINGOL = Code_Thingol;  |