| author | wenzelm | 
| Sun, 18 Mar 2012 13:04:22 +0100 | |
| changeset 47005 | 421760a1efe7 | 
| parent 46961 | 5c6955f487e5 | 
| child 47437 | 4625ee486ff6 | 
| 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 = | |
| 41782 | 18 | Dict of string list * plain_dict | 
| 41118 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 19 | and plain_dict = | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 20 | Dict_Const of string * dict list list | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 21 | | Dict_Var of vname * (int * int) | 
| 24219 | 22 | datatype itype = | 
| 23 | `%% of string * itype list | |
| 24 | | ITyVar of vname; | |
| 44789 
5a062c23c7db
adding the body type as well to the code generation for constants as it is required for type annotations of constants
 bulwahn parents: 
44788diff
changeset | 25 | type const = string * (((itype list * dict list list) * (itype list * itype)) * bool) | 
| 45009 
99e1965f9c21
syntactic improvements and tuning names in the code generator due to Florian's code review
 bulwahn parents: 
45000diff
changeset | 26 |     (* (f [T1..Tn] {dicts} (_::S1) .. (_::Sm)) :: S =^= (f, ((([T1..Tn], dicts), ([S1..Sm], S)), ambiguous)) *)
 | 
| 24219 | 27 | datatype iterm = | 
| 24591 | 28 | IConst of const | 
| 31889 | 29 | | IVar of vname option | 
| 24219 | 30 | | `$ of iterm * iterm | 
| 31888 | 31 | | `|=> of (vname option * itype) * iterm | 
| 24219 | 32 | | ICase of ((iterm * itype) * (iterm * iterm) list) * iterm; | 
| 33 | (*((term, type), [(selector pattern, body term )]), primitive term)*) | |
| 34 | val `$$ : iterm * iterm list -> iterm; | |
| 31888 | 35 | val `|==> : (vname option * itype) list * iterm -> iterm; | 
| 24219 | 36 | type typscheme = (vname * sort) list * itype; | 
| 37 | end; | |
| 38 | ||
| 39 | signature CODE_THINGOL = | |
| 40 | sig | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 41 | include BASIC_CODE_THINGOL | 
| 34084 | 42 | val fun_tyco: string | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 43 |   val unfoldl: ('a -> ('a * 'b) option) -> 'a -> 'a * 'b list
 | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 44 |   val unfoldr: ('a -> ('b * 'a) option) -> 'a -> 'b list * 'a
 | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 45 | val unfold_fun: itype -> itype list * itype | 
| 37640 | 46 | val unfold_fun_n: int -> itype -> itype list * itype | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 47 | val unfold_app: iterm -> iterm * iterm list | 
| 31888 | 48 | val unfold_abs: iterm -> (vname option * itype) list * iterm | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 49 | val split_let: iterm -> (((iterm * itype) * iterm) * iterm) option | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 50 | val unfold_let: iterm -> ((iterm * itype) * iterm) list * iterm | 
| 31889 | 51 | val split_pat_abs: iterm -> ((iterm * itype) * iterm) option | 
| 52 | val unfold_pat_abs: iterm -> (iterm * itype) list * iterm | |
| 31049 | 53 | val unfold_const_app: iterm -> (const * iterm list) option | 
| 32909 | 54 | val is_IVar: iterm -> bool | 
| 41782 | 55 | val is_IAbs: iterm -> bool | 
| 31049 | 56 | val eta_expand: int -> const * iterm list -> iterm | 
| 41100 
6c0940392fb4
dictionary constants must permit explicit weakening of classes;
 haftmann parents: 
40844diff
changeset | 57 | val contains_dict_var: iterm -> bool | 
| 31935 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 58 | val add_constnames: iterm -> string list -> string list | 
| 32917 | 59 | val add_tyconames: iterm -> string list -> string list | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 60 | val fold_varnames: (string -> 'a -> 'a) -> iterm -> 'a -> 'a | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 61 | |
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 62 | type naming | 
| 28706 | 63 | val empty_naming: naming | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 64 | val lookup_class: naming -> class -> string option | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 65 | val lookup_classrel: naming -> class * class -> string option | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 66 | val lookup_tyco: naming -> string -> string option | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 67 | val lookup_instance: naming -> class * string -> string option | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 68 | val lookup_const: naming -> string -> string option | 
| 31054 | 69 | val ensure_declared_const: theory -> string -> naming -> string * naming | 
| 24219 | 70 | |
| 24918 | 71 | datatype stmt = | 
| 27024 | 72 | NoStmt | 
| 37437 | 73 | | Fun of string * ((typscheme * ((iterm list * iterm) * (thm option * bool)) list) * thm option) | 
| 37448 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 74 | | Datatype of string * ((vname * sort) list * | 
| 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 75 | ((string * vname list (*type argument wrt. canonical order*)) * itype list) list) | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 76 | | Datatypecons of string * string | 
| 37447 
ad3e04f289b6
transitive superclasses were also only a misunderstanding
 haftmann parents: 
37446diff
changeset | 77 | | Class of class * (vname * ((class * string) list * (string * itype) list)) | 
| 24219 | 78 | | Classrel of class * class | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 79 | | Classparam of string * class | 
| 37448 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 80 | | Classinst of (class * (string * (vname * sort) list) (*class and arity*)) | 
| 37445 
e372fa3c7239
dropped obscure type argument weakening mapping -- was only a misunderstanding
 haftmann parents: 
37440diff
changeset | 81 | * ((class * (string * (string * dict list list))) list (*super instances*) | 
| 37448 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 82 | * (((string * const) * (thm * bool)) list (*class parameter instances*) | 
| 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 83 | * ((string * const) * (thm * bool)) list (*super class parameter instances*))) | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 84 | type program = stmt Graph.T | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 85 | val empty_funs: program -> string list | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 86 | val map_terms_bottom_up: (iterm -> iterm) -> iterm -> iterm | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 87 | val map_terms_stmt: (iterm -> iterm) -> stmt -> stmt | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 88 | val is_cons: program -> string -> bool | 
| 37440 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 haftmann parents: 
37437diff
changeset | 89 | val is_case: stmt -> bool | 
| 32895 | 90 | val labelled_name: theory -> program -> string -> string | 
| 91 | val group_stmts: theory -> program | |
| 92 | -> ((string * stmt) list * (string * stmt) list | |
| 93 | * ((string * stmt) list * (string * stmt) list)) list | |
| 24219 | 94 | |
| 31036 | 95 | 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: 
36210diff
changeset | 96 | val consts_program: theory -> bool -> string list -> string list * (naming * program) | 
| 41184 | 97 | val dynamic_conv: theory -> (naming -> program | 
| 39487 
80b2cf3b0779
proper closures for static evaluation; no need for FIXMEs any longer
 haftmann parents: 
39475diff
changeset | 98 | -> ((string * sort) list * typscheme) * iterm -> string list -> conv) | 
| 38672 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 haftmann parents: 
38669diff
changeset | 99 | -> conv | 
| 41184 | 100 | val dynamic_value: theory -> ((term -> term) -> 'a -> 'a) -> (naming -> program | 
| 39487 
80b2cf3b0779
proper closures for static evaluation; no need for FIXMEs any longer
 haftmann parents: 
39475diff
changeset | 101 | -> ((string * sort) list * typscheme) * iterm -> string list -> 'a) | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 102 | -> term -> 'a | 
| 41346 | 103 | val static_conv: theory -> string list -> (naming -> program -> string list | 
| 104 | -> ((string * sort) list * typscheme) * iterm -> string list -> conv) | |
| 38672 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 haftmann parents: 
38669diff
changeset | 105 | -> conv | 
| 41184 | 106 | val static_conv_simple: theory -> string list | 
| 41346 | 107 | -> (program -> (string * sort) list -> term -> conv) -> conv | 
| 41184 | 108 | val static_value: theory -> ((term -> term) -> 'a -> 'a) -> string list -> | 
| 41346 | 109 | (naming -> program -> string list | 
| 110 | -> ((string * sort) list * typscheme) * iterm -> string list -> 'a) | |
| 39475 | 111 | -> term -> 'a | 
| 24219 | 112 | end; | 
| 113 | ||
| 28054 | 114 | structure Code_Thingol: CODE_THINGOL = | 
| 24219 | 115 | struct | 
| 116 | ||
| 117 | (** auxiliary **) | |
| 118 | ||
| 119 | fun unfoldl dest x = | |
| 120 | case dest x | |
| 121 | of NONE => (x, []) | |
| 122 | | SOME (x1, x2) => | |
| 123 | let val (x', xs') = unfoldl dest x1 in (x', xs' @ [x2]) end; | |
| 124 | ||
| 125 | fun unfoldr dest x = | |
| 126 | case dest x | |
| 127 | of NONE => ([], x) | |
| 128 | | SOME (x1, x2) => | |
| 129 | let val (xs', x') = unfoldr dest x2 in (x1::xs', x') end; | |
| 130 | ||
| 131 | ||
| 29962 
bd4dc7fa742d
tuned comments, stripped ID, deleted superfluous code
 haftmann parents: 
29952diff
changeset | 132 | (** language core - types, terms **) | 
| 24219 | 133 | |
| 134 | type vname = string; | |
| 135 | ||
| 136 | datatype dict = | |
| 41782 | 137 | Dict of string list * plain_dict | 
| 41118 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 138 | and plain_dict = | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 139 | Dict_Const of string * dict list list | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 140 | | Dict_Var of vname * (int * int) | 
| 24219 | 141 | |
| 142 | datatype itype = | |
| 143 | `%% of string * itype list | |
| 144 | | ITyVar of vname; | |
| 145 | ||
| 44789 
5a062c23c7db
adding the body type as well to the code generation for constants as it is required for type annotations of constants
 bulwahn parents: 
