| author | bulwahn | 
| Tue, 09 Jun 2009 13:40:57 +0200 | |
| changeset 31515 | 62fc203eed88 | 
| parent 31156 | 90fed3d4430f | 
| child 31724 | 9b5a128cdb5c | 
| permissions | -rw-r--r-- | 
| 24219 | 1 | (* Title: Tools/code/code_thingol.ML | 
| 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 `$$; | |
| 11 | infixr 3 `|->; | |
| 12 | infixr 3 `|-->; | |
| 13 | ||
| 14 | signature BASIC_CODE_THINGOL = | |
| 15 | sig | |
| 16 | type vname = string; | |
| 17 | datatype dict = | |
| 18 | DictConst of string * dict list list | |
| 19 | | DictVar of string list * (vname * (int * int)); | |
| 20 | datatype itype = | |
| 21 | `%% of string * itype list | |
| 22 | | ITyVar of vname; | |
| 31049 | 23 | type const = string * ((itype list * dict list list) * itype list (*types of arguments*)) | 
| 24219 | 24 | datatype iterm = | 
| 24591 | 25 | IConst of const | 
| 24219 | 26 | | IVar of vname | 
| 27 | | `$ of iterm * iterm | |
| 28 | | `|-> of (vname * itype) * iterm | |
| 29 | | ICase of ((iterm * itype) * (iterm * iterm) list) * iterm; | |
| 30 | (*((term, type), [(selector pattern, body term )]), primitive term)*) | |
| 31 | val `$$ : iterm * iterm list -> iterm; | |
| 32 | val `|--> : (vname * itype) list * iterm -> iterm; | |
| 33 | type typscheme = (vname * sort) list * itype; | |
| 34 | end; | |
| 35 | ||
| 36 | signature CODE_THINGOL = | |
| 37 | sig | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 38 | include BASIC_CODE_THINGOL | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 39 |   val unfoldl: ('a -> ('a * 'b) option) -> 'a -> 'a * 'b list
 | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 40 |   val unfoldr: ('a -> ('b * 'a) option) -> 'a -> 'b list * 'a
 | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 41 | val unfold_fun: itype -> itype list * itype | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 42 | val unfold_app: iterm -> iterm * iterm list | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 43 | val split_abs: iterm -> (((vname * iterm option) * itype) * iterm) option | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 44 | val unfold_abs: iterm -> ((vname * iterm option) * itype) list * iterm | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 45 | val split_let: iterm -> (((iterm * itype) * iterm) * iterm) option | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 46 | val unfold_let: iterm -> ((iterm * itype) * iterm) list * iterm | 
| 31049 | 47 | val unfold_const_app: iterm -> (const * iterm list) option | 
| 24219 | 48 | val collapse_let: ((vname * itype) * iterm) * iterm | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 49 | -> (iterm * itype) * (iterm * iterm) list | 
| 31049 | 50 | val eta_expand: int -> const * iterm list -> iterm | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 51 | val contains_dictvar: iterm -> bool | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 52 | val locally_monomorphic: iterm -> bool | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 53 | val fold_constnames: (string -> 'a -> 'a) -> iterm -> 'a -> 'a | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 54 | val fold_varnames: (string -> 'a -> 'a) -> iterm -> 'a -> 'a | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 55 | val fold_unbound_varnames: (string -> 'a -> 'a) -> iterm -> 'a -> 'a | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 56 | |
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 57 | type naming | 
| 28706 | 58 | val empty_naming: naming | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 59 | val lookup_class: naming -> class -> string option | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 60 | val lookup_classrel: naming -> class * class -> string option | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 61 | val lookup_tyco: naming -> string -> string option | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 62 | val lookup_instance: naming -> class * string -> string option | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 63 | val lookup_const: naming -> string -> string option | 
| 31054 | 64 | val ensure_declared_const: theory -> string -> naming -> string * naming | 
| 24219 | 65 | |
| 24918 | 66 | datatype stmt = | 
| 27024 | 67 | NoStmt | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 68 | | Fun of string * (typscheme * ((iterm list * iterm) * (thm * bool)) list) | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 69 | | Datatype of string * ((vname * sort) list * (string * itype list) list) | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 70 | | Datatypecons of string * string | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 71 | | Class of class * (vname * ((class * string) list * (string * itype) list)) | 
| 24219 | 72 | | Classrel of class * class | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 73 | | Classparam of string * class | 
| 24219 | 74 | | Classinst of (class * (string * (vname * sort) list)) | 
| 75 | * ((class * (string * (string * dict list list))) list | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 76 | * ((string * const) * (thm * bool)) list) | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 77 | type program = stmt Graph.T | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 78 | val empty_funs: program -> string list | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 79 | val map_terms_bottom_up: (iterm -> iterm) -> iterm -> iterm | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 80 | val map_terms_stmt: (iterm -> iterm) -> stmt -> stmt | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 81 | val is_cons: program -> string -> bool | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 82 | val contr_classparam_typs: program -> string -> itype option list | 
| 24219 | 83 | |
| 31036 | 84 | val read_const_exprs: theory -> string list -> string list * string list | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 85 | val consts_program: theory -> string list -> string list * (naming * program) | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 86 | val cached_program: theory -> naming * program | 
| 30947 | 87 | val eval_conv: theory -> (sort -> sort) | 
| 31063 
88aaab83b6fc
dropped explicit suppport for frees in evaluation conversion stack
 haftmann parents: 
31054diff
changeset | 88 | -> (naming -> program -> ((string * sort) list * typscheme) * iterm -> string list -> cterm -> thm) | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 89 | -> cterm -> thm | 
| 30970 
3fe2e418a071
generic postprocessing scheme for term evaluations
 haftmann parents: 
30947diff
changeset | 90 | val eval: theory -> (sort -> sort) -> ((term -> term) -> 'a -> 'a) | 
| 31063 
88aaab83b6fc
dropped explicit suppport for frees in evaluation conversion stack
 haftmann parents: 
31054diff
changeset | 91 | -> (naming -> program -> ((string * sort) list * typscheme) * iterm -> string list -> 'a) | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 92 | -> term -> 'a | 
| 24219 | 93 | end; | 
| 94 | ||
| 28054 | 95 | structure Code_Thingol: CODE_THINGOL = | 
| 24219 | 96 | struct | 
| 97 | ||
| 98 | (** auxiliary **) | |
| 99 | ||
| 100 | fun unfoldl dest x = | |
| 101 | case dest x | |
| 102 | of NONE => (x, []) | |
| 103 | | SOME (x1, x2) => | |
| 104 | let val (x', xs') = unfoldl dest x1 in (x', xs' @ [x2]) end; | |
| 105 | ||
| 106 | fun unfoldr dest x = | |
| 107 | case dest x | |
| 108 | of NONE => ([], x) | |
| 109 | | SOME (x1, x2) => | |
| 110 | let val (xs', x') = unfoldr dest x2 in (x1::xs', x') end; | |
| 111 | ||
| 112 | ||
| 29962 
bd4dc7fa742d
tuned comments, stripped ID, deleted superfluous code
 haftmann parents: 
29952diff
changeset | 113 | (** language core - types, terms **) | 
| 24219 | 114 | |
| 115 | type vname = string; | |
| 116 | ||
| 117 | datatype dict = | |
| 118 | DictConst of string * dict list list | |
| 119 | | DictVar of string list * (vname * (int * int)); | |
| 120 | ||
| 121 | datatype itype = | |
| 122 | `%% of string * itype list | |
| 123 | | ITyVar of vname; | |
| 124 | ||
| 31049 | 125 | type const = string * ((itype list * dict list list) * itype list (*types of arguments*)) | 
| 24591 | 126 | |
| 24219 | 127 | datatype iterm = | 
| 24591 | 128 | IConst of const | 
| 24219 | 129 | | IVar of vname | 
| 130 | | `$ of iterm * iterm | |
| 131 | | `|-> of (vname * itype) * iterm | |
| 132 | | ICase of ((iterm * itype) * (iterm * iterm) list) * iterm; | |
| 133 | (*see also signature*) | |
| 134 | ||
| 135 | val op `$$ = Library.foldl (op `$); | |
| 136 | val op `|--> = Library.foldr (op `|->); | |
| 137 | ||
| 138 | val unfold_app = unfoldl | |
| 139 | (fn op `$ t => SOME t | |
| 140 | | _ => NONE); | |
| 141 | ||
| 142 | val split_abs = | |
| 143 | (fn (v, ty) `|-> (t as ICase (((IVar w, _), [(p, t')]), _)) => | |
| 144 | if v = w then SOME (((v, SOME p), ty), t') else SOME (((v, NONE), ty), t) | |
| 145 | | (v, ty) `|-> t => SOME (((v, NONE), ty), t) | |
| 146 | | _ => NONE); | |
| 147 | ||
| 148 | val unfold_abs = unfoldr split_abs; | |
| 149 | ||
| 150 | val split_let = | |
| 151 | (fn ICase (((td, ty), [(p, t)]), _) => SOME (((p, ty), td), t) | |
| 152 | | _ => NONE); | |
| 153 | ||
| 154 | val unfold_let = unfoldr split_let; | |
| 155 | ||
| 156 | fun unfold_const_app t = | |
| 157 | case unfold_app t | |
| 158 | of (IConst c, ts) => SOME (c, ts) | |
| 159 | | _ => NONE; | |
| 160 | ||
| 161 | fun fold_aiterms f (t as IConst _) = f t | |
| 162 | | fold_aiterms f (t as IVar _) = f t | |
| 163 | | fold_aiterms f (t1 `$ t2) = fold_aiterms f t1 #> fold_aiterms f t2 | |
| 164 | | fold_aiterms f (t as _ `|-> t') = f t #> fold_aiterms f t' | |
| 165 | | fold_aiterms f (ICase (_, t)) = fold_aiterms f t; | |
| 166 | ||
| 167 | fun fold_constnames f = | |
| 168 | let | |
| 169 | fun add (IConst (c, _)) = f c | |
| 170 | | add _ = I; | |
| 171 | in fold_aiterms add end; | |
| 172 | ||
| 173 | fun fold_varnames f = | |
| 174 | let | |
| 175 | fun add (IVar v) = f v | |
| 176 | | add ((v, _) `|-> _) = f v | |
| 177 | | add _ = I; | |
| 178 | in fold_aiterms add end; | |
| 179 | ||
| 180 | fun fold_unbound_varnames f = | |
| 181 | let | |
| 182 | fun add _ (IConst _) = I | |
| 183 | | add vs (IVar v) = if not (member (op =) vs v) then f v else I | |
| 184 | | add vs (t1 `$ t2) = add vs t1 #> add vs t2 | |
| 185 | | add vs ((v, _) `|-> t) = add (insert (op =) v vs) t | |
| 186 | | add vs (ICase (_, t)) = add vs t; | |
| 187 | in add [] end; | |
| 188 | ||
| 189 | fun collapse_let (((v, ty), se), be as ICase (((IVar w, _), ds), _)) = | |
| 190 | let | |
| 191 | fun exists_v t = fold_unbound_varnames (fn w => fn b => | |
| 192 | b orelse v = w) t false; | |
| 193 | in if v = w andalso forall (fn (t1, t2) => | |
| 194 | exists_v t1 orelse not (exists_v t2)) ds | |
| 195 | then ((se, ty), ds) | |
| 196 | else ((se, ty), [(IVar v, be)]) | |
| 197 | end | |
| 198 | | collapse_let (((v, ty), se), be) = | |
| 199 | ((se, ty), [(IVar v, be)]) | |
| 200 | ||
| 27711 | 201 | fun eta_expand k (c as (_, (_, tys)), ts) = | 
| 24219 | 202 | let | 
| 203 | val j = length ts; | |
| 204 | val l = k - j; | |
| 205 | val ctxt = (fold o fold_varnames) Name.declare ts Name.context; | |
| 206 | val vs_tys = Name.names ctxt "a" ((curry Library.take l o curry Library.drop j) tys); | |
| 207 | in vs_tys `|--> IConst c `$$ ts @ map (fn (v, _) => IVar v) vs_tys end; | |
| 208 | ||
| 24662 
f79f6061525c
more precise treatment of free dictionary parameters for evaluation
 haftmann parents: 
24591diff
changeset | 209 | fun contains_dictvar t = | 
| 
f79f6061525c
more precise treatment of free dictionary parameters for evaluation
 haftmann parents: 
24591diff
changeset | 210 | let | 
| 
f79f6061525c
more precise treatment of free dictionary parameters for evaluation
 haftmann parents: 
24591diff
changeset | 211 | fun contains (DictConst (_, dss)) = (fold o fold) contains dss | 
| 
f79f6061525c
more precise treatment of free dictionary parameters for evaluation
 haftmann parents: 
24591diff
changeset | 212 | | contains (DictVar _) = K true; | 
| 
f79f6061525c
more precise treatment of free dictionary parameters for evaluation
 haftmann parents: 
24591diff
changeset | 213 | in | 
| 
f79f6061525c
more precise treatment of free dictionary parameters for evaluation
 haftmann parents: 
24591diff
changeset | 214 | fold_aiterms | 
| 31049 | 215 | (fn IConst (_, ((_, dss), _)) => (fold o fold) contains dss | _ => I) t false | 
| 24662 
f79f6061525c
more precise treatment of free dictionary parameters for evaluation
 haftmann parents: 
24591diff
changeset | 216 | end; | 
| 
f79f6061525c
more precise treatment of free dictionary parameters for evaluation
 haftmann parents: 
24591diff
changeset | 217 | |
| 25621 | 218 | fun locally_monomorphic (IConst _) = false | 
| 219 | | locally_monomorphic (IVar _) = true | |
| 220 | | locally_monomorphic (t `$ _) = locally_monomorphic t | |
| 221 | | locally_monomorphic (_ `|-> t) = locally_monomorphic t | |
| 222 | | locally_monomorphic (ICase ((_, ds), _)) = exists (locally_monomorphic o snd) ds; | |
| 223 | ||
| 224 | ||
| 28688 | 225 | (** namings **) | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 226 | |
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 227 | (* policies *) | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 228 | |
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 229 | local | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 230 | fun thyname_of thy f x = the (AList.lookup (op =) (f x) Markup.theory_nameN); | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 231 | fun thyname_of_class thy = | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 232 | thyname_of thy (ProofContext.query_class (ProofContext.init thy)); | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 233 | fun thyname_of_tyco thy = | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 234 | thyname_of thy (Type.the_tags (Sign.tsig_of thy)); | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 235 | fun thyname_of_instance thy inst = case AxClass.arity_property thy inst Markup.theory_nameN | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 236 |    of [] => error ("no such instance: " ^ quote (snd inst ^ " :: " ^ fst inst))
 | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 237 | | thyname :: _ => thyname; | 
| 28706 | 238 | fun thyname_of_const thy c = case AxClass.class_of_param thy c | 
| 239 | of SOME class => thyname_of_class thy class | |
| 240 | | NONE => (case Code.get_datatype_of_constr thy c | |
| 241 | of SOME dtco => thyname_of_tyco thy dtco | |
| 242 | | NONE => thyname_of thy (Consts.the_tags (Sign.consts_of thy)) c); | |
| 31036 | 243 | fun purify_base "op &" = "and" | 
| 244 | | purify_base "op |" = "or" | |
| 245 | | purify_base "op -->" = "implies" | |
| 246 | | purify_base "op :" = "member" | |
| 247 | | purify_base "op =" = "eq" | |
| 248 | | purify_base "*" = "product" | |
| 249 | | purify_base "+" = "sum" | |
| 250 | | purify_base s = Name.desymbolize false s; | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 251 | fun namify thy get_basename get_thyname name = | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 252 | let | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 253 | val prefix = get_thyname thy name; | 
| 31036 | 254 | 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 | 255 | in Long_Name.append prefix base end; | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 256 | in | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 257 | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 258 | fun namify_class thy = namify thy Long_Name.base_name thyname_of_class; | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 259 | fun namify_classrel thy = namify thy (fn (class1, class2) => | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 260 | Long_Name.base_name class2 ^ "_" ^ Long_Name.base_name class1) (fn thy => thyname_of_class thy o fst); | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 261 | (*order fits nicely with composed projections*) | 
| 28688 | 262 | fun namify_tyco thy "fun" = "Pure.fun" | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 263 | | namify_tyco thy tyco = namify thy Long_Name.base_name thyname_of_tyco tyco; | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 264 | 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 | 265 | 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 | 266 | 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 | 267 | |
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 268 | end; (* local *) | 
| 
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 | |
| 28688 | 271 | (* data *) | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 272 | |
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 273 | datatype naming = Naming of {
 | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 274 | class: class Symtab.table * Name.context, | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 275 | classrel: string Symreltab.table * Name.context, | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 276 | tyco: string Symtab.table * Name.context, | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 277 | instance: string Symreltab.table * Name.context, | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 278 | const: string Symtab.table * Name.context | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 279 | } | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 280 | |
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 281 | fun dest_Naming (Naming naming) = naming; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 282 | |
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 283 | val empty_naming = Naming {
 | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 284 | class = (Symtab.empty, Name.context), | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 285 | classrel = (Symreltab.empty, Name.context), | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 286 | tyco = (Symtab.empty, Name.context), | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 287 | instance = (Symreltab.empty, Name.context), | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 288 | const = (Symtab.empty, Name.context) | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 289 | }; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 290 | |
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 291 | local | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 292 | fun mk_naming (class, classrel, tyco, instance, const) = | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 293 |     Naming { class = class, classrel = classrel,
 | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 294 | tyco = tyco, instance = instance, const = const }; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 295 |   fun map_naming f (Naming { class, classrel, tyco, instance, const }) =
 | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 296 | mk_naming (f (class, classrel, tyco, instance, const)); | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 297 | in | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 298 | fun map_class f = map_naming | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 299 | (fn (class, classrel, tyco, inst, const) => | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 300 | (f class, classrel, tyco, inst, const)); | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 301 | fun map_classrel f = map_naming | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 302 | (fn (class, classrel, tyco, inst, const) => | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 303 | (class, f classrel, tyco, inst, const)); | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 304 | fun map_tyco f = map_naming | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 305 | (fn (class, classrel, tyco, inst, const) => | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 306 | (class, classrel, f tyco, inst, const)); | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 307 | fun map_instance f = map_naming | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 308 | (fn (class, classrel, tyco, inst, const) => | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 309 | (class, classrel, tyco, f inst, const)); | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 310 | fun map_const f = map_naming | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 311 | (fn (class, classrel, tyco, inst, const) => | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 312 | (class, classrel, tyco, inst, f const)); | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 313 | end; (*local*) | 
| 
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 | fun add_variant update (thing, name) (tab, used) = | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 316 | let | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 317 | val (name', used') = yield_singleton Name.variants name used; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 318 | val tab' = update (thing, name') tab; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 319 | in (tab', used') end; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 320 | |
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 321 | fun declare thy mapp lookup update namify thing = | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 322 | mapp (add_variant update (thing, namify thy thing)) | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 323 | #> `(fn naming => the (lookup naming thing)); | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 324 | |
| 28688 | 325 | |
| 326 | (* lookup and declare *) | |
| 327 | ||
| 328 | local | |
| 329 | ||
| 330 | val suffix_class = "class"; | |
| 331 | val suffix_classrel = "classrel" | |
| 332 | val suffix_tyco = "tyco"; | |
| 333 | val suffix_instance = "inst"; | |
| 334 | val suffix_const = "const"; | |
| 335 | ||
| 336 | fun add_suffix nsp NONE = NONE | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 337 | | add_suffix nsp (SOME name) = SOME (Long_Name.append name nsp); | 
| 28688 | 338 | |
| 339 | in | |
| 340 | ||
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 341 | val lookup_class = add_suffix suffix_class | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 342 | oo Symtab.lookup o fst o #class o dest_Naming; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 343 | val lookup_classrel = add_suffix suffix_classrel | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 344 | oo Symreltab.lookup o fst o #classrel o dest_Naming; | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 345 | val lookup_tyco = add_suffix suffix_tyco | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 346 | oo Symtab.lookup o fst o #tyco o dest_Naming; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 347 | val lookup_instance = add_suffix suffix_instance | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 348 | oo Symreltab.lookup o fst o #instance o dest_Naming; | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 349 | val lookup_const = add_suffix suffix_const | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 350 | oo Symtab.lookup o fst o #const o dest_Naming; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 351 | |
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 352 | fun declare_class thy = declare thy map_class | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 353 | lookup_class Symtab.update_new namify_class; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 354 | fun declare_classrel thy = declare thy map_classrel | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 355 | lookup_classrel Symreltab.update_new namify_classrel; | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 356 | fun declare_tyco thy = declare thy map_tyco | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 357 | lookup_tyco Symtab.update_new namify_tyco; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 358 | fun declare_instance thy = declare thy map_instance | 
| 30648 
17365ef082f3
clarified relationship of modules Code_Name and Code_Printer
 haftmann parents: 
30364diff
changeset | 359 | lookup_instance Symreltab.update_new namify_instance; | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 360 | fun declare_const thy = declare thy map_const | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 361 | lookup_const Symtab.update_new namify_const; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 362 | |
| 31054 | 363 | fun ensure_declared_const thy const naming = | 
| 364 | case lookup_const naming const | |
| 365 | of SOME const' => (const', naming) | |
| 366 | | NONE => declare_const thy const naming; | |
| 367 | ||
| 28688 | 368 | val unfold_fun = unfoldr | 
| 369 | (fn "Pure.fun.tyco" `%% [ty1, ty2] => SOME (ty1, ty2) | |
| 370 | | _ => NONE); (*depends on suffix_tyco and namify_tyco!*) | |
| 371 | ||
| 372 | end; (* local *) | |
| 373 | ||
| 24219 | 374 | |
| 27103 | 375 | (** statements, abstract programs **) | 
| 24219 | 376 | |
| 377 | type typscheme = (vname * sort) list * itype; | |
| 24918 | 378 | datatype stmt = | 
| 27024 | 379 | NoStmt | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 380 | | Fun of string * (typscheme * ((iterm list * iterm) * (thm * bool)) list) | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 381 | | Datatype of string * ((vname * sort) list * (string * itype list) list) | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 382 | | Datatypecons of string * string | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 383 | | Class of class * (vname * ((class * string) list * (string * itype) list)) | 
| 24219 | 384 | | Classrel of class * class | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 385 | | Classparam of string * class | 
| 24219 | 386 | | Classinst of (class * (string * (vname * sort) list)) | 
| 387 | * ((class * (string * (string * dict list list))) list | |
| 28350 | 388 | * ((string * const) * (thm * bool)) list); | 
| 24219 | 389 | |
| 27103 | 390 | type program = stmt Graph.T; | 
| 24219 | 391 | |
| 27103 | 392 | fun empty_funs program = | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 393 | Graph.fold (fn (name, (Fun (c, (_, [])), _)) => cons c | 
| 27103 | 394 | | _ => I) program []; | 
| 24219 | 395 | |
| 27711 | 396 | fun map_terms_bottom_up f (t as IConst _) = f t | 
| 397 | | map_terms_bottom_up f (t as IVar _) = f t | |
| 398 | | map_terms_bottom_up f (t1 `$ t2) = f | |
| 399 | (map_terms_bottom_up f t1 `$ map_terms_bottom_up f t2) | |
| 400 | | map_terms_bottom_up f ((v, ty) `|-> t) = f | |
| 401 | ((v, ty) `|-> map_terms_bottom_up f t) | |
| 402 | | map_terms_bottom_up f (ICase (((t, ty), ps), t0)) = f | |
| 403 | (ICase (((map_terms_bottom_up f t, ty), (map o pairself) | |
| 404 | (map_terms_bottom_up f) ps), map_terms_bottom_up f t0)); | |
| 405 | ||
| 406 | fun map_terms_stmt f NoStmt = NoStmt | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 407 | | map_terms_stmt f (Fun (c, (tysm, eqs))) = Fun (c, (tysm, (map o apfst) | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 408 | (fn (ts, t) => (map f ts, f t)) eqs)) | 
| 27711 | 409 | | map_terms_stmt f (stmt as Datatype _) = stmt | 
| 410 | | map_terms_stmt f (stmt as Datatypecons _) = stmt | |
| 411 | | map_terms_stmt f (stmt as Class _) = stmt | |
| 412 | | map_terms_stmt f (stmt as Classrel _) = stmt | |
| 413 | | map_terms_stmt f (stmt as Classparam _) = stmt | |
| 414 | | map_terms_stmt f (Classinst (arity, (superarities, classparms))) = | |
| 415 | Classinst (arity, (superarities, (map o apfst o apsnd) (fn const => | |
| 416 | case f (IConst const) of IConst const' => const') classparms)); | |
| 417 | ||
| 27103 | 418 | fun is_cons program name = case Graph.get_node program name | 
| 24219 | 419 | of Datatypecons _ => true | 
| 420 | | _ => false; | |
| 421 | ||
| 27103 | 422 | fun contr_classparam_typs program name = case Graph.get_node program name | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 423 | of Classparam (_, class) => let | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 424 | val Class (_, (_, (_, params))) = Graph.get_node program class; | 
| 25621 | 425 | val SOME ty = AList.lookup (op =) params name; | 
| 426 | val (tys, res_ty) = unfold_fun ty; | |
| 427 | fun no_tyvar (_ `%% tys) = forall no_tyvar tys | |
| 428 | | no_tyvar (ITyVar _) = false; | |
| 429 | in if no_tyvar res_ty | |
| 430 | then map (fn ty => if no_tyvar ty then NONE else SOME ty) tys | |
| 431 | else [] | |
| 432 | end | |
| 433 | | _ => []; | |
| 434 | ||
| 24219 | 435 | |
| 27103 | 436 | (** translation kernel **) | 
| 24219 | 437 | |
| 28724 | 438 | (* generic mechanisms *) | 
| 439 | ||
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 440 | fun ensure_stmt lookup declare generate thing (dep, (naming, program)) = | 
| 24219 | 441 | let | 
| 28706 | 442 | fun add_dep name = case dep of NONE => I | 
| 443 | | SOME dep => Graph.add_edge (dep, name); | |
| 444 | val (name, naming') = case lookup naming thing | |
| 445 | of SOME name => (name, naming) | |
| 446 | | NONE => declare thing naming; | |
| 447 | in case try (Graph.get_node program) name | |
| 448 | of SOME stmt => program | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 449 | |> add_dep name | 
| 28706 | 450 | |> pair naming' | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 451 | |> pair dep | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 452 | |> pair name | 
| 28706 | 453 | | NONE => program | 
| 454 | |> Graph.default_node (name, NoStmt) | |
| 455 | |> add_dep name | |
| 456 | |> pair naming' | |
| 457 | |> curry generate (SOME name) | |
| 458 | ||> snd | |
| 459 | |-> (fn stmt => (apsnd o Graph.map_node name) (K stmt)) | |
| 460 | |> pair dep | |
| 461 | |> pair name | |
| 24219 | 462 | end; | 
| 463 | ||
| 26972 | 464 | fun not_wellsorted thy thm ty sort e = | 
| 465 | let | |
| 466 | val err_class = Sorts.class_error (Syntax.pp_global thy) e; | |
| 467 | val err_thm = case thm | |
| 30023 
55954f726043
permissive check for pattern discipline in case schemes
 haftmann parents: 
30009diff
changeset | 468 | of SOME thm => "\n(in code equation " ^ Display.string_of_thm thm ^ ")" | NONE => ""; | 
| 26972 | 469 | val err_typ = "Type " ^ Syntax.string_of_typ_global thy ty ^ " not of sort " | 
| 470 | ^ Syntax.string_of_sort_global thy sort; | |
| 471 |   in error ("Wellsortedness error" ^ err_thm ^ ":\n" ^ err_typ ^ "\n" ^ err_class) end;
 | |
| 472 | ||
| 28724 | 473 | |
| 474 | (* translation *) | |
| 475 | ||
| 30932 | 476 | fun ensure_tyco thy algbr funcgr tyco = | 
| 477 | let | |
| 478 | val stmt_datatype = | |
| 479 | let | |
| 480 | val (vs, cos) = Code.get_datatype thy tyco; | |
| 481 | in | |
| 482 | fold_map (translate_tyvar_sort thy algbr funcgr) vs | |
| 483 | ##>> fold_map (fn (c, tys) => | |
| 484 | ensure_const thy algbr funcgr c | |
| 485 | ##>> fold_map (translate_typ thy algbr funcgr) tys) cos | |
| 486 | #>> (fn info => Datatype (tyco, info)) | |
| 487 | end; | |
| 488 | in ensure_stmt lookup_tyco (declare_tyco thy) stmt_datatype tyco end | |
| 489 | and ensure_const thy algbr funcgr c = | |
| 490 | let | |
| 491 | fun stmt_datatypecons tyco = | |
| 492 | ensure_tyco thy algbr funcgr tyco | |
| 493 | #>> (fn tyco => Datatypecons (c, tyco)); | |
| 494 | fun stmt_classparam class = | |
| 495 | ensure_class thy algbr funcgr class | |
| 496 | #>> (fn class => Classparam (c, class)); | |
| 497 | fun stmt_fun ((vs, ty), raw_thms) = | |
| 498 | let | |
| 499 | val thms = if null (Term.add_tfreesT ty []) orelse (null o fst o strip_type) ty | |
| 500 | then raw_thms | |
| 31156 | 501 | else (map o apfst) (Code.expand_eta thy 1) raw_thms; | 
| 30932 | 502 | in | 
| 503 | fold_map (translate_tyvar_sort thy algbr funcgr) vs | |
| 504 | ##>> translate_typ thy algbr funcgr ty | |
| 505 | ##>> fold_map (translate_eq thy algbr funcgr) thms | |
| 506 | #>> (fn info => Fun (c, info)) | |
| 507 | end; | |
| 508 | val stmt_const = case Code.get_datatype_of_constr thy c | |
| 509 | of SOME tyco => stmt_datatypecons tyco | |
| 510 | | NONE => (case AxClass.class_of_param thy c | |
| 511 | of SOME class => stmt_classparam class | |
| 31125 
80218ee73167
transferred code generator preprocessor into separate module
 haftmann parents: 
31088diff
changeset | 512 | | NONE => stmt_fun (Code_Preproc.typ funcgr c, Code_Preproc.eqns funcgr c)) | 
| 30932 | 513 | in ensure_stmt lookup_const (declare_const thy) stmt_const c end | 
| 514 | and ensure_class thy (algbr as (_, algebra)) funcgr class = | |
| 24918 | 515 | let | 
| 28924 
5c8781b7d6a4
code_funcgr interface includes also sort algebra
 haftmann parents: 
28724diff
changeset | 516 | val superclasses = (Sorts.minimize_sort algebra o Sorts.super_classes algebra) class; | 
| 24969 | 517 | val cs = #params (AxClass.get_info thy class); | 
| 24918 | 518 | val stmt_class = | 
| 519 | fold_map (fn superclass => ensure_class thy algbr funcgr superclass | |
| 520 | ##>> ensure_classrel thy algbr funcgr (class, superclass)) superclasses | |
| 521 | ##>> fold_map (fn (c, ty) => ensure_const thy algbr funcgr c | |
| 28688 | 522 | ##>> translate_typ thy algbr funcgr ty) cs | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 523 | #>> (fn info => Class (class, (unprefix "'" Name.aT, info))) | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 524 | in ensure_stmt lookup_class (declare_class thy) stmt_class class end | 
| 24918 | 525 | and ensure_classrel thy algbr funcgr (subclass, superclass) = | 
| 526 | let | |
| 527 | val stmt_classrel = | |
| 528 | ensure_class thy algbr funcgr subclass | |
| 529 | ##>> ensure_class thy algbr funcgr superclass | |
| 530 | #>> Classrel; | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 531 | in ensure_stmt lookup_classrel (declare_classrel thy) stmt_classrel (subclass, superclass) end | 
| 26972 | 532 | and ensure_inst thy (algbr as (_, algebra)) funcgr (class, tyco) = | 
| 24918 | 533 | let | 
| 28924 
5c8781b7d6a4
code_funcgr interface includes also sort algebra
 haftmann parents: 
28724diff
changeset | 534 | val superclasses = (Sorts.minimize_sort algebra o Sorts.super_classes algebra) class; | 
| 24969 | 535 | val classparams = these (try (#params o AxClass.get_info thy) class); | 
| 24918 | 536 | val vs = Name.names Name.context "'a" (Sorts.mg_domain algebra tyco [class]); | 
| 537 | val sorts' = Sorts.mg_domain (Sign.classes_of thy) tyco [class]; | |
| 538 | val vs' = map2 (fn (v, sort1) => fn sort2 => (v, | |
| 539 | Sorts.inter_sort (Sign.classes_of thy) (sort1, sort2))) vs sorts'; | |
| 540 | val arity_typ = Type (tyco, map TFree vs); | |
| 541 | val arity_typ' = Type (tyco, map (fn (v, sort) => TVar ((v, 0), sort)) vs'); | |
| 28688 | 542 | fun translate_superarity superclass = | 
| 24918 | 543 | ensure_class thy algbr funcgr superclass | 
| 544 | ##>> ensure_classrel thy algbr funcgr (class, superclass) | |
| 28688 | 545 | ##>> translate_dicts thy algbr funcgr NONE (arity_typ, [superclass]) | 
| 24918 | 546 | #>> (fn ((superclass, classrel), [DictConst (inst, dss)]) => | 
| 547 | (superclass, (classrel, (inst, dss)))); | |
| 28688 | 548 | fun translate_classparam_inst (c, ty) = | 
| 24918 | 549 | let | 
| 550 | val c_inst = Const (c, map_type_tfree (K arity_typ') ty); | |
| 25597 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 haftmann parents: 
25485diff
changeset | 551 | val thm = AxClass.unoverload_conv thy (Thm.cterm_of thy c_inst); | 
| 24918 | 552 | val c_ty = (apsnd Logic.unvarifyT o dest_Const o snd | 
| 553 | o Logic.dest_equals o Thm.prop_of) thm; | |
| 554 | in | |
| 555 | ensure_const thy algbr funcgr c | |
| 28688 | 556 | ##>> translate_const thy algbr funcgr (SOME thm) c_ty | 
| 28350 | 557 | #>> (fn (c, IConst c_inst) => ((c, c_inst), (thm, true))) | 
| 24918 | 558 | end; | 
| 559 | val stmt_inst = | |
| 560 | ensure_class thy algbr funcgr class | |
| 561 | ##>> ensure_tyco thy algbr funcgr tyco | |
| 28688 | 562 | ##>> fold_map (translate_tyvar_sort thy algbr funcgr) vs | 
| 563 | ##>> fold_map translate_superarity superclasses | |
| 564 | ##>> fold_map translate_classparam_inst classparams | |
| 24918 | 565 | #>> (fn ((((class, tyco), arity), superarities), classparams) => | 
| 566 | Classinst ((class, (tyco, arity)), (superarities, classparams))); | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 567 | in ensure_stmt lookup_instance (declare_instance thy) stmt_inst (class, tyco) end | 
| 30932 | 568 | and translate_typ thy algbr funcgr (TFree (v, _)) = | 
| 569 | pair (ITyVar (unprefix "'" v)) | |
| 570 | | translate_typ thy algbr funcgr (Type (tyco, tys)) = | |
| 24918 | 571 | ensure_tyco thy algbr funcgr tyco | 
| 30932 | 572 | ##>> fold_map (translate_typ thy algbr funcgr) tys | 
| 573 | #>> (fn (tyco, tys) => tyco `%% tys) | |
| 28688 | 574 | and translate_term thy algbr funcgr thm (Const (c, ty)) = | 
| 575 | translate_app thy algbr funcgr thm ((c, ty), []) | |
| 576 | | translate_term thy algbr funcgr thm (Free (v, _)) = | |
| 24918 | 577 | pair (IVar v) | 
| 28688 | 578 | | translate_term thy algbr funcgr thm (Abs (abs as (_, ty, _))) = | 
| 24918 | 579 | let | 
| 580 | val (v, t) = Syntax.variant_abs abs; | |
| 581 | in | |
| 28688 | 582 | translate_typ thy algbr funcgr ty | 
| 583 | ##>> translate_term thy algbr funcgr thm t | |
| 24918 | 584 | #>> (fn (ty, t) => (v, ty) `|-> t) | 
| 585 | end | |
| 28688 | 586 | | translate_term thy algbr funcgr thm (t as _ $ _) = | 
| 24918 | 587 | case strip_comb t | 
| 588 | of (Const (c, ty), ts) => | |
| 28688 | 589 | translate_app thy algbr funcgr thm ((c, ty), ts) | 
| 24918 | 590 | | (t', ts) => | 
| 28688 | 591 | translate_term thy algbr funcgr thm t' | 
| 592 | ##>> fold_map (translate_term thy algbr funcgr thm) ts | |
| 24918 | 593 | #>> (fn (t, ts) => t `$$ ts) | 
| 31088 | 594 | and translate_eq thy algbr funcgr (thm, proper) = | 
| 30932 | 595 | let | 
| 596 | val (args, rhs) = (apfst (snd o strip_comb) o Logic.dest_equals | |
| 597 | o Logic.unvarify o prop_of) thm; | |
| 598 | in | |
| 599 | fold_map (translate_term thy algbr funcgr (SOME thm)) args | |
| 600 | ##>> translate_term thy algbr funcgr (SOME thm) rhs | |
| 31088 | 601 | #>> rpair (thm, proper) | 
| 30932 | 602 | end | 
| 28688 | 603 | and translate_const thy algbr funcgr thm (c, ty) = | 
| 26972 | 604 | let | 
| 605 | val tys = Sign.const_typargs thy (c, ty); | |
| 31125 
80218ee73167
transferred code generator preprocessor into separate module
 haftmann parents: 
31088diff
changeset | 606 | val sorts = (map snd o fst o Code_Preproc.typ funcgr) c; | 
| 26972 | 607 | val tys_args = (fst o Term.strip_type) ty; | 
| 608 | in | |
| 609 | ensure_const thy algbr funcgr c | |
| 31049 | 610 | ##>> fold_map (translate_typ thy algbr funcgr) tys | 
| 28688 | 611 | ##>> fold_map (translate_dicts thy algbr funcgr thm) (tys ~~ sorts) | 
| 612 | ##>> fold_map (translate_typ thy algbr funcgr) tys_args | |
| 31049 | 613 | #>> (fn (((c, tys), iss), tys_args) => IConst (c, ((tys, iss), tys_args))) | 
| 26972 | 614 | end | 
| 29973 | 615 | and translate_app_const thy algbr funcgr thm (c_ty, ts) = | 
| 28688 | 616 | translate_const thy algbr funcgr thm c_ty | 
| 617 | ##>> fold_map (translate_term thy algbr funcgr thm) ts | |
| 24918 | 618 | #>> (fn (t, ts) => t `$$ ts) | 
| 29952 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 619 | and translate_case thy algbr funcgr thm (num_args, (t_pos, case_pats)) (c_ty, ts) = | 
| 24918 | 620 | let | 
| 29952 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 621 | val (tys, _) = (chop num_args o fst o strip_type o snd) c_ty; | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 622 | val t = nth ts t_pos; | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 623 | val ty = nth tys t_pos; | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 624 | val ts_clause = nth_drop t_pos ts; | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 625 | fun mk_clause (co, num_co_args) t = | 
| 24918 | 626 | let | 
| 29952 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 627 | val (vs, body) = Term.strip_abs_eta num_co_args t; | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 628 | val not_undefined = case body | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 629 | of (Const (c, _)) => not (Code.is_undefined thy c) | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 630 | | _ => true; | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 631 | val pat = list_comb (Const (co, map snd vs ---> ty), map Free vs); | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 632 | in (not_undefined, (pat, body)) end; | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 633 | val clauses = if null case_pats then let val ([v_ty], body) = | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 634 | Term.strip_abs_eta 1 (the_single ts_clause) | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 635 | in [(true, (Free v_ty, body))] end | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 636 | else map (uncurry mk_clause) | 
| 31156 | 637 | (AList.make (Code.no_args thy) case_pats ~~ ts_clause); | 
| 29952 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 638 | fun retermify ty (_, (IVar x, body)) = | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 639 | (x, ty) `|-> body | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 640 | | retermify _ (_, (pat, body)) = | 
| 24918 | 641 | let | 
| 29952 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 642 | val (IConst (_, (_, tys)), ts) = unfold_app pat; | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 643 | val vs = map2 (fn IVar x => fn ty => (x, ty)) ts tys; | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 644 | in vs `|--> body end; | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 645 | fun mk_icase const t ty clauses = | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 646 | let | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 647 | val (ts1, ts2) = chop t_pos (map (retermify ty) clauses); | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 648 | in | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 649 | ICase (((t, ty), map_filter (fn (b, d) => if b then SOME d else NONE) clauses), | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 650 | const `$$ (ts1 @ t :: ts2)) | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 651 | end; | 
| 24918 | 652 | in | 
| 29952 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 653 | translate_const thy algbr funcgr thm c_ty | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 654 | ##>> translate_term thy algbr funcgr thm t | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 655 | ##>> translate_typ thy algbr funcgr ty | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 656 | ##>> fold_map (fn (b, (pat, body)) => translate_term thy algbr funcgr thm pat | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 657 | ##>> translate_term thy algbr funcgr thm body | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 658 | #>> pair b) clauses | 
| 
9aed85067721
unified variable names in case expressions; no exponential fork in translation of case expressions
 haftmann parents: 
29277diff
changeset | 659 | #>> (fn (((const, t), ty), ds) => mk_icase const t ty ds) | 
| 24918 | 660 | end | 
| 29973 | 661 | and translate_app_case thy algbr funcgr thm (case_scheme as (num_args, _)) ((c, ty), ts) = | 
| 662 | if length ts < num_args then | |
| 663 | let | |
| 664 | val k = length ts; | |
| 665 | val tys = (curry Library.take (num_args - k) o curry Library.drop k o fst o strip_type) ty; | |
| 666 | val ctxt = (fold o fold_aterms) Term.declare_term_frees ts Name.context; | |
| 667 | val vs = Name.names ctxt "a" tys; | |
| 668 | in | |
| 669 | fold_map (translate_typ thy algbr funcgr) tys | |
| 670 | ##>> translate_case thy algbr funcgr thm case_scheme ((c, ty), ts @ map Free vs) | |
| 671 | #>> (fn (tys, t) => map2 (fn (v, _) => pair v) vs tys `|--> t) | |
| 672 | end | |
| 673 | else if length ts > num_args then | |
| 674 | translate_case thy algbr funcgr thm case_scheme ((c, ty), Library.take (num_args, ts)) | |
| 675 | ##>> fold_map (translate_term thy algbr funcgr thm) (Library.drop (num_args, ts)) | |
| 676 | #>> (fn (t, ts) => t `$$ ts) | |
| 677 | else | |
| 678 | translate_case thy algbr funcgr thm case_scheme ((c, ty), ts) | |
| 679 | and translate_app thy algbr funcgr thm (c_ty_ts as ((c, _), _)) = | |
| 680 | case Code.get_case_scheme thy c | |
| 30023 
55954f726043
permissive check for pattern discipline in case schemes
 haftmann parents: 
30009diff
changeset | 681 | of SOME case_scheme => translate_app_case thy algbr funcgr thm case_scheme c_ty_ts | 
| 30932 | 682 | | NONE => translate_app_const thy algbr funcgr thm c_ty_ts | 
| 683 | and translate_tyvar_sort thy (algbr as (proj_sort, _)) funcgr (v, sort) = | |
| 684 | fold_map (ensure_class thy algbr funcgr) (proj_sort sort) | |
| 685 | #>> (fn sort => (unprefix "'" v, sort)) | |
| 686 | and translate_dicts thy (algbr as (proj_sort, algebra)) funcgr thm (ty, sort) = | |
| 687 | let | |
| 688 | val pp = Syntax.pp_global thy; | |
| 689 | datatype typarg = | |
| 690 | Global of (class * string) * typarg list list | |
| 691 | | Local of (class * class) list * (string * (int * sort)); | |
| 692 | fun class_relation (Global ((_, tyco), yss), _) class = | |
| 693 | Global ((class, tyco), yss) | |
| 694 | | class_relation (Local (classrels, v), subclass) superclass = | |
| 695 | Local ((subclass, superclass) :: classrels, v); | |
| 696 | fun type_constructor tyco yss class = | |
| 697 | Global ((class, tyco), (map o map) fst yss); | |
| 698 | fun type_variable (TFree (v, sort)) = | |
| 699 | let | |
| 700 | val sort' = proj_sort sort; | |
| 701 | in map_index (fn (n, class) => (Local ([], (v, (n, sort'))), class)) sort' end; | |
| 702 | val typargs = Sorts.of_sort_derivation pp algebra | |
| 703 |       {class_relation = class_relation, type_constructor = type_constructor,
 | |
| 704 | type_variable = type_variable} (ty, proj_sort sort) | |
| 705 | handle Sorts.CLASS_ERROR e => not_wellsorted thy thm ty sort e; | |
| 706 | fun mk_dict (Global (inst, yss)) = | |
| 707 | ensure_inst thy algbr funcgr inst | |
| 708 | ##>> (fold_map o fold_map) mk_dict yss | |
| 709 | #>> (fn (inst, dss) => DictConst (inst, dss)) | |
| 710 | | mk_dict (Local (classrels, (v, (k, sort)))) = | |
| 711 | fold_map (ensure_classrel thy algbr funcgr) classrels | |
| 712 | #>> (fn classrels => DictVar (classrels, (unprefix "'" v, (k, length sort)))) | |
| 713 | in fold_map mk_dict typargs end; | |
| 24918 | 714 | |
| 25969 | 715 | |
| 27103 | 716 | (* store *) | 
| 717 | ||
| 718 | structure Program = CodeDataFun | |
| 719 | ( | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 720 | type T = naming * program; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 721 | val empty = (empty_naming, Graph.empty); | 
| 28706 | 722 | fun purge thy cs (naming, program) = | 
| 27609 | 723 | let | 
| 28706 | 724 | val names_delete = cs | 
| 725 | |> map_filter (lookup_const naming) | |
| 726 | |> filter (can (Graph.get_node program)) | |
| 727 | |> Graph.all_preds program; | |
| 728 | val program' = Graph.del_nodes names_delete program; | |
| 729 | in (naming, program') end; | |
| 27103 | 730 | ); | 
| 731 | ||
| 732 | val cached_program = Program.get; | |
| 25969 | 733 | |
| 28924 
5c8781b7d6a4
code_funcgr interface includes also sort algebra
 haftmann parents: 
28724diff
changeset | 734 | fun invoke_generation thy (algebra, funcgr) f name = | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 735 | Program.change_yield thy (fn naming_program => (NONE, naming_program) | 
| 28924 
5c8781b7d6a4
code_funcgr interface includes also sort algebra
 haftmann parents: 
28724diff
changeset | 736 | |> f thy algebra funcgr name | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 737 | |-> (fn name => fn (_, naming_program) => (name, naming_program))); | 
| 27103 | 738 | |
| 739 | ||
| 740 | (* program generation *) | |
| 741 | ||
| 742 | fun consts_program thy cs = | |
| 743 | let | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 744 | fun project_consts cs (naming, program) = | 
| 27103 | 745 | let | 
| 746 | val cs_all = Graph.all_succs program cs; | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 747 | in (cs, (naming, Graph.subgraph (member (op =) cs_all) program)) end; | 
| 27103 | 748 | fun generate_consts thy algebra funcgr = | 
| 749 | fold_map (ensure_const thy algebra funcgr); | |
| 750 | in | |
| 31125 
80218ee73167
transferred code generator preprocessor into separate module
 haftmann parents: 
31088diff
changeset | 751 | invoke_generation thy (Code_Preproc.obtain thy cs []) generate_consts cs | 
| 27103 | 752 | |-> project_consts | 
| 753 | end; | |
| 754 | ||
| 755 | ||
| 756 | (* value evaluation *) | |
| 25969 | 757 | |
| 31063 
88aaab83b6fc
dropped explicit suppport for frees in evaluation conversion stack
 haftmann parents: 
31054diff
changeset | 758 | fun ensure_value thy algbr funcgr t = | 
| 24918 | 759 | let | 
| 760 | val ty = fastype_of t; | |
| 761 | val vs = fold_term_types (K (fold_atyps (insert (eq_fst op =) | |
| 762 | o dest_TFree))) t []; | |
| 763 | val stmt_value = | |
| 28688 | 764 | fold_map (translate_tyvar_sort thy algbr funcgr) vs | 
| 765 | ##>> translate_typ thy algbr funcgr ty | |
| 766 | ##>> translate_term thy algbr funcgr NONE t | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 767 | #>> (fn ((vs, ty), t) => Fun | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 768 | (Term.dummy_patternN, ((vs, ty), [(([], t), (Drule.dummy_thm, true))]))); | 
| 31063 
88aaab83b6fc
dropped explicit suppport for frees in evaluation conversion stack
 haftmann parents: 
31054diff
changeset | 769 | fun term_value (dep, (naming, program1)) = | 
| 25969 | 770 | let | 
| 30947 | 771 | val Fun (_, (vs_ty, [(([], t), _)])) = | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 772 | Graph.get_node program1 Term.dummy_patternN; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 773 | val deps = Graph.imm_succs program1 Term.dummy_patternN; | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 774 | val program2 = Graph.del_nodes [Term.dummy_patternN] program1; | 
| 27103 | 775 | val deps_all = Graph.all_succs program2 deps; | 
| 776 | val program3 = Graph.subgraph (member (op =) deps_all) program2; | |
| 31063 
88aaab83b6fc
dropped explicit suppport for frees in evaluation conversion stack
 haftmann parents: 
31054diff
changeset | 777 | in (((naming, program3), ((vs_ty, t), deps)), (dep, (naming, program2))) end; | 
| 26011 | 778 | in | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 779 | ensure_stmt ((K o K) NONE) pair stmt_value Term.dummy_patternN | 
| 26011 | 780 | #> snd | 
| 31063 
88aaab83b6fc
dropped explicit suppport for frees in evaluation conversion stack
 haftmann parents: 
31054diff
changeset | 781 | #> term_value | 
| 26011 | 782 | end; | 
| 24219 | 783 | |
| 30947 | 784 | fun base_evaluator thy evaluator algebra funcgr vs t = | 
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 785 | let | 
| 31063 
88aaab83b6fc
dropped explicit suppport for frees in evaluation conversion stack
 haftmann parents: 
31054diff
changeset | 786 | val (((naming, program), (((vs', ty'), t'), deps)), _) = | 
| 
88aaab83b6fc
dropped explicit suppport for frees in evaluation conversion stack
 haftmann parents: 
31054diff
changeset | 787 | invoke_generation thy (algebra, funcgr) ensure_value t; | 
| 30947 | 788 | val vs'' = map (fn (v, _) => (v, (the o AList.lookup (op =) vs o prefix "'") v)) vs'; | 
| 31063 
88aaab83b6fc
dropped explicit suppport for frees in evaluation conversion stack
 haftmann parents: 
31054diff
changeset | 789 | in evaluator naming program ((vs'', (vs', ty')), t') deps end; | 
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 790 | |
| 31125 
80218ee73167
transferred code generator preprocessor into separate module
 haftmann parents: 
31088diff
changeset | 791 | fun eval_conv thy prep_sort = Code_Preproc.eval_conv thy prep_sort o base_evaluator thy; | 
| 
80218ee73167
transferred code generator preprocessor into separate module
 haftmann parents: 
31088diff
changeset | 792 | fun eval thy prep_sort postproc = Code_Preproc.eval thy prep_sort postproc o base_evaluator thy; | 
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 793 | |
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 794 | |
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 795 | (** diagnostic commands **) | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 796 | |
| 31036 | 797 | fun read_const_exprs thy = | 
| 798 | let | |
| 799 | fun consts_of some_thyname = | |
| 800 | let | |
| 801 | val thy' = case some_thyname | |
| 802 | of SOME thyname => ThyInfo.the_theory thyname thy | |
| 803 | | NONE => thy; | |
| 804 | val cs = Symtab.fold (fn (c, (_, NONE)) => cons c | _ => I) | |
| 805 | ((snd o #constants o Consts.dest o #consts o Sign.rep_sg) thy') []; | |
| 806 | fun belongs_here c = | |
| 807 | not (exists (fn thy'' => Sign.declared_const thy'' c) (Theory.parents_of thy')) | |
| 808 | in case some_thyname | |
| 809 | of NONE => cs | |
| 810 | | SOME thyname => filter belongs_here cs | |
| 811 | end; | |
| 812 | fun read_const_expr "*" = ([], consts_of NONE) | |
| 813 | | read_const_expr s = if String.isSuffix ".*" s | |
| 814 | then ([], consts_of (SOME (unsuffix ".*" s))) | |
| 31156 | 815 | else ([Code.read_const thy s], []); | 
| 31036 | 816 | in pairself flat o split_list o map read_const_expr end; | 
| 817 | ||
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 818 | fun code_depgr thy consts = | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 819 | let | 
| 31125 
80218ee73167
transferred code generator preprocessor into separate module
 haftmann parents: 
31088diff
changeset | 820 | val (_, eqngr) = Code_Preproc.obtain thy consts []; | 
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 821 | val select = Graph.all_succs eqngr consts; | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 822 | in | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 823 | eqngr | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 824 | |> not (null consts) ? Graph.subgraph (member (op =) select) | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 825 | |> Graph.map_nodes ((apsnd o map o apfst) (AxClass.overload thy)) | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 826 | end; | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 827 | |
| 31125 
80218ee73167
transferred code generator preprocessor into separate module
 haftmann parents: 
31088diff
changeset | 828 | 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 | 829 | |
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 830 | fun code_deps thy consts = | 
| 27103 | 831 | let | 
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 832 | val eqngr = code_depgr thy consts; | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 833 | val constss = Graph.strong_conn eqngr; | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 834 | 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 | 835 | Symtab.update (const, consts)) consts) constss; | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 836 | fun succs consts = consts | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 837 | |> maps (Graph.imm_succs eqngr) | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 838 | |> subtract (op =) consts | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 839 | |> map (the o Symtab.lookup mapping) | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 840 | |> distinct (op =); | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 841 | val conn = [] |> fold (fn consts => cons (consts, succs consts)) constss; | 
| 31156 | 842 | 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 | 843 | |> commas; | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 844 | val prgr = map (fn (consts, constss) => | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 845 |       { 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 | 846 | path = "", parents = map namify constss }) conn; | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 847 | in Present.display_graph prgr end; | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 848 | |
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 849 | local | 
| 27103 | 850 | |
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 851 | structure P = OuterParse | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 852 | and K = OuterKeyword | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 853 | |
| 31036 | 854 | fun code_thms_cmd thy = code_thms thy o op @ o read_const_exprs thy; | 
| 855 | 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 | 856 | |
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 857 | in | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 858 | |
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 859 | val _ = | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 860 | OuterSyntax.improper_command "code_thms" "print system of code equations for code" OuterKeyword.diag | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 861 | (Scan.repeat P.term_group | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 862 | >> (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 | 863 | 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 | 864 | |
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 865 | val _ = | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 866 | OuterSyntax.improper_command "code_deps" "visualize dependencies of code equations for code" OuterKeyword.diag | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 867 | (Scan.repeat P.term_group | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 868 | >> (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 | 869 | 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 | 870 | |
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 871 | end; | 
| 27103 | 872 | |
| 24219 | 873 | end; (*struct*) | 
| 874 | ||
| 875 | ||
| 28054 | 876 | structure Basic_Code_Thingol: BASIC_CODE_THINGOL = Code_Thingol; |