44788diff
changeset | 146 | type const = string * (((itype list * dict list list) * | 
| 45009 
99e1965f9c21
syntactic improvements and tuning names in the code generator due to Florian's code review
 bulwahn parents: 
45000diff
changeset | 147 | (itype list (*types of arguments*) * itype (*result type*))) * bool (*requires type annotation*)) | 
| 24591 | 148 | |
| 24219 | 149 | datatype iterm = | 
| 24591 | 150 | IConst of const | 
| 31889 | 151 | | IVar of vname option | 
| 24219 | 152 | | `$ of iterm * iterm | 
| 31888 | 153 | | `|=> of (vname option * itype) * iterm | 
| 24219 | 154 | | ICase of ((iterm * itype) * (iterm * iterm) list) * iterm; | 
| 155 | (*see also signature*) | |
| 156 | ||
| 32909 | 157 | fun is_IVar (IVar _) = true | 
| 158 | | is_IVar _ = false; | |
| 159 | ||
| 41782 | 160 | fun is_IAbs (_ `|=> _) = true | 
| 161 | | is_IAbs _ = false; | |
| 162 | ||
| 24219 | 163 | val op `$$ = Library.foldl (op `$); | 
| 31724 | 164 | val op `|==> = Library.foldr (op `|=>); | 
| 24219 | 165 | |
| 166 | val unfold_app = unfoldl | |
| 167 | (fn op `$ t => SOME t | |
| 168 | | _ => NONE); | |
| 169 | ||
| 31874 | 170 | val unfold_abs = unfoldr | 
| 171 | (fn op `|=> t => SOME t | |
| 24219 | 172 | | _ => NONE); | 
| 173 | ||
| 174 | val split_let = | |
| 175 | (fn ICase (((td, ty), [(p, t)]), _) => SOME (((p, ty), td), t) | |
| 176 | | _ => NONE); | |
| 177 | ||
| 178 | val unfold_let = unfoldr split_let; | |
| 179 | ||
| 180 | fun unfold_const_app t = | |
| 181 | case unfold_app t | |
| 182 | of (IConst c, ts) => SOME (c, ts) | |
| 183 | | _ => NONE; | |
| 184 | ||
| 32917 | 185 | fun fold_constexprs f = | 
| 186 | let | |
| 187 | fun fold' (IConst c) = f c | |
| 188 | | fold' (IVar _) = I | |
| 189 | | fold' (t1 `$ t2) = fold' t1 #> fold' t2 | |
| 190 | | fold' (_ `|=> t) = fold' t | |
| 191 | | fold' (ICase (((t, _), ds), _)) = fold' t | |
| 192 | #> fold (fn (pat, body) => fold' pat #> fold' body) ds | |
| 193 | in fold' end; | |
| 194 | ||
| 195 | val add_constnames = fold_constexprs (fn (c, _) => insert (op =) c); | |
| 196 | ||
| 197 | fun add_tycos (tyco `%% tys) = insert (op =) tyco #> fold add_tycos tys | |
| 198 | | add_tycos (ITyVar _) = I; | |
| 199 | ||
| 44788 
8b935f1b3cf8
changing const type to pass along if typing annotations are necessary for disambigous terms
 bulwahn parents: 
44338diff
changeset | 200 | val add_tyconames = fold_constexprs (fn (_, (((tys, _), _), _)) => fold add_tycos tys); | 
| 24219 | 201 | |
| 202 | fun fold_varnames f = | |
| 203 | let | |
| 31935 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 204 | fun fold_aux add f = | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 205 | let | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 206 | fun fold_term _ (IConst _) = I | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 207 | | 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: 
31892diff
changeset | 208 | | fold_term _ (IVar NONE) = I | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 209 | | fold_term vs (t1 `$ t2) = fold_term vs t1 #> fold_term vs t2 | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 210 | | fold_term vs ((SOME v, _) `|=> t) = fold_term (insert (op =) v vs) t | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 211 | | fold_term vs ((NONE, _) `|=> t) = fold_term vs t | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 212 | | 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: 
31892diff
changeset | 213 | and fold_case vs (p, t) = fold_term (add p vs) t; | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 214 | in fold_term [] end; | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 215 | fun add t = fold_aux add (insert (op =)) t; | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 216 | in fold_aux add f end; | 
| 24219 | 217 | |
| 31935 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 218 | fun exists_var t v = fold_varnames (fn w => fn b => v = w orelse b) t false; | 
| 31874 | 219 | |
| 31889 | 220 | fun split_pat_abs ((NONE, ty) `|=> t) = SOME ((IVar NONE, ty), t) | 
| 31888 | 221 | | split_pat_abs ((SOME v, ty) `|=> t) = SOME (case t | 
| 31889 | 222 | of ICase (((IVar (SOME w), _), [(p, t')]), _) => | 
| 31888 | 223 | if v = w andalso (exists_var p v orelse not (exists_var t' v)) | 
| 31889 | 224 | then ((p, ty), t') | 
| 225 | else ((IVar (SOME v), ty), t) | |
| 226 | | _ => ((IVar (SOME v), ty), t)) | |
| 31888 | 227 | | split_pat_abs _ = NONE; | 
| 31874 | 228 | |
| 229 | val unfold_pat_abs = unfoldr split_pat_abs; | |
| 24219 | 230 | |
| 31890 
e943b039f0ac
an intermediate step towards a refined translation of cases
 haftmann parents: 
31889diff
changeset | 231 | fun unfold_abs_eta [] t = ([], t) | 
| 
e943b039f0ac
an intermediate step towards a refined translation of cases
 haftmann parents: 
31889diff
changeset | 232 | | unfold_abs_eta (_ :: tys) (v_ty `|=> t) = | 
| 
e943b039f0ac
an intermediate step towards a refined translation of cases
 haftmann parents: 
31889diff
changeset | 233 | let | 
| 
e943b039f0ac
an intermediate step towards a refined translation of cases
 haftmann parents: 
31889diff
changeset | 234 | val (vs_tys, t') = unfold_abs_eta tys t; | 
| 
e943b039f0ac
an intermediate step towards a refined translation of cases
 haftmann parents: 
31889diff
changeset | 235 | in (v_ty :: vs_tys, t') end | 
| 31892 | 236 | | unfold_abs_eta tys t = | 
| 31890 
e943b039f0ac
an intermediate step towards a refined translation of cases
 haftmann parents: 
31889diff
changeset | 237 | let | 
| 
e943b039f0ac
an intermediate step towards a refined translation of cases
 haftmann parents: 
31889diff
changeset | 238 | val ctxt = fold_varnames Name.declare t Name.context; | 
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43326diff
changeset | 239 | val vs_tys = (map o apfst) SOME (Name.invent_names ctxt "a" tys); | 
| 31890 
e943b039f0ac
an intermediate step towards a refined translation of cases
 haftmann parents: 
31889diff
changeset | 240 | in (vs_tys, t `$$ map (IVar o fst) vs_tys) end; | 
| 
e943b039f0ac
an intermediate step towards a refined translation of cases
 haftmann parents: 
31889diff
changeset | 241 | |
| 44789 
5a062c23c7db
adding the body type as well to the code generation for constants as it is required for type annotations of constants
 bulwahn parents: 
44788diff
changeset | 242 | fun eta_expand k (c as (name, ((_, (tys, _)), _)), ts) = | 
| 24219 | 243 | let | 
| 244 | val j = length ts; | |
| 245 | val l = k - j; | |
| 37841 | 246 | val _ = if l > length tys | 
| 247 |       then error ("Impossible eta-expansion for constant " ^ quote name) else ();
 | |
| 24219 | 248 | val ctxt = (fold o fold_varnames) Name.declare ts Name.context; | 
| 31889 | 249 | val vs_tys = (map o apfst) SOME | 
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43326diff
changeset | 250 | (Name.invent_names ctxt "a" ((take l o drop j) tys)); | 
| 31889 | 251 | in vs_tys `|==> IConst c `$$ ts @ map (IVar o fst) vs_tys end; | 
| 24219 | 252 | |
| 41100 
6c0940392fb4
dictionary constants must permit explicit weakening of classes;
 haftmann parents: 
40844diff
changeset | 253 | fun contains_dict_var t = | 
| 24662 
f79f6061525c
more precise treatment of free dictionary parameters for evaluation
 haftmann parents: 
24591diff
changeset | 254 | let | 
| 41118 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 255 | fun cont_dict (Dict (_, d)) = cont_plain_dict d | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 256 | and cont_plain_dict (Dict_Const (_, dss)) = (exists o exists) cont_dict dss | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 257 | | cont_plain_dict (Dict_Var _) = true; | 
| 44788 
8b935f1b3cf8
changing const type to pass along if typing annotations are necessary for disambigous terms
 bulwahn parents: 
44338diff
changeset | 258 | fun cont_term (IConst (_, (((_, dss), _), _))) = (exists o exists) cont_dict dss | 
| 31935 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 259 | | cont_term (IVar _) = false | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 260 | | cont_term (t1 `$ t2) = cont_term t1 orelse cont_term t2 | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 261 | | cont_term (_ `|=> t) = cont_term t | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 262 | | cont_term (ICase (_, t)) = cont_term t; | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 263 | in cont_term t end; | 
| 25621 | 264 | |
| 265 | ||
| 28688 | 266 | (** namings **) | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 267 | |
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 268 | (* policies *) | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 269 | |
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 270 | local | 
| 45188 
35870ec62ec7
removing dependency of the generic code generator to old code generator functions thyname_of_type and thyname_of_const
 bulwahn parents: 
45128diff
changeset | 271 | fun thyname_of_type thy = #theory_name o Name_Space.the_entry (Sign.type_space thy); | 
| 33172 | 272 | fun thyname_of_class thy = #theory_name o Name_Space.the_entry (Sign.class_space thy); | 
| 45188 
35870ec62ec7
removing dependency of the generic code generator to old code generator functions thyname_of_type and thyname_of_const
 bulwahn parents: 
45128diff
changeset | 273 | fun thyname_of_const' thy = #theory_name o Name_Space.the_entry (Sign.const_space thy); | 
| 33172 | 274 | fun thyname_of_instance thy inst = case AxClass.thynames_of_arity thy inst | 
| 275 |    of [] => error ("No such instance: " ^ quote (snd inst ^ " :: " ^ fst inst))
 | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 276 | | thyname :: _ => thyname; | 
| 28706 | 277 | fun thyname_of_const thy c = case AxClass.class_of_param thy c | 
| 278 | of SOME class => thyname_of_class thy class | |
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 279 | | NONE => (case Code.get_type_of_constr_or_abstr thy c | 
| 45188 
35870ec62ec7
removing dependency of the generic code generator to old code generator functions thyname_of_type and thyname_of_const
 bulwahn parents: 
45128diff
changeset | 280 | of SOME (tyco, _) => thyname_of_type thy tyco | 
| 
35870ec62ec7
removing dependency of the generic code generator to old code generator functions thyname_of_type and thyname_of_const
 bulwahn parents: 
45128diff
changeset | 281 | | NONE => thyname_of_const' thy c); | 
| 33420 | 282 | fun purify_base "==>" = "follows" | 
| 39566 | 283 | | purify_base "==" = "meta_eq" | 
| 31036 | 284 | | purify_base s = Name.desymbolize false s; | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 285 | fun namify thy get_basename get_thyname name = | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 286 | let | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 287 | val prefix = get_thyname thy name; | 
| 31036 | 288 | val base = (purify_base o get_basename) name; | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 289 | in Long_Name.append prefix base end; | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 290 | in | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 291 | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 292 | 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: 
37216diff
changeset | 293 | fun namify_classrel thy = namify thy (fn (sub_class, super_class) => | 
| 37431 | 294 | Long_Name.base_name super_class ^ "_" ^ Long_Name.base_name sub_class) | 
| 33172 | 295 | (fn thy => thyname_of_class thy o fst); | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 296 | (*order fits nicely with composed projections*) | 
| 28688 | 297 | fun namify_tyco thy "fun" = "Pure.fun" | 
| 45188 
35870ec62ec7
removing dependency of the generic code generator to old code generator functions thyname_of_type and thyname_of_const
 bulwahn parents: 
45128diff
changeset | 298 | | namify_tyco thy tyco = namify thy Long_Name.base_name thyname_of_type tyco; | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 299 | 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: 
30280diff
changeset | 300 | 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: 
30280diff
changeset | 301 | fun namify_const thy = namify thy Long_Name.base_name thyname_of_const; | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 302 | |
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 303 | end; (* local *) | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 304 | |
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 305 | |
| 28688 | 306 | (* data *) | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 307 | |
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 308 | datatype naming = Naming of {
 | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 309 | class: class Symtab.table * Name.context, | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 310 | classrel: string Symreltab.table * Name.context, | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 311 | tyco: string Symtab.table * Name.context, | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 312 | instance: string Symreltab.table * Name.context, | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 313 | const: string Symtab.table * Name.context | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 314 | } | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 315 | |
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 316 | fun dest_Naming (Naming naming) = naming; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 317 | |
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 318 | val empty_naming = Naming {
 | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 319 | class = (Symtab.empty, Name.context), | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 320 | classrel = (Symreltab.empty, Name.context), | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 321 | tyco = (Symtab.empty, Name.context), | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 322 | instance = (Symreltab.empty, Name.context), | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 323 | const = (Symtab.empty, Name.context) | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 324 | }; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 325 | |
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 326 | local | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 327 | fun mk_naming (class, classrel, tyco, instance, const) = | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 328 |     Naming { class = class, classrel = classrel,
 | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 329 | tyco = tyco, instance = instance, const = const }; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 330 |   fun map_naming f (Naming { class, classrel, tyco, instance, const }) =
 | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 331 | mk_naming (f (class, classrel, tyco, instance, const)); | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 332 | in | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 333 | fun map_class f = map_naming | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 334 | (fn (class, classrel, tyco, inst, const) => | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 335 | (f class, classrel, tyco, inst, const)); | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 336 | fun map_classrel f = map_naming | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 337 | (fn (class, classrel, tyco, inst, const) => | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 338 | (class, f classrel, tyco, inst, const)); | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 339 | fun map_tyco f = map_naming | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 340 | (fn (class, classrel, tyco, inst, const) => | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 341 | (class, classrel, f tyco, inst, const)); | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 342 | fun map_instance f = map_naming | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 343 | (fn (class, classrel, tyco, inst, const) => | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 344 | (class, classrel, tyco, f inst, const)); | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 345 | fun map_const f = map_naming | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 346 | (fn (class, classrel, tyco, inst, const) => | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 347 | (class, classrel, tyco, inst, f const)); | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 348 | end; (*local*) | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 349 | |
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 350 | fun add_variant update (thing, name) (tab, used) = | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 351 | let | 
| 43326 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 wenzelm parents: 
43048diff
changeset | 352 | val (name', used') = Name.variant name used; | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 353 | val tab' = update (thing, name') tab; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 354 | in (tab', used') end; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 355 | |
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 356 | fun declare thy mapp lookup update namify thing = | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 357 | mapp (add_variant update (thing, namify thy thing)) | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 358 | #> `(fn naming => the (lookup naming thing)); | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 359 | |
| 28688 | 360 | |
| 361 | (* lookup and declare *) | |
| 362 | ||
| 363 | local | |
| 364 | ||
| 365 | val suffix_class = "class"; | |
| 366 | val suffix_classrel = "classrel" | |
| 367 | val suffix_tyco = "tyco"; | |
| 368 | val suffix_instance = "inst"; | |
| 369 | val suffix_const = "const"; | |
| 370 | ||
| 371 | fun add_suffix nsp NONE = NONE | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 372 | | add_suffix nsp (SOME name) = SOME (Long_Name.append name nsp); | 
| 28688 | 373 | |
| 374 | in | |
| 375 | ||
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 376 | val lookup_class = add_suffix suffix_class | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 377 | oo Symtab.lookup o fst o #class o dest_Naming; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 378 | val lookup_classrel = add_suffix suffix_classrel | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 379 | oo Symreltab.lookup o fst o #classrel o dest_Naming; | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 380 | val lookup_tyco = add_suffix suffix_tyco | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 381 | oo Symtab.lookup o fst o #tyco o dest_Naming; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 382 | val lookup_instance = add_suffix suffix_instance | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 383 | oo Symreltab.lookup o fst o #instance o dest_Naming; | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 384 | val lookup_const = add_suffix suffix_const | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 385 | oo Symtab.lookup o fst o #const o dest_Naming; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 386 | |
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 387 | fun declare_class thy = declare thy map_class | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 388 | lookup_class Symtab.update_new namify_class; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 389 | fun declare_classrel thy = declare thy map_classrel | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 390 | lookup_classrel Symreltab.update_new namify_classrel; | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 391 | fun declare_tyco thy = declare thy map_tyco | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 392 | lookup_tyco Symtab.update_new namify_tyco; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 393 | fun declare_instance thy = declare thy map_instance | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 394 | lookup_instance Symreltab.update_new namify_instance; | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 395 | fun declare_const thy = declare thy map_const | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 396 | lookup_const Symtab.update_new namify_const; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 397 | |
| 31054 | 398 | fun ensure_declared_const thy const naming = | 
| 399 | case lookup_const naming const | |
| 400 | of SOME const' => (const', naming) | |
| 401 | | NONE => declare_const thy const naming; | |
| 402 | ||
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 403 | val fun_tyco = Long_Name.append (namify_tyco Pure.thy "fun") suffix_tyco | 
| 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 404 | (*depends on add_suffix*); | 
| 34084 | 405 | |
| 28688 | 406 | val unfold_fun = unfoldr | 
| 34084 | 407 | (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: 
37216diff
changeset | 408 | | _ => NONE); | 
| 28688 | 409 | |
| 37640 | 410 | fun unfold_fun_n n ty = | 
| 411 | let | |
| 412 | val (tys1, ty1) = unfold_fun ty; | |
| 413 | val (tys3, tys2) = chop n tys1; | |
| 414 | val ty3 = Library.foldr (fn (ty1, ty2) => fun_tyco `%% [ty1, ty2]) (tys2, ty1); | |
| 415 | in (tys3, ty3) end; | |
| 416 | ||
| 28688 | 417 | end; (* local *) | 
| 418 | ||
| 24219 | 419 | |
| 27103 | 420 | (** statements, abstract programs **) | 
| 24219 | 421 | |
| 422 | type typscheme = (vname * sort) list * itype; | |
| 37447 
ad3e04f289b6
transitive superclasses were also only a misunderstanding
 haftmann parents: 
37446diff
changeset | 423 | datatype stmt = | 
| 27024 | 424 | NoStmt | 
| 37437 | 425 | | Fun of string * ((typscheme * ((iterm list * iterm) * (thm option * bool)) list) * thm option) | 
| 37448 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 426 | | Datatype of string * ((vname * sort) list * ((string * vname list) * itype list) list) | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 427 | | Datatypecons of string * string | 
| 37447 
ad3e04f289b6
transitive superclasses were also only a misunderstanding
 haftmann parents: 
37446diff
changeset | 428 | | Class of class * (vname * ((class * string) list * (string * itype) list)) | 
| 24219 | 429 | | Classrel of class * class | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 430 | | Classparam of string * class | 
| 24219 | 431 | | Classinst of (class * (string * (vname * sort) list)) | 
| 37445 
e372fa3c7239
dropped obscure type argument weakening mapping -- was only a misunderstanding
 haftmann parents: 
37440diff
changeset | 432 | * ((class * (string * (string * dict list list))) list | 
| 37448 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 433 | * (((string * const) * (thm * bool)) list | 
| 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 434 | * ((string * const) * (thm * bool)) list)) | 
| 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 435 | (*see also signature*); | 
| 24219 | 436 | |
| 27103 | 437 | type program = stmt Graph.T; | 
| 24219 | 438 | |
| 27103 | 439 | fun empty_funs program = | 
| 37437 | 440 | Graph.fold (fn (name, (Fun (c, ((_, []), _)), _)) => cons c | 
| 27103 | 441 | | _ => I) program []; | 
| 24219 | 442 | |
| 27711 | 443 | fun map_terms_bottom_up f (t as IConst _) = f t | 
| 444 | | map_terms_bottom_up f (t as IVar _) = f t | |
| 445 | | map_terms_bottom_up f (t1 `$ t2) = f | |
| 446 | (map_terms_bottom_up f t1 `$ map_terms_bottom_up f t2) | |
| 31724 | 447 | | map_terms_bottom_up f ((v, ty) `|=> t) = f | 
| 448 | ((v, ty) `|=> map_terms_bottom_up f t) | |
| 27711 | 449 | | map_terms_bottom_up f (ICase (((t, ty), ps), t0)) = f | 
| 450 | (ICase (((map_terms_bottom_up f t, ty), (map o pairself) | |
| 451 | (map_terms_bottom_up f) ps), map_terms_bottom_up f t0)); | |
| 452 | ||
| 37448 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 453 | fun map_classparam_instances_as_term f = | 
| 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 454 | (map o apfst o apsnd) (fn const => case f (IConst const) of IConst const' => const') | 
| 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 455 | |
| 27711 | 456 | fun map_terms_stmt f NoStmt = NoStmt | 
| 37437 | 457 | | map_terms_stmt f (Fun (c, ((tysm, eqs), case_cong))) = Fun (c, ((tysm, (map o apfst) | 
| 458 | (fn (ts, t) => (map f ts, f t)) eqs), case_cong)) | |
| 27711 | 459 | | map_terms_stmt f (stmt as Datatype _) = stmt | 
| 460 | | map_terms_stmt f (stmt as Datatypecons _) = stmt | |
| 461 | | map_terms_stmt f (stmt as Class _) = stmt | |
| 462 | | map_terms_stmt f (stmt as Classrel _) = stmt | |
| 463 | | map_terms_stmt f (stmt as Classparam _) = stmt | |
| 37448 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 464 | | map_terms_stmt f (Classinst (arity, (super_instances, classparam_instances))) = | 
| 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 465 | Classinst (arity, (super_instances, (pairself o map_classparam_instances_as_term) f classparam_instances)); | 
| 27711 | 466 | |
| 27103 | 467 | fun is_cons program name = case Graph.get_node program name | 
| 24219 | 468 | of Datatypecons _ => true | 
| 469 | | _ => false; | |
| 470 | ||
| 37440 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 haftmann parents: 
37437diff
changeset | 471 | fun is_case (Fun (_, (_, SOME _))) = true | 
| 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 haftmann parents: 
37437diff
changeset | 472 | | is_case _ = false; | 
| 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 haftmann parents: 
37437diff
changeset | 473 | |
| 43048 
c62bed03fbce
improving heuristics of type annotations in contravariant positions for the special case with instances of the class partial_term_of
 bulwahn parents: 
42385diff
changeset | 474 | fun lookup_classparam_instance program name = program |> Graph.get_first | 
| 
c62bed03fbce
improving heuristics of type annotations in contravariant positions for the special case with instances of the class partial_term_of
 bulwahn parents: 
42385diff
changeset | 475 | (fn (_, (Classinst ((class, _), (_, (param_insts, _))), _)) => | 
| 
c62bed03fbce
improving heuristics of type annotations in contravariant positions for the special case with instances of the class partial_term_of
 bulwahn parents: 
42385diff
changeset | 476 | Option.map (fn ((const, _), _) => (class, const)) | 
| 
c62bed03fbce
improving heuristics of type annotations in contravariant positions for the special case with instances of the class partial_term_of
 bulwahn parents: 
42385diff
changeset | 477 | (find_first (fn ((_, (inst_const, _)), _) => inst_const = name) param_insts) | _ => NONE) | 
| 
c62bed03fbce
improving heuristics of type annotations in contravariant positions for the special case with instances of the class partial_term_of
 bulwahn parents: 
42385diff
changeset | 478 | |
| 42359 | 479 | fun labelled_name thy program name = | 
| 42361 | 480 | let val ctxt = Proof_Context.init_global thy in | 
| 42359 | 481 | case Graph.get_node program name of | 
| 482 | Fun (c, _) => quote (Code.string_of_const thy c) | |
| 42361 | 483 | | Datatype (tyco, _) => "type " ^ quote (Proof_Context.extern_type ctxt tyco) | 
| 42359 | 484 | | Datatypecons (c, _) => quote (Code.string_of_const thy c) | 
| 42361 | 485 | | Class (class, _) => "class " ^ quote (Proof_Context.extern_class ctxt class) | 
| 42359 | 486 | | Classrel (sub, super) => | 
| 487 | let | |
| 488 | val Class (sub, _) = Graph.get_node program sub; | |
| 489 | val Class (super, _) = Graph.get_node program super; | |
| 490 | in | |
| 42361 | 491 | quote (Proof_Context.extern_class ctxt sub ^ " < " ^ Proof_Context.extern_class ctxt super) | 
| 42359 | 492 | end | 
| 493 | | Classparam (c, _) => quote (Code.string_of_const thy c) | |
| 494 | | Classinst ((class, (tyco, _)), _) => | |
| 495 | let | |
| 496 | val Class (class, _) = Graph.get_node program class; | |
| 497 | val Datatype (tyco, _) = Graph.get_node program tyco; | |
| 498 | in | |
| 42361 | 499 | quote (Proof_Context.extern_type ctxt tyco ^ " :: " ^ Proof_Context.extern_class ctxt class) | 
| 42359 | 500 | end | 
| 501 | end; | |
| 32895 | 502 | |
| 37440 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 haftmann parents: 
37437diff
changeset | 503 | fun linear_stmts program = | 
| 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 haftmann parents: 
37437diff
changeset | 504 | rev (Graph.strong_conn program) | 
| 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 haftmann parents: 
37437diff
changeset | 505 | |> map (AList.make (Graph.get_node program)); | 
| 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 haftmann parents: 
37437diff
changeset | 506 | |
| 32895 | 507 | fun group_stmts thy program = | 
| 508 | let | |
| 509 | fun is_fun (_, Fun _) = true | is_fun _ = false; | |
| 510 | fun is_datatypecons (_, Datatypecons _) = true | is_datatypecons _ = false; | |
| 511 | fun is_datatype (_, Datatype _) = true | is_datatype _ = false; | |
| 512 | fun is_class (_, Class _) = true | is_class _ = false; | |
| 513 | fun is_classrel (_, Classrel _) = true | is_classrel _ = false; | |
| 514 | fun is_classparam (_, Classparam _) = true | is_classparam _ = false; | |
| 515 | fun is_classinst (_, Classinst _) = true | is_classinst _ = false; | |
| 516 | fun group stmts = | |
| 517 | if forall (is_datatypecons orf is_datatype) stmts | |
| 518 | then (filter is_datatype stmts, [], ([], [])) | |
| 519 | else if forall (is_class orf is_classrel orf is_classparam) stmts | |
| 520 | then ([], filter is_class stmts, ([], [])) | |
| 521 | else if forall (is_fun orf is_classinst) stmts | |
| 522 | then ([], [], List.partition is_fun stmts) | |
| 523 |       else error ("Illegal mutual dependencies: " ^
 | |
| 524 | (commas o map (labelled_name thy program o fst)) stmts) | |
| 525 | in | |
| 37440 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 haftmann parents: 
37437diff
changeset | 526 | linear_stmts program | 
| 32895 | 527 | |> map group | 
| 528 | end; | |
| 529 | ||
| 24219 | 530 | |
| 27103 | 531 | (** translation kernel **) | 
| 24219 | 532 | |
| 28724 | 533 | (* generic mechanisms *) | 
| 534 | ||
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 535 | fun ensure_stmt lookup declare generate thing (dep, (naming, program)) = | 
| 24219 | 536 | let | 
| 28706 | 537 | fun add_dep name = case dep of NONE => I | 
| 538 | | SOME dep => Graph.add_edge (dep, name); | |
| 539 | val (name, naming') = case lookup naming thing | |
| 540 | of SOME name => (name, naming) | |
| 541 | | NONE => declare thing naming; | |
| 542 | in case try (Graph.get_node program) name | |
| 543 | of SOME stmt => program | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 544 | |> add_dep name | 
| 28706 | 545 | |> pair naming' | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 546 | |> pair dep | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 547 | |> pair name | 
| 28706 | 548 | | NONE => program | 
| 549 | |> Graph.default_node (name, NoStmt) | |
| 550 | |> add_dep name | |
| 551 | |> pair naming' | |
| 552 | |> curry generate (SOME name) | |
| 553 | ||> snd | |
| 554 | |-> (fn stmt => (apsnd o Graph.map_node name) (K stmt)) | |
| 555 | |> pair dep | |
| 556 | |> pair name | |
| 24219 | 557 | end; | 
| 558 | ||
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 559 | exception PERMISSIVE of unit; | 
| 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 560 | |
| 37698 | 561 | 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: 
36210diff
changeset | 562 | if permissive | 
| 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 563 | then raise PERMISSIVE () | 
| 42385 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 564 | else | 
| 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 565 | let | 
| 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 566 | val err_thm = | 
| 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 567 | (case some_thm of | 
| 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 568 | SOME thm => "\n(in code equation " ^ Display.string_of_thm_global thy thm ^ ")" | 
| 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 569 | | NONE => ""); | 
| 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 570 | in error (msg ^ err_thm ^ ":\n" ^ sub_msg) end; | 
| 37698 | 571 | |
| 572 | fun not_wellsorted thy permissive some_thm ty sort e = | |
| 573 | let | |
| 47005 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 wenzelm parents: 
46961diff
changeset | 574 | val err_class = Sorts.class_error (Context.pretty_global thy) e; | 
| 42385 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 575 | val err_typ = | 
| 47005 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 wenzelm parents: 
46961diff
changeset | 576 | "Type " ^ Syntax.string_of_typ_global thy ty ^ " not of sort " ^ | 
| 42385 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 577 | Syntax.string_of_sort_global thy sort; | 
| 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 578 | in | 
| 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 579 | translation_error thy permissive some_thm "Wellsortedness error" | 
| 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 580 | (err_typ ^ "\n" ^ err_class) | 
| 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 581 | end; | 
| 26972 | 582 | |
| 44790 
c13fdf710a40
adding type inference for disambiguation annotations in code equation
 bulwahn parents: 
44789diff
changeset | 583 | (* inference of type annotations for disambiguation with type classes *) | 
| 
c13fdf710a40
adding type inference for disambiguation annotations in code equation
 bulwahn parents: 
44789diff
changeset | 584 | |
| 45000 
0febe2089425
adding abstraction layer; more precise function names
 bulwahn parents: 
44999diff
changeset | 585 | fun mk_tagged_type (true, T) = Type ("", [T])
 | 
| 
0febe2089425
adding abstraction layer; more precise function names
 bulwahn parents: 
44999diff
changeset | 586 | | mk_tagged_type (false, T) = T | 
| 
0febe2089425
adding abstraction layer; more precise function names
 bulwahn parents: 
44999diff
changeset | 587 | |
| 44998 
f12ef61ea76e
determining the fastype of a case-pattern but ignoring dummy type constructors that were added as markers for type annotations
 bulwahn parents: 
44997diff
changeset | 588 | fun dest_tagged_type (Type ("", [T])) = (true, T)
 | 
| 
f12ef61ea76e
determining the fastype of a case-pattern but ignoring dummy type constructors that were added as markers for type annotations
 bulwahn parents: 
44997diff
changeset | 589 | | dest_tagged_type T = (false, T) | 
| 
f12ef61ea76e
determining the fastype of a case-pattern but ignoring dummy type constructors that were added as markers for type annotations
 bulwahn parents: 
44997diff
changeset | 590 | |
| 
f12ef61ea76e
determining the fastype of a case-pattern but ignoring dummy type constructors that were added as markers for type annotations
 bulwahn parents: 
44997diff
changeset | 591 | val untag_term = map_types (snd o dest_tagged_type) | 
| 
f12ef61ea76e
determining the fastype of a case-pattern but ignoring dummy type constructors that were added as markers for type annotations
 bulwahn parents: 
44997diff
changeset | 592 | |
| 45000 
0febe2089425
adding abstraction layer; more precise function names
 bulwahn parents: 
44999diff
changeset | 593 | fun tag_term (proj_sort, _) eqngr = | 
| 44997 | 594 | let | 
| 595 | val has_sort_constraints = exists (not o null) o map proj_sort o Code_Preproc.sortargs eqngr | |
| 45000 
0febe2089425
adding abstraction layer; more precise function names
 bulwahn parents: 
44999diff
changeset | 596 | fun tag (Const (c', T')) (Const (c, T)) = | 
| 
0febe2089425
adding abstraction layer; more precise function names
 bulwahn parents: 
44999diff
changeset | 597 | Const (c, | 
| 
0febe2089425
adding abstraction layer; more precise function names
 bulwahn parents: 
44999diff
changeset | 598 | mk_tagged_type (not (null (Term.add_tvarsT T' [])) andalso has_sort_constraints c, T)) | 
| 
0febe2089425
adding abstraction layer; more precise function names
 bulwahn parents: 
44999diff
changeset | 599 | | tag (t1 $ u1) (t $ u) = tag t1 t $ tag u1 u | 
| 
0febe2089425
adding abstraction layer; more precise function names
 bulwahn parents: 
44999diff
changeset | 600 | | tag (Abs (_, _, t1)) (Abs (x, T, t)) = Abs (x, T, tag t1 t) | 
| 
0febe2089425
adding abstraction layer; more precise function names
 bulwahn parents: 
44999diff
changeset | 601 | | tag (Free _) (t as Free _) = t | 
| 
0febe2089425
adding abstraction layer; more precise function names
 bulwahn parents: 
44999diff
changeset | 602 | | tag (Var _) (t as Var _) = t | 
| 
0febe2089425
adding abstraction layer; more precise function names
 bulwahn parents: 
44999diff
changeset | 603 | | tag (Bound _) (t as Bound _) = t | 
| 44997 | 604 | in | 
| 45000 
0febe2089425
adding abstraction layer; more precise function names
 bulwahn parents: 
44999diff
changeset | 605 | tag | 
| 44997 | 606 | end | 
| 44790 
c13fdf710a40
adding type inference for disambiguation annotations in code equation
 bulwahn parents: 
44789diff
changeset | 607 | |
| 44997 | 608 | fun annotate thy algbr eqngr (c, ty) args rhs = | 
| 44790 
c13fdf710a40
adding type inference for disambiguation annotations in code equation
 bulwahn parents: 
44789diff
changeset | 609 | let | 
| 45128 
5af3a3203a76
discontinued obsolete alias structure ProofContext;
 wenzelm parents: 
45009diff
changeset | 610 | val ctxt = Proof_Context.init_global thy |> Config.put Type_Infer_Context.const_sorts false | 
| 44790 
c13fdf710a40
adding type inference for disambiguation annotations in code equation
 bulwahn parents: 
44789diff
changeset | 611 | val erase = map_types (fn _ => Type_Infer.anyT []) | 
| 
c13fdf710a40
adding type inference for disambiguation annotations in code equation
 bulwahn parents: 
44789diff
changeset | 612 | val reinfer = singleton (Type_Infer_Context.infer_types ctxt) | 
| 44996 
410eea28b0f7
also adding type annotations for the dynamic invocation
 bulwahn parents: 
44855diff
changeset | 613 | val lhs = list_comb (Const (c, ty), map (map_types Type.strip_sorts o fst) args) | 
| 
410eea28b0f7
also adding type annotations for the dynamic invocation
 bulwahn parents: 
44855diff
changeset | 614 | val reinferred_rhs = snd (Logic.dest_equals (reinfer (Logic.mk_equals (lhs, erase rhs)))) | 
| 44790 
c13fdf710a40
adding type inference for disambiguation annotations in code equation
 bulwahn parents: 
44789diff
changeset | 615 | in | 
| 45000 
0febe2089425
adding abstraction layer; more precise function names
 bulwahn parents: 
44999diff
changeset | 616 | tag_term algbr eqngr reinferred_rhs rhs | 
| 44996 
410eea28b0f7
also adding type annotations for the dynamic invocation
 bulwahn parents: 
44855diff
changeset | 617 | end | 
| 
410eea28b0f7
also adding type annotations for the dynamic invocation
 bulwahn parents: 
44855diff
changeset | 618 | |
| 44997 | 619 | fun annotate_eqns thy algbr eqngr (c, ty) eqns = | 
| 620 | map (apfst (fn (args, (rhs, some_abs)) => (args, | |
| 621 | (annotate thy algbr eqngr (c, ty) args rhs, some_abs)))) eqns | |
| 28724 | 622 | |
| 623 | (* translation *) | |
| 624 | ||
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 625 | fun ensure_tyco thy algbr eqngr permissive tyco = | 
| 30932 | 626 | let | 
| 40726 
16dcfedc4eb7
keep type variable arguments of datatype constructors in bookkeeping
 haftmann parents: 
40711diff
changeset | 627 | val ((vs, cos), _) = Code.get_type thy tyco; | 
| 30932 | 628 | val stmt_datatype = | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 629 | fold_map (translate_tyvar_sort thy algbr eqngr permissive) vs | 
| 40726 
16dcfedc4eb7
keep type variable arguments of datatype constructors in bookkeeping
 haftmann parents: 
40711diff
changeset | 630 | ##>> fold_map (fn (c, (vs, tys)) => | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 631 | ensure_const thy algbr eqngr permissive c | 
| 40726 
16dcfedc4eb7
keep type variable arguments of datatype constructors in bookkeeping
 haftmann parents: 
40711diff
changeset | 632 | ##>> pair (map (unprefix "'" o fst) vs) | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 633 | ##>> 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: 
36210diff
changeset | 634 | #>> (fn info => Datatype (tyco, info)); | 
| 30932 | 635 | 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: 
36210diff
changeset | 636 | and ensure_const thy algbr eqngr permissive c = | 
| 30932 | 637 | let | 
| 638 | fun stmt_datatypecons tyco = | |
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 639 | ensure_tyco thy algbr eqngr permissive tyco | 
| 30932 | 640 | #>> (fn tyco => Datatypecons (c, tyco)); | 
| 641 | fun stmt_classparam class = | |
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 642 | ensure_class thy algbr eqngr permissive class | 
| 30932 | 643 | #>> (fn class => Classparam (c, class)); | 
| 34891 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34251diff
changeset | 644 | fun stmt_fun cert = | 
| 32872 
019201eb7e07
variables in type schemes must be renamed simultaneously with variables in equations
 haftmann parents: 
32795diff
changeset | 645 | let | 
| 35226 | 646 | val ((vs, ty), eqns) = Code.equations_of_cert thy cert; | 
| 44997 | 647 | val eqns' = annotate_eqns thy algbr eqngr (c, ty) eqns | 
| 37440 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 haftmann parents: 
37437diff
changeset | 648 | 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: 
32795diff
changeset | 649 | in | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 650 | 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: 
36210diff
changeset | 651 | ##>> translate_typ thy algbr eqngr permissive ty | 
| 44791 | 652 | ##>> translate_eqns thy algbr eqngr permissive eqns' | 
| 37440 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 haftmann parents: 
37437diff
changeset | 653 | #>> (fn info => Fun (c, (info, some_case_cong))) | 
| 32872 
019201eb7e07
variables in type schemes must be renamed simultaneously with variables in equations
 haftmann parents: 
32795diff
changeset | 654 | end; | 
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 655 | val stmt_const = case Code.get_type_of_constr_or_abstr thy c | 
| 35226 | 656 | of SOME (tyco, _) => stmt_datatypecons tyco | 
| 30932 | 657 | | NONE => (case AxClass.class_of_param thy c | 
| 658 | of SOME class => stmt_classparam class | |
| 34891 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34251diff
changeset | 659 | | NONE => stmt_fun (Code_Preproc.cert eqngr c)) | 
| 30932 | 660 | 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: 
36210diff
changeset | 661 | and ensure_class thy (algbr as (_, algebra)) eqngr permissive class = | 
| 24918 | 662 | let | 
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 663 | val super_classes = (Sorts.minimize_sort algebra o Sorts.super_classes algebra) class; | 
| 24969 | 664 | val cs = #params (AxClass.get_info thy class); | 
| 24918 | 665 | val stmt_class = | 
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 666 | fold_map (fn super_class => ensure_class thy algbr eqngr permissive super_class | 
| 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 667 | ##>> 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: 
36210diff
changeset | 668 | ##>> 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: 
36210diff
changeset | 669 | ##>> translate_typ thy algbr eqngr permissive ty) cs | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 670 | #>> (fn info => Class (class, (unprefix "'" Name.aT, info))) | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 671 | in ensure_stmt lookup_class (declare_class thy) stmt_class class end | 
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 672 | and ensure_classrel thy algbr eqngr permissive (sub_class, super_class) = | 
| 24918 | 673 | let | 
| 674 | val stmt_classrel = | |
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 675 | ensure_class thy algbr eqngr permissive sub_class | 
| 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 676 | ##>> ensure_class thy algbr eqngr permissive super_class | 
| 24918 | 677 | #>> Classrel; | 
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 678 | 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: 
36210diff
changeset | 679 | and ensure_inst thy (algbr as (_, algebra)) eqngr permissive (class, tyco) = | 
| 24918 | 680 | let | 
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 681 | val super_classes = (Sorts.minimize_sort algebra o Sorts.super_classes algebra) class; | 
| 37448 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 682 | val these_classparams = these o try (#params o AxClass.get_info thy); | 
| 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 683 | val classparams = these_classparams class; | 
| 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 684 | val further_classparams = maps these_classparams | 
| 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 685 | ((Sorts.complete_sort algebra o Sorts.super_classes algebra) class); | 
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43326diff
changeset | 686 | val vs = Name.invent_names Name.context "'a" (Sorts.mg_domain algebra tyco [class]); | 
| 24918 | 687 | val sorts' = Sorts.mg_domain (Sign.classes_of thy) tyco [class]; | 
| 688 | val vs' = map2 (fn (v, sort1) => fn sort2 => (v, | |
| 689 | Sorts.inter_sort (Sign.classes_of thy) (sort1, sort2))) vs sorts'; | |
| 690 | val arity_typ = Type (tyco, map TFree vs); | |
| 691 | 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: 
37216diff
changeset | 692 | fun translate_super_instance super_class = | 
| 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 693 | ensure_class thy algbr eqngr permissive super_class | 
| 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 694 | ##>> ensure_classrel thy algbr eqngr permissive (class, super_class) | 
| 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 695 | ##>> translate_dicts thy algbr eqngr permissive NONE (arity_typ, [super_class]) | 
| 41118 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 696 | #>> (fn ((super_class, classrel), [Dict ([], Dict_Const (inst, dss))]) => | 
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 697 | (super_class, (classrel, (inst, dss)))); | 
| 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 698 | fun translate_classparam_instance (c, ty) = | 
| 24918 | 699 | let | 
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 700 | val raw_const = Const (c, map_type_tfree (K arity_typ') ty); | 
| 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 701 | val thm = AxClass.unoverload_conv thy (Thm.cterm_of thy raw_const); | 
| 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 702 | val const = (apsnd Logic.unvarifyT_global o dest_Const o snd | 
| 24918 | 703 | o Logic.dest_equals o Thm.prop_of) thm; | 
| 704 | in | |
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 705 | ensure_const thy algbr eqngr permissive c | 
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 706 | ##>> translate_const thy algbr eqngr permissive (SOME thm) (const, NONE) | 
| 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 707 | #>> (fn (c, IConst const') => ((c, const'), (thm, true))) | 
| 24918 | 708 | end; | 
| 709 | val stmt_inst = | |
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 710 | ensure_class thy algbr eqngr permissive class | 
| 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 711 | ##>> ensure_tyco thy algbr eqngr permissive tyco | 
| 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 712 | ##>> fold_map (translate_tyvar_sort thy algbr eqngr permissive) vs | 
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 713 | ##>> fold_map translate_super_instance super_classes | 
| 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 714 | ##>> fold_map translate_classparam_instance classparams | 
| 37448 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 715 | ##>> fold_map translate_classparam_instance further_classparams | 
| 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 716 | #>> (fn (((((class, tyco), arity_args), super_instances), | 
| 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 717 | classparam_instances), further_classparam_instances) => | 
| 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 718 | Classinst ((class, (tyco, arity_args)), (super_instances, | 
| 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 719 | (classparam_instances, further_classparam_instances)))); | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 720 | 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: 
36210diff
changeset | 721 | and translate_typ thy algbr eqngr permissive (TFree (v, _)) = | 
| 30932 | 722 | pair (ITyVar (unprefix "'" v)) | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 723 | | translate_typ thy algbr eqngr permissive (Type (tyco, tys)) = | 
| 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 724 | ensure_tyco thy algbr eqngr permissive tyco | 
| 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 725 | ##>> fold_map (translate_typ thy algbr eqngr permissive) tys | 
| 30932 | 726 | #>> (fn (tyco, tys) => tyco `%% tys) | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 727 | 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: 
36210diff
changeset | 728 | 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: 
36210diff
changeset | 729 | | translate_term thy algbr eqngr permissive some_thm (Free (v, _), some_abs) = | 
| 31889 | 730 | pair (IVar (SOME v)) | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 731 | | translate_term thy algbr eqngr permissive some_thm (Abs (v, ty, t), some_abs) = | 
| 24918 | 732 | let | 
| 42284 | 733 | val (v', t') = Syntax_Trans.variant_abs (Name.desymbolize false v, ty, t); | 
| 32273 | 734 | val v'' = if member (op =) (Term.add_free_names t' []) v' | 
| 735 | then SOME v' else NONE | |
| 24918 | 736 | in | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 737 | translate_typ thy algbr eqngr permissive ty | 
| 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 738 | ##>> translate_term thy algbr eqngr permissive some_thm (t', some_abs) | 
| 32273 | 739 | #>> (fn (ty, t) => (v'', ty) `|=> t) | 
| 24918 | 740 | end | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 741 | | translate_term thy algbr eqngr permissive some_thm (t as _ $ _, some_abs) = | 
| 24918 | 742 | case strip_comb t | 
| 743 | of (Const (c, ty), ts) => | |
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 744 | translate_app thy algbr eqngr permissive some_thm (((c, ty), ts), some_abs) | 
| 24918 | 745 | | (t', ts) => | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 746 | 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: 
36210diff
changeset | 747 | ##>> fold_map (translate_term thy algbr eqngr permissive some_thm o rpair NONE) ts | 
| 24918 | 748 | #>> (fn (t, ts) => t `$$ ts) | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 749 | 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: 
36210diff
changeset | 750 | 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: 
36210diff
changeset | 751 | ##>> translate_term thy algbr eqngr permissive some_thm (rhs, some_abs) | 
| 35226 | 752 | #>> rpair (some_thm, proper) | 
| 37440 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 haftmann parents: 
37437diff
changeset | 753 | and translate_eqns thy algbr eqngr permissive eqns prgrm = | 
| 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 haftmann parents: 
37437diff
changeset | 754 | prgrm |> fold_map (translate_eqn thy algbr eqngr permissive) eqns | 
| 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 haftmann parents: 
37437diff
changeset | 755 | handle PERMISSIVE () => ([], prgrm) | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 756 | and translate_const thy algbr eqngr permissive some_thm ((c, ty), some_abs) = | 
| 30932 | 757 | let | 
| 37698 | 758 | val _ = if (case some_abs of NONE => true | SOME abs => not (c = abs)) | 
| 35226 | 759 | andalso Code.is_abstr thy c | 
| 37698 | 760 | then translation_error thy permissive some_thm | 
| 761 |           "Abstraction violation" ("constant " ^ Code.string_of_const thy c)
 | |
| 762 | else () | |
| 45000 
0febe2089425
adding abstraction layer; more precise function names
 bulwahn parents: 
44999diff
changeset | 763 | val (annotate, ty') = dest_tagged_type ty | 
| 44792 
26b19918e670
adding minimalistic implementation for printing the type annotations
 bulwahn parents: 
44791diff
changeset | 764 | val arg_typs = Sign.const_typargs thy (c, ty'); | 
| 32873 | 765 | val sorts = Code_Preproc.sortargs eqngr c; | 
| 44792 
26b19918e670
adding minimalistic implementation for printing the type annotations
 bulwahn parents: 
44791diff
changeset | 766 | val (function_typs, body_typ) = Term.strip_type ty'; | 
| 26972 | 767 | in | 
| 37698 | 768 | ensure_const thy algbr eqngr permissive c | 
| 37448 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 769 | ##>> fold_map (translate_typ thy algbr eqngr permissive) arg_typs | 
| 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 770 | ##>> fold_map (translate_dicts thy algbr eqngr permissive some_thm) (arg_typs ~~ sorts) | 
| 44789 
5a062c23c7db
adding the body type as well to the code generation for constants as it is required for type annotations of constants
 bulwahn parents: 
44788diff
changeset | 771 | ##>> fold_map (translate_typ thy algbr eqngr permissive) (body_typ :: function_typs) | 
| 
5a062c23c7db
adding the body type as well to the code generation for constants as it is required for type annotations of constants
 bulwahn parents: 
44788diff
changeset | 772 | #>> (fn (((c, arg_typs), dss), body_typ :: function_typs) => | 
| 44792 
26b19918e670
adding minimalistic implementation for printing the type annotations
 bulwahn parents: 
44791diff
changeset | 773 | IConst (c, (((arg_typs, dss), (function_typs, body_typ)), annotate))) | 
| 26972 | 774 | end | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 775 | 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: 
36210diff
changeset | 776 | 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: 
36210diff
changeset | 777 | ##>> fold_map (translate_term thy algbr eqngr permissive some_thm o rpair NONE) ts | 
| 24918 | 778 | #>> (fn (t, ts) => t `$$ ts) | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 779 | and translate_case thy algbr eqngr permissive some_thm (num_args, (t_pos, case_pats)) (c_ty, ts) = | 
| 24918 | 780 | let | 
| 40844 | 781 | fun arg_types num_args ty = fst (chop num_args (binder_types ty)); | 
| 31892 | 782 | 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: 
29277diff
changeset | 783 | val ty = nth tys t_pos; | 
| 31957 | 784 | fun mk_constr c t = let val n = Code.args_number thy c | 
| 44998 
f12ef61ea76e
determining the fastype of a case-pattern but ignoring dummy type constructors that were added as markers for type annotations
 bulwahn parents: 
44997diff
changeset | 785 | in ((c, arg_types n (fastype_of (untag_term t)) ---> ty), n) end; | 
| 31892 | 786 | val constrs = if null case_pats then [] | 
| 787 | else map2 mk_constr case_pats (nth_drop t_pos ts); | |
| 788 | fun casify naming constrs ty ts = | |
| 24918 | 789 | let | 
| 31935 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 790 | val undefineds = map_filter (lookup_const naming) (Code.undefineds thy); | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 791 | fun collapse_clause vs_map ts body = | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 792 | let | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 793 | in case body | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 794 | of IConst (c, _) => if member (op =) undefineds c | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 795 | then [] | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 796 | else [(ts, body)] | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 797 | | ICase (((IVar (SOME v), _), subclauses), _) => | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 798 | if forall (fn (pat', body') => exists_var pat' v | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 799 | orelse not (exists_var body' v)) subclauses | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 800 | then case AList.lookup (op =) vs_map v | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 801 | of SOME i => maps (fn (pat', body') => | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 802 | collapse_clause (AList.delete (op =) v vs_map) | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 803 | (nth_map i (K pat') ts) body') subclauses | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 804 | | NONE => [(ts, body)] | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 805 | else [(ts, body)] | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 806 | | _ => [(ts, body)] | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 807 | end; | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 808 | fun mk_clause mk tys t = | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 809 | let | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 810 | val (vs, body) = unfold_abs_eta tys t; | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 811 | 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: 
31892diff
changeset | 812 | val ts = map (IVar o fst) vs; | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 813 | in map mk (collapse_clause vs_map ts body) end; | 
| 31892 | 814 | val t = nth ts t_pos; | 
| 815 | val ts_clause = nth_drop t_pos ts; | |
| 31935 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 816 | val clauses = if null case_pats | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 817 | then mk_clause (fn ([t], body) => (t, body)) [ty] (the_single ts_clause) | 
| 44789 
5a062c23c7db
adding the body type as well to the code generation for constants as it is required for type annotations of constants
 bulwahn parents: 
44788diff
changeset | 818 | else maps (fn ((constr as IConst (_, ((_, (tys, _)), _)), n), t) => | 
| 33957 | 819 | mk_clause (fn (ts, body) => (constr `$$ ts, body)) (take n tys) t) | 
| 31935 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 820 | (constrs ~~ ts_clause); | 
| 31892 | 821 | in ((t, ty), clauses) end; | 
| 24918 | 822 | in | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 823 | 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: 
36210diff
changeset | 824 | ##>> 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: 
36210diff
changeset | 825 | #>> rpair n) constrs | 
| 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 826 | ##>> translate_typ thy algbr eqngr permissive ty | 
| 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 827 | ##>> fold_map (translate_term thy algbr eqngr permissive some_thm o rpair NONE) ts | 
| 31892 | 828 | #-> (fn (((t, constrs), ty), ts) => | 
| 829 | `(fn (_, (naming, _)) => ICase (casify naming constrs ty ts, t `$$ ts))) | |
| 24918 | 830 | end | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 831 | and translate_app_case thy algbr eqngr permissive some_thm (case_scheme as (num_args, _)) ((c, ty), ts) = | 
| 29973 | 832 | if length ts < num_args then | 
| 833 | let | |
| 834 | val k = length ts; | |
| 33957 | 835 | val tys = (take (num_args - k) o drop k o fst o strip_type) ty; | 
| 29973 | 836 | val ctxt = (fold o fold_aterms) Term.declare_term_frees ts Name.context; | 
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43326diff
changeset | 837 | val vs = Name.invent_names ctxt "a" tys; | 
| 29973 | 838 | in | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 839 | fold_map (translate_typ thy algbr eqngr permissive) tys | 
| 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 840 | ##>> translate_case thy algbr eqngr permissive some_thm case_scheme ((c, ty), ts @ map Free vs) | 
| 31888 | 841 | #>> (fn (tys, t) => map2 (fn (v, _) => pair (SOME v)) vs tys `|==> t) | 
| 29973 | 842 | end | 
| 843 | else if length ts > num_args then | |
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 844 | 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: 
36210diff
changeset | 845 | ##>> fold_map (translate_term thy algbr eqngr permissive some_thm o rpair NONE) (drop num_args ts) | 
| 29973 | 846 | #>> (fn (t, ts) => t `$$ ts) | 
| 847 | else | |
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 848 | 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: 
36210diff
changeset | 849 | and translate_app thy algbr eqngr permissive some_thm (c_ty_ts as ((c, _), _), some_abs) = | 
| 29973 | 850 | case Code.get_case_scheme thy c | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 851 | 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: 
36210diff
changeset | 852 | | 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: 
36210diff
changeset | 853 | 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: 
36210diff
changeset | 854 | fold_map (ensure_class thy algbr eqngr permissive) (proj_sort sort) | 
| 30932 | 855 | #>> (fn sort => (unprefix "'" v, sort)) | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 856 | and translate_dicts thy (algbr as (proj_sort, algebra)) eqngr permissive some_thm (ty, sort) = | 
| 30932 | 857 | let | 
| 41100 
6c0940392fb4
dictionary constants must permit explicit weakening of classes;
 haftmann parents: 
40844diff
changeset | 858 | datatype typarg_witness = | 
| 41118 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 859 | Weakening of (class * class) list * plain_typarg_witness | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 860 | and plain_typarg_witness = | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 861 | Global of (class * string) * typarg_witness list list | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 862 | | Local of string * (int * sort); | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 863 | fun class_relation ((Weakening (classrels, x)), sub_class) super_class = | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 864 | Weakening ((sub_class, super_class) :: classrels, x); | 
| 41100 
6c0940392fb4
dictionary constants must permit explicit weakening of classes;
 haftmann parents: 
40844diff
changeset | 865 | fun type_constructor (tyco, _) dss class = | 
| 41118 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 866 | Weakening ([], Global ((class, tyco), (map o map) fst dss)); | 
| 30932 | 867 | fun type_variable (TFree (v, sort)) = | 
| 868 | let | |
| 869 | val sort' = proj_sort sort; | |
| 41118 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 870 | in map_index (fn (n, class) => (Weakening ([], Local (v, (n, sort'))), class)) sort' end; | 
| 41100 
6c0940392fb4
dictionary constants must permit explicit weakening of classes;
 haftmann parents: 
40844diff
changeset | 871 | val typarg_witnesses = Sorts.of_sort_derivation algebra | 
| 36102 
a51d1d154c71
of_sort_derivation: pass-through full type information -- following the version by krauss/schropp;
 wenzelm parents: 
35961diff
changeset | 872 |       {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: 
35845diff
changeset | 873 | type_constructor = type_constructor, | 
| 37698 | 874 | type_variable = type_variable} (ty, proj_sort sort) | 
| 875 | handle Sorts.CLASS_ERROR e => not_wellsorted thy permissive some_thm ty sort e; | |
| 41118 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 876 | fun mk_dict (Weakening (classrels, x)) = | 
| 41100 
6c0940392fb4
dictionary constants must permit explicit weakening of classes;
 haftmann parents: 
40844diff
changeset | 877 | fold_map (ensure_classrel thy algbr eqngr permissive) classrels | 
| 41118 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 878 | ##>> mk_plain_dict x | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 879 | #>> Dict | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 880 | and mk_plain_dict (Global (inst, dss)) = | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 881 | ensure_inst thy algbr eqngr permissive inst | 
| 41100 
6c0940392fb4
dictionary constants must permit explicit weakening of classes;
 haftmann parents: 
40844diff
changeset | 882 | ##>> (fold_map o fold_map) mk_dict dss | 
| 41118 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 883 | #>> (fn (inst, dss) => Dict_Const (inst, dss)) | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 884 | | mk_plain_dict (Local (v, (n, sort))) = | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 885 | pair (Dict_Var (unprefix "'" v, (n, length sort))) | 
| 41100 
6c0940392fb4
dictionary constants must permit explicit weakening of classes;
 haftmann parents: 
40844diff
changeset | 886 | in fold_map mk_dict typarg_witnesses end; | 
| 24918 | 887 | |
| 25969 | 888 | |
| 27103 | 889 | (* store *) | 
| 890 | ||
| 34173 
458ced35abb8
reduced code generator cache to the baremost minimum
 haftmann parents: 
34084diff
changeset | 891 | structure Program = Code_Data | 
| 27103 | 892 | ( | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 893 | type T = naming * program; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 894 | val empty = (empty_naming, Graph.empty); | 
| 27103 | 895 | ); | 
| 896 | ||
| 39397 | 897 | fun invoke_generation ignore_cache thy (algebra, eqngr) f name = | 
| 898 | Program.change_yield (if ignore_cache then NONE else SOME thy) | |
| 899 | (fn naming_program => (NONE, naming_program) | |
| 900 | |> f thy algebra eqngr name | |
| 901 | |-> (fn name => fn (_, naming_program) => (name, naming_program))); | |
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 902 | |
| 27103 | 903 | |
| 904 | (* program generation *) | |
| 905 | ||
| 39487 
80b2cf3b0779
proper closures for static evaluation; no need for FIXMEs any longer
 haftmann parents: 
39475diff
changeset | 906 | fun consts_program thy permissive consts = | 
| 27103 | 907 | let | 
| 39487 
80b2cf3b0779
proper closures for static evaluation; no need for FIXMEs any longer
 haftmann parents: 
39475diff
changeset | 908 | fun project_consts consts (naming, program) = | 
| 
80b2cf3b0779
proper closures for static evaluation; no need for FIXMEs any longer
 haftmann parents: 
39475diff
changeset | 909 | if permissive then (consts, (naming, program)) | 
| 46614 
165886a4fe64
clarified Graph.restrict (formerly Graph.subgraph) based on public graph operations;
 wenzelm parents: 
45987diff
changeset | 910 | else (consts, (naming, Graph.restrict | 
| 39487 
80b2cf3b0779
proper closures for static evaluation; no need for FIXMEs any longer
 haftmann parents: 
39475diff
changeset | 911 | (member (op =) (Graph.all_succs program consts)) program)); | 
| 32873 | 912 | fun generate_consts thy algebra eqngr = | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 913 | fold_map (ensure_const thy algebra eqngr permissive); | 
| 27103 | 914 | in | 
| 39487 
80b2cf3b0779
proper closures for static evaluation; no need for FIXMEs any longer
 haftmann parents: 
39475diff
changeset | 915 | invoke_generation permissive thy (Code_Preproc.obtain false thy consts []) | 
| 
80b2cf3b0779
proper closures for static evaluation; no need for FIXMEs any longer
 haftmann parents: 
39475diff
changeset | 916 | generate_consts consts | 
| 27103 | 917 | |-> project_consts | 
| 918 | end; | |
| 919 | ||
| 920 | ||
| 921 | (* value evaluation *) | |
| 25969 | 922 | |
| 32873 | 923 | fun ensure_value thy algbr eqngr t = | 
| 24918 | 924 | let | 
| 925 | val ty = fastype_of t; | |
| 926 | val vs = fold_term_types (K (fold_atyps (insert (eq_fst op =) | |
| 927 | o dest_TFree))) t []; | |
| 45987 | 928 | val t' = annotate thy algbr eqngr (Term.dummy_patternN, ty) [] t; | 
| 24918 | 929 | val stmt_value = | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 930 | 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: 
36210diff
changeset | 931 | ##>> translate_typ thy algbr eqngr false ty | 
| 44996 
410eea28b0f7
also adding type annotations for the dynamic invocation
 bulwahn parents: 
44855diff
changeset | 932 | ##>> translate_term thy algbr eqngr false NONE (t', NONE) | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 933 | #>> (fn ((vs, ty), t) => Fun | 
| 37437 | 934 | (Term.dummy_patternN, (((vs, ty), [(([], t), (NONE, true))]), NONE))); | 
| 31063 
88aaab83b6fc
dropped explicit suppport for frees in evaluation conversion stack
 haftmann parents: 
31054diff
changeset | 935 | fun term_value (dep, (naming, program1)) = | 
| 25969 | 936 | let | 
| 37437 | 937 | val Fun (_, ((vs_ty, [(([], t), _)]), _)) = | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 938 | Graph.get_node program1 Term.dummy_patternN; | 
| 44338 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 wenzelm parents: 
43329diff
changeset | 939 | val deps = Graph.immediate_succs program1 Term.dummy_patternN; | 
| 46665 
919dfcdf6d8a
discontinued slightly odd Graph.del_nodes (inefficient due to full Table.map);
 wenzelm parents: 
46614diff
changeset | 940 | val program2 = Graph.del_node Term.dummy_patternN program1; | 
| 27103 | 941 | val deps_all = Graph.all_succs program2 deps; | 
| 46614 
165886a4fe64
clarified Graph.restrict (formerly Graph.subgraph) based on public graph operations;
 wenzelm parents: 
45987diff
changeset | 942 | val program3 = Graph.restrict (member (op =) deps_all) program2; | 
| 31063 
88aaab83b6fc
dropped explicit suppport for frees in evaluation conversion stack
 haftmann parents: 
31054diff
changeset | 943 | in (((naming, program3), ((vs_ty, t), deps)), (dep, (naming, program2))) end; | 
| 26011 | 944 | in | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 945 | ensure_stmt ((K o K) NONE) pair stmt_value Term.dummy_patternN | 
| 26011 | 946 | #> snd | 
| 31063 
88aaab83b6fc
dropped explicit suppport for frees in evaluation conversion stack
 haftmann parents: 
31054diff
changeset | 947 | #> term_value | 
| 26011 | 948 | end; | 
| 24219 | 949 | |
| 39487 
80b2cf3b0779
proper closures for static evaluation; no need for FIXMEs any longer
 haftmann parents: 
39475diff
changeset | 950 | fun original_sorts vs = | 
| 
80b2cf3b0779
proper closures for static evaluation; no need for FIXMEs any longer
 haftmann parents: 
39475diff
changeset | 951 | map (fn (v, _) => (v, (the o AList.lookup (op =) vs o prefix "'") v)); | 
| 
80b2cf3b0779
proper closures for static evaluation; no need for FIXMEs any longer
 haftmann parents: 
39475diff
changeset | 952 | |
| 
80b2cf3b0779
proper closures for static evaluation; no need for FIXMEs any longer
 haftmann parents: 
39475diff
changeset | 953 | fun dynamic_evaluator thy evaluator algebra eqngr vs t = | 
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 954 | let | 
| 31063 
88aaab83b6fc
dropped explicit suppport for frees in evaluation conversion stack
 haftmann parents: 
31054diff
changeset | 955 | val (((naming, program), (((vs', ty'), t'), deps)), _) = | 
| 39397 | 956 | invoke_generation false thy (algebra, eqngr) ensure_value t; | 
| 39487 
80b2cf3b0779
proper closures for static evaluation; no need for FIXMEs any longer
 haftmann parents: 
39475diff
changeset | 957 | in evaluator naming program ((original_sorts vs vs', (vs', ty')), t') deps end; | 
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 958 | |
| 41184 | 959 | fun dynamic_conv thy evaluator = | 
| 960 | Code_Preproc.dynamic_conv thy (dynamic_evaluator thy evaluator); | |
| 39475 | 961 | |
| 41184 | 962 | fun dynamic_value thy postproc evaluator = | 
| 963 | Code_Preproc.dynamic_value thy postproc (dynamic_evaluator thy evaluator); | |
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 964 | |
| 41365 | 965 | fun lift_evaluation thy evaluation' algebra eqngr naming program vs t = | 
| 39487 
80b2cf3b0779
proper closures for static evaluation; no need for FIXMEs any longer
 haftmann parents: 
39475diff
changeset | 966 | let | 
| 
80b2cf3b0779
proper closures for static evaluation; no need for FIXMEs any longer
 haftmann parents: 
39475diff
changeset | 967 | val (((_, program'), (((vs', ty'), t'), deps)), _) = | 
| 
80b2cf3b0779
proper closures for static evaluation; no need for FIXMEs any longer
 haftmann parents: 
39475diff
changeset | 968 | ensure_value thy algebra eqngr t (NONE, (naming, program)); | 
| 41365 | 969 | in evaluation' ((original_sorts vs vs', (vs', ty')), t') deps end; | 
| 970 | ||
| 971 | fun lift_evaluator thy evaluator' consts algebra eqngr = | |
| 972 | let | |
| 973 | fun generate_consts thy algebra eqngr = | |
| 974 | fold_map (ensure_const thy algebra eqngr false); | |
| 975 | val (consts', (naming, program)) = | |
| 976 | invoke_generation true thy (algebra, eqngr) generate_consts consts; | |
| 977 | val evaluation' = evaluator' naming program consts'; | |
| 978 | in lift_evaluation thy evaluation' algebra eqngr naming program end; | |
| 979 | ||
| 980 | fun lift_evaluator_simple thy evaluator' consts algebra eqngr = | |
| 981 | let | |
| 982 | fun generate_consts thy algebra eqngr = | |
| 983 | fold_map (ensure_const thy algebra eqngr false); | |
| 984 | val (consts', (naming, program)) = | |
| 985 | invoke_generation true thy (algebra, eqngr) generate_consts consts; | |
| 986 | in evaluator' program end; | |
| 39487 
80b2cf3b0779
proper closures for static evaluation; no need for FIXMEs any longer
 haftmann parents: 
39475diff
changeset | 987 | |
| 41184 | 988 | fun static_conv thy consts conv = | 
| 41365 | 989 | Code_Preproc.static_conv thy consts (lift_evaluator thy conv consts); | 
| 38672 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 haftmann parents: 
38669diff
changeset | 990 | |
| 41184 | 991 | fun static_conv_simple thy consts conv = | 
| 41365 | 992 | Code_Preproc.static_conv thy consts (lift_evaluator_simple thy conv consts); | 
| 38672 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 haftmann parents: 
38669diff
changeset | 993 | |
| 41184 | 994 | fun static_value thy postproc consts evaluator = | 
| 41365 | 995 | Code_Preproc.static_value thy postproc consts (lift_evaluator thy evaluator consts); | 
| 39487 
80b2cf3b0779
proper closures for static evaluation; no need for FIXMEs any longer
 haftmann parents: 
39475diff
changeset | 996 | |
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 997 | |
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 998 | (** diagnostic commands **) | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 999 | |
| 31036 | 1000 | fun read_const_exprs thy = | 
| 1001 | let | |
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 1002 | fun consts_of thy' = Symtab.fold (fn (c, (_, NONE)) => cons c | _ => I) | 
| 47005 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 wenzelm parents: 
46961diff
changeset | 1003 | ((snd o #constants o Consts.dest o Sign.consts_of) thy') []; | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 1004 | fun belongs_here thy' c = forall | 
| 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 1005 | (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: 
36210diff
changeset | 1006 | fun consts_of_select thy' = filter (belongs_here thy') (consts_of thy'); | 
| 40711 
81bc73585eec
globbing constant expressions use more idiomatic underscore rather than star
 haftmann parents: 
39566diff
changeset | 1007 | fun read_const_expr "_" = ([], consts_of thy) | 
| 
81bc73585eec
globbing constant expressions use more idiomatic underscore rather than star
 haftmann parents: 
39566diff
changeset | 1008 | | read_const_expr s = if String.isSuffix "._" s | 
| 
81bc73585eec
globbing constant expressions use more idiomatic underscore rather than star
 haftmann parents: 
39566diff
changeset | 1009 | then ([], consts_of_select (Context.this_theory thy (unsuffix "._" s))) | 
| 31156 | 1010 | else ([Code.read_const thy s], []); | 
| 31036 | 1011 | in pairself flat o split_list o map read_const_expr end; | 
| 1012 | ||
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1013 | fun code_depgr thy consts = | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1014 | let | 
| 39397 | 1015 | val (_, eqngr) = Code_Preproc.obtain true thy consts []; | 
| 34173 
458ced35abb8
reduced code generator cache to the baremost minimum
 haftmann parents: 
34084diff
changeset | 1016 | val all_consts = Graph.all_succs eqngr consts; | 
| 46614 
165886a4fe64
clarified Graph.restrict (formerly Graph.subgraph) based on public graph operations;
 wenzelm parents: 
45987diff
changeset | 1017 | in Graph.restrict (member (op =) all_consts) eqngr end; | 
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1018 | |
| 31125 
80218ee73167
transferred code generator preprocessor into separate module
 haftmann parents: 
31088diff
changeset | 1019 | 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: 
30932diff
changeset | 1020 | |
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1021 | fun code_deps thy consts = | 
| 27103 | 1022 | let | 
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1023 | val eqngr = code_depgr thy consts; | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1024 | val constss = Graph.strong_conn eqngr; | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1025 | val mapping = Symtab.empty |> fold (fn consts => fold (fn const => | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1026 | Symtab.update (const, consts)) consts) constss; | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1027 | fun succs consts = consts | 
| 44338 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 wenzelm parents: 
43329diff
changeset | 1028 | |> maps (Graph.immediate_succs eqngr) | 
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1029 | |> subtract (op =) consts | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1030 | |> map (the o Symtab.lookup mapping) | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1031 | |> distinct (op =); | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1032 | val conn = [] |> fold (fn consts => cons (consts, succs consts)) constss; | 
| 31156 | 1033 | 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: 
30932diff
changeset | 1034 | |> commas; | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1035 | val prgr = map (fn (consts, constss) => | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1036 |       { name = namify consts, ID = namify consts, dir = "", unfold = true,
 | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1037 | path = "", parents = map namify constss }) conn; | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1038 | in Present.display_graph prgr end; | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1039 | |
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1040 | local | 
| 27103 | 1041 | |
| 31036 | 1042 | fun code_thms_cmd thy = code_thms thy o op @ o read_const_exprs thy; | 
| 1043 | 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: 
30932diff
changeset | 1044 | |
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1045 | in | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1046 | |
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1047 | val _ = | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46665diff
changeset | 1048 |   Outer_Syntax.improper_command @{command_spec "code_thms"}
 | 
| 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46665diff
changeset | 1049 | "print system of code equations for code" | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36272diff
changeset | 1050 | (Scan.repeat1 Parse.term_group | 
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1051 | >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1052 | 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: 
30932diff
changeset | 1053 | |
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1054 | val _ = | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46665diff
changeset | 1055 |   Outer_Syntax.improper_command @{command_spec "code_deps"}
 | 
| 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46665diff
changeset | 1056 | "visualize dependencies of code equations for code" | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36272diff
changeset | 1057 | (Scan.repeat1 Parse.term_group | 
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1058 | >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1059 | 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: 
30932diff
changeset | 1060 | |
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1061 | end; | 
| 27103 | 1062 | |
| 24219 | 1063 | end; (*struct*) | 
| 1064 | ||
| 1065 | ||
| 28054 | 1066 | structure Basic_Code_Thingol: BASIC_CODE_THINGOL = Code_Thingol; |