| author | haftmann | 
| Thu, 10 Jan 2019 12:07:05 +0000 | |
| changeset 69623 | ef02c5e051e5 | 
| parent 69593 | 3dda49e08b9d | 
| child 69645 | e4e5bc6ac214 | 
| 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 `$$; | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 11 | infixr 3 `->; | 
| 31724 | 12 | infixr 3 `|=>; | 
| 13 | infixr 3 `|==>; | |
| 24219 | 14 | |
| 15 | signature BASIC_CODE_THINGOL = | |
| 16 | sig | |
| 17 | type vname = string; | |
| 18 | datatype dict = | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 19 | Dict of (class * class) list * plain_dict | 
| 41118 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 20 | and plain_dict = | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 21 | Dict_Const of (string * class) * dict list list | 
| 63303 | 22 |     | Dict_Var of { var: vname, index: int, length: int, class: class, unique: bool };
 | 
| 24219 | 23 | datatype itype = | 
| 24 | `%% of string * itype list | |
| 25 | | ITyVar of vname; | |
| 55150 | 26 |   type const = { sym: Code_Symbol.T, typargs: itype list, dicts: dict list list,
 | 
| 58397 | 27 | dom: itype list, annotation: itype option }; | 
| 24219 | 28 | datatype iterm = | 
| 24591 | 29 | IConst of const | 
| 31889 | 30 | | IVar of vname option | 
| 24219 | 31 | | `$ of iterm * iterm | 
| 31888 | 32 | | `|=> of (vname option * itype) * iterm | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 33 |     | ICase of { term: iterm, typ: itype, clauses: (iterm * iterm) list, primitive: iterm };
 | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 34 | val `-> : itype * itype -> itype; | 
| 24219 | 35 | val `$$ : iterm * iterm list -> iterm; | 
| 31888 | 36 | val `|==> : (vname option * itype) list * iterm -> iterm; | 
| 24219 | 37 | type typscheme = (vname * sort) list * itype; | 
| 38 | end; | |
| 39 | ||
| 40 | signature CODE_THINGOL = | |
| 41 | sig | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 42 | include BASIC_CODE_THINGOL | 
| 
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 | 
| 65483 
1cb9fd58d55e
for generated Haskell code, never use let-binds with pattern matching: irrefutable patterns destroy partial correctness
 haftmann parents: 
64957diff
changeset | 50 | val split_let_no_pat: iterm -> (((string option * itype) * iterm) * iterm) option | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 51 | val unfold_let: iterm -> ((iterm * itype) * iterm) list * iterm | 
| 65483 
1cb9fd58d55e
for generated Haskell code, never use let-binds with pattern matching: irrefutable patterns destroy partial correctness
 haftmann parents: 
64957diff
changeset | 52 | val unfold_let_no_pat: iterm -> ((string option * itype) * iterm) list * iterm | 
| 31889 | 53 | val split_pat_abs: iterm -> ((iterm * itype) * iterm) option | 
| 54 | val unfold_pat_abs: iterm -> (iterm * itype) list * iterm | |
| 31049 | 55 | val unfold_const_app: iterm -> (const * iterm list) option | 
| 32909 | 56 | val is_IVar: iterm -> bool | 
| 41782 | 57 | val is_IAbs: iterm -> bool | 
| 31049 | 58 | val eta_expand: int -> const * iterm list -> iterm | 
| 41100 
6c0940392fb4
dictionary constants must permit explicit weakening of classes;
 haftmann parents: 
40844diff
changeset | 59 | val contains_dict_var: iterm -> bool | 
| 63303 | 60 | val unambiguous_dictss: dict list list -> bool | 
| 55150 | 61 | val add_constsyms: iterm -> Code_Symbol.T list -> Code_Symbol.T list | 
| 32917 | 62 | val add_tyconames: iterm -> string list -> string list | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 63 | val fold_varnames: (string -> 'a -> 'a) -> iterm -> 'a -> 'a | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 64 | |
| 24918 | 65 | datatype stmt = | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 66 | NoStmt | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 67 | | Fun of (typscheme * ((iterm list * iterm) * (thm option * bool)) list) * thm option | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 68 | | Datatype of vname list * | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 69 | ((string * vname list (*type argument wrt. canonical order*)) * itype list) list | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 70 | | Datatypecons of string | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 71 | | Class of vname * ((class * class) list * (string * itype) list) | 
| 24219 | 72 | | Classrel of class * class | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 73 | | Classparam of class | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 74 |     | Classinst of { class: string, tyco: string, vs: (vname * sort) list,
 | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 75 | superinsts: (class * dict list list) list, | 
| 52519 
598addf65209
explicit hint for domain of class parameters in instance statements
 haftmann parents: 
52161diff
changeset | 76 | inst_params: ((string * (const * int)) * (thm * bool)) list, | 
| 
598addf65209
explicit hint for domain of class parameters in instance statements
 haftmann parents: 
52161diff
changeset | 77 | superinst_params: ((string * (const * int)) * (thm * bool)) list }; | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 78 | type program = stmt Code_Symbol.Graph.T | 
| 54889 
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
 haftmann parents: 
52801diff
changeset | 79 | val unimplemented: program -> string list | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 80 | val implemented_deps: program -> string list | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 81 | val map_terms_bottom_up: (iterm -> iterm) -> iterm -> iterm | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 82 | val map_terms_stmt: (iterm -> iterm) -> stmt -> stmt | 
| 55150 | 83 | val is_constr: program -> Code_Symbol.T -> bool | 
| 37440 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 haftmann parents: 
37437diff
changeset | 84 | val is_case: stmt -> bool | 
| 55757 | 85 | val group_stmts: Proof.context -> program | 
| 55150 | 86 | -> ((Code_Symbol.T * stmt) list * (Code_Symbol.T * stmt) list | 
| 87 | * ((Code_Symbol.T * stmt) list * (Code_Symbol.T * stmt) list)) list | |
| 24219 | 88 | |
| 55757 | 89 | val read_const_exprs: Proof.context -> string list -> string list | 
| 63159 | 90 | val consts_program: Proof.context -> string list -> program | 
| 55757 | 91 | val dynamic_conv: Proof.context -> (program | 
| 56969 
7491932da574
dropped obsolete hand-waving adjustion of type variables: safely done in preprocessor
 haftmann parents: 
56920diff
changeset | 92 | -> typscheme * iterm -> Code_Symbol.T list -> conv) | 
| 38672 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 haftmann parents: 
38669diff
changeset | 93 | -> conv | 
| 55757 | 94 | val dynamic_value: Proof.context -> ((term -> term) -> 'a -> 'a) -> (program | 
| 56969 
7491932da574
dropped obsolete hand-waving adjustion of type variables: safely done in preprocessor
 haftmann parents: 
56920diff
changeset | 95 | -> term -> typscheme * iterm -> Code_Symbol.T list -> 'a) | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 96 | -> term -> 'a | 
| 63156 | 97 |   val static_conv_thingol: { ctxt: Proof.context, consts: string list }
 | 
| 56973 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56969diff
changeset | 98 |     -> ({ program: program, deps: string list }
 | 
| 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56969diff
changeset | 99 | -> Proof.context -> typscheme * iterm -> Code_Symbol.T list -> conv) | 
| 55757 | 100 | -> Proof.context -> conv | 
| 63156 | 101 |   val static_conv_isa: { ctxt: Proof.context, consts: string list }
 | 
| 56920 
d651b944c67e
normalizing of type variables before evaluation with explicit resubstitution function: make nbe work with funny type variables like \<AA>;
 haftmann parents: 
56811diff
changeset | 102 | -> (program -> Proof.context -> term -> conv) | 
| 55757 | 103 | -> Proof.context -> conv | 
| 56973 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56969diff
changeset | 104 |   val static_value: { ctxt: Proof.context, lift_postproc: ((term -> term) -> 'a -> 'a), consts: string list }
 | 
| 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56969diff
changeset | 105 |     -> ({ program: program, deps: string list }
 | 
| 56969 
7491932da574
dropped obsolete hand-waving adjustion of type variables: safely done in preprocessor
 haftmann parents: 
56920diff
changeset | 106 | -> Proof.context -> term -> typscheme * iterm -> Code_Symbol.T list -> 'a) | 
| 55757 | 107 | -> Proof.context -> term -> 'a | 
| 24219 | 108 | end; | 
| 109 | ||
| 63160 | 110 | structure Code_Thingol : CODE_THINGOL = | 
| 24219 | 111 | struct | 
| 112 | ||
| 55150 | 113 | open Basic_Code_Symbol; | 
| 114 | ||
| 24219 | 115 | (** auxiliary **) | 
| 116 | ||
| 117 | fun unfoldl dest x = | |
| 118 | case dest x | |
| 119 | of NONE => (x, []) | |
| 120 | | SOME (x1, x2) => | |
| 121 | let val (x', xs') = unfoldl dest x1 in (x', xs' @ [x2]) end; | |
| 122 | ||
| 123 | fun unfoldr dest x = | |
| 124 | case dest x | |
| 125 | of NONE => ([], x) | |
| 126 | | SOME (x1, x2) => | |
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 127 | let val (xs', x') = unfoldr dest x2 in (x1 :: xs', x') end; | 
| 24219 | 128 | |
| 129 | ||
| 29962 
bd4dc7fa742d
tuned comments, stripped ID, deleted superfluous code
 haftmann parents: 
29952diff
changeset | 130 | (** language core - types, terms **) | 
| 24219 | 131 | |
| 132 | type vname = string; | |
| 133 | ||
| 134 | datatype dict = | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 135 | Dict of (class * class) list * plain_dict | 
| 41118 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 136 | and plain_dict = | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 137 | Dict_Const of (string * class) * dict list list | 
| 63303 | 138 |   | Dict_Var of { var: vname, index: int, length: int, class: class, unique: bool };
 | 
| 24219 | 139 | |
| 140 | datatype itype = | |
| 141 | `%% of string * itype list | |
| 142 | | ITyVar of vname; | |
| 143 | ||
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 144 | fun ty1 `-> ty2 = "fun" `%% [ty1, ty2]; | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 145 | |
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 146 | val unfold_fun = unfoldr | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 147 | (fn "fun" `%% [ty1, ty2] => SOME (ty1, ty2) | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 148 | | _ => NONE); | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 149 | |
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 150 | fun unfold_fun_n n ty = | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 151 | let | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 152 | val (tys1, ty1) = unfold_fun ty; | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 153 | val (tys3, tys2) = chop n tys1; | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 154 | val ty3 = Library.foldr (op `->) (tys2, ty1); | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 155 | in (tys3, ty3) end; | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 156 | |
| 55150 | 157 | type const = { sym: Code_Symbol.T, typargs: itype list, dicts: dict list list,
 | 
| 58397 | 158 | dom: itype list, annotation: itype option }; | 
| 24591 | 159 | |
| 24219 | 160 | datatype iterm = | 
| 24591 | 161 | IConst of const | 
| 31889 | 162 | | IVar of vname option | 
| 24219 | 163 | | `$ of iterm * iterm | 
| 31888 | 164 | | `|=> of (vname option * itype) * iterm | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 165 |   | ICase of { term: iterm, typ: itype, clauses: (iterm * iterm) list, primitive: iterm };
 | 
| 24219 | 166 | (*see also signature*) | 
| 167 | ||
| 32909 | 168 | fun is_IVar (IVar _) = true | 
| 169 | | is_IVar _ = false; | |
| 170 | ||
| 41782 | 171 | fun is_IAbs (_ `|=> _) = true | 
| 172 | | is_IAbs _ = false; | |
| 173 | ||
| 24219 | 174 | val op `$$ = Library.foldl (op `$); | 
| 31724 | 175 | val op `|==> = Library.foldr (op `|=>); | 
| 24219 | 176 | |
| 177 | val unfold_app = unfoldl | |
| 178 | (fn op `$ t => SOME t | |
| 179 | | _ => NONE); | |
| 180 | ||
| 31874 | 181 | val unfold_abs = unfoldr | 
| 182 | (fn op `|=> t => SOME t | |
| 24219 | 183 | | _ => NONE); | 
| 184 | ||
| 185 | val split_let = | |
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 186 |   (fn ICase { term = t, typ = ty, clauses = [(p, body)], ... } => SOME (((p, ty), t), body)
 | 
| 24219 | 187 | | _ => NONE); | 
| 188 | ||
| 65483 
1cb9fd58d55e
for generated Haskell code, never use let-binds with pattern matching: irrefutable patterns destroy partial correctness
 haftmann parents: 
64957diff
changeset | 189 | val split_let_no_pat = | 
| 
1cb9fd58d55e
for generated Haskell code, never use let-binds with pattern matching: irrefutable patterns destroy partial correctness
 haftmann parents: 
64957diff
changeset | 190 |   (fn ICase { term = t, typ = ty, clauses = [(IVar v, body)], ... } => SOME (((v, ty), t), body)
 | 
| 
1cb9fd58d55e
for generated Haskell code, never use let-binds with pattern matching: irrefutable patterns destroy partial correctness
 haftmann parents: 
64957diff
changeset | 191 | | _ => NONE); | 
| 
1cb9fd58d55e
for generated Haskell code, never use let-binds with pattern matching: irrefutable patterns destroy partial correctness
 haftmann parents: 
64957diff
changeset | 192 | |
| 24219 | 193 | val unfold_let = unfoldr split_let; | 
| 194 | ||
| 65483 
1cb9fd58d55e
for generated Haskell code, never use let-binds with pattern matching: irrefutable patterns destroy partial correctness
 haftmann parents: 
64957diff
changeset | 195 | val unfold_let_no_pat = unfoldr split_let_no_pat; | 
| 
1cb9fd58d55e
for generated Haskell code, never use let-binds with pattern matching: irrefutable patterns destroy partial correctness
 haftmann parents: 
64957diff
changeset | 196 | |
| 24219 | 197 | fun unfold_const_app t = | 
| 198 | case unfold_app t | |
| 199 | of (IConst c, ts) => SOME (c, ts) | |
| 200 | | _ => NONE; | |
| 201 | ||
| 32917 | 202 | fun fold_constexprs f = | 
| 203 | let | |
| 204 | fun fold' (IConst c) = f c | |
| 205 | | fold' (IVar _) = I | |
| 206 | | fold' (t1 `$ t2) = fold' t1 #> fold' t2 | |
| 207 | | fold' (_ `|=> t) = fold' t | |
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 208 |       | fold' (ICase { term = t, clauses = clauses, ... }) = fold' t
 | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 209 | #> fold (fn (p, body) => fold' p #> fold' body) clauses | 
| 32917 | 210 | in fold' end; | 
| 211 | ||
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 212 | val add_constsyms = fold_constexprs (fn { sym, ... } => insert (op =) sym);
 | 
| 32917 | 213 | |
| 214 | fun add_tycos (tyco `%% tys) = insert (op =) tyco #> fold add_tycos tys | |
| 215 | | add_tycos (ITyVar _) = I; | |
| 216 | ||
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 217 | val add_tyconames = fold_constexprs (fn { typargs = tys, ... } => fold add_tycos tys);
 | 
| 24219 | 218 | |
| 219 | fun fold_varnames f = | |
| 220 | let | |
| 59541 | 221 | fun fold_aux add_vars f = | 
| 31935 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 222 | let | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 223 | fun fold_term _ (IConst _) = I | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 224 | | 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 | 225 | | fold_term _ (IVar NONE) = I | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 226 | | 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 | 227 | | 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 | 228 | | fold_term vs ((NONE, _) `|=> t) = fold_term vs t | 
| 59541 | 229 |           | fold_term vs (ICase { term = t, clauses = clauses, ... }) =
 | 
| 230 | fold_term vs t #> fold (fold_clause vs) clauses | |
| 231 | and fold_clause vs (p, t) = fold_term (add_vars p vs) t; | |
| 232 | in fold_term [] end | |
| 233 | fun add_vars t = fold_aux add_vars (insert (op =)) t; | |
| 234 | in fold_aux add_vars f end; | |
| 24219 | 235 | |
| 31935 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 236 | fun exists_var t v = fold_varnames (fn w => fn b => v = w orelse b) t false; | 
| 31874 | 237 | |
| 31889 | 238 | fun split_pat_abs ((NONE, ty) `|=> t) = SOME ((IVar NONE, ty), t) | 
| 31888 | 239 | | split_pat_abs ((SOME v, ty) `|=> t) = SOME (case t | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 240 |      of ICase { term = IVar (SOME w), clauses = [(p, body)], ... } =>
 | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 241 | if v = w andalso (exists_var p v orelse not (exists_var body v)) | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 242 | then ((p, ty), body) | 
| 31889 | 243 | else ((IVar (SOME v), ty), t) | 
| 244 | | _ => ((IVar (SOME v), ty), t)) | |
| 31888 | 245 | | split_pat_abs _ = NONE; | 
| 31874 | 246 | |
| 247 | val unfold_pat_abs = unfoldr split_pat_abs; | |
| 24219 | 248 | |
| 31890 
e943b039f0ac
an intermediate step towards a refined translation of cases
 haftmann parents: 
31889diff
changeset | 249 | fun unfold_abs_eta [] t = ([], t) | 
| 
e943b039f0ac
an intermediate step towards a refined translation of cases
 haftmann parents: 
31889diff
changeset | 250 | | unfold_abs_eta (_ :: tys) (v_ty `|=> t) = | 
| 
e943b039f0ac
an intermediate step towards a refined translation of cases
 haftmann parents: 
31889diff
changeset | 251 | let | 
| 
e943b039f0ac
an intermediate step towards a refined translation of cases
 haftmann parents: 
31889diff
changeset | 252 | val (vs_tys, t') = unfold_abs_eta tys t; | 
| 
e943b039f0ac
an intermediate step towards a refined translation of cases
 haftmann parents: 
31889diff
changeset | 253 | in (v_ty :: vs_tys, t') end | 
| 31892 | 254 | | unfold_abs_eta tys t = | 
| 31890 
e943b039f0ac
an intermediate step towards a refined translation of cases
 haftmann parents: 
31889diff
changeset | 255 | let | 
| 
e943b039f0ac
an intermediate step towards a refined translation of cases
 haftmann parents: 
31889diff
changeset | 256 | val ctxt = fold_varnames Name.declare t Name.context; | 
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43326diff
changeset | 257 | 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 | 258 | 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 | 259 | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 260 | fun eta_expand k (const as { dom = tys, ... }, ts) =
 | 
| 24219 | 261 | let | 
| 262 | val j = length ts; | |
| 263 | val l = k - j; | |
| 37841 | 264 | val _ = if l > length tys | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 265 | then error "Impossible eta-expansion" else (); | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 266 | val vars = (fold o fold_varnames) Name.declare ts Name.context; | 
| 31889 | 267 | val vs_tys = (map o apfst) SOME | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 268 | (Name.invent_names vars "a" ((take l o drop j) tys)); | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 269 | in vs_tys `|==> IConst const `$$ ts @ map (IVar o fst) vs_tys end; | 
| 24219 | 270 | |
| 63303 | 271 | fun exists_dict_var f (Dict (_, d)) = exists_plain_dict_var_pred f d | 
| 272 | and exists_plain_dict_var_pred f (Dict_Const (_, dss)) = exists_dictss_var f dss | |
| 273 | | exists_plain_dict_var_pred f (Dict_Var x) = f x | |
| 274 | and exists_dictss_var f dss = (exists o exists) (exists_dict_var f) dss; | |
| 275 | ||
| 276 | fun contains_dict_var (IConst { dicts = dss, ... }) = exists_dictss_var (K true) dss
 | |
| 277 | | contains_dict_var (IVar _) = false | |
| 278 | | contains_dict_var (t1 `$ t2) = contains_dict_var t1 orelse contains_dict_var t2 | |
| 279 | | contains_dict_var (_ `|=> t) = contains_dict_var t | |
| 280 |   | contains_dict_var (ICase { primitive = t, ... }) = contains_dict_var t;
 | |
| 281 | ||
| 282 | val unambiguous_dictss = not o exists_dictss_var (fn { unique, ... } => not unique);
 | |
| 25621 | 283 | |
| 284 | ||
| 27103 | 285 | (** statements, abstract programs **) | 
| 24219 | 286 | |
| 287 | type typscheme = (vname * sort) list * itype; | |
| 37447 
ad3e04f289b6
transitive superclasses were also only a misunderstanding
 haftmann parents: 
37446diff
changeset | 288 | datatype stmt = | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 289 | NoStmt | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 290 | | Fun of (typscheme * ((iterm list * iterm) * (thm option * bool)) list) * thm option | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 291 | | Datatype of vname list * ((string * vname list) * itype list) list | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 292 | | Datatypecons of string | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 293 | | Class of vname * ((class * class) list * (string * itype) list) | 
| 24219 | 294 | | Classrel of class * class | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 295 | | Classparam of class | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 296 |   | Classinst of { class: string, tyco: string, vs: (vname * sort) list,
 | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 297 | superinsts: (class * dict list list) list, | 
| 52519 
598addf65209
explicit hint for domain of class parameters in instance statements
 haftmann parents: 
52161diff
changeset | 298 | inst_params: ((string * (const * int)) * (thm * bool)) list, | 
| 
598addf65209
explicit hint for domain of class parameters in instance statements
 haftmann parents: 
52161diff
changeset | 299 | superinst_params: ((string * (const * int)) * (thm * bool)) list }; | 
| 24219 | 300 | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 301 | type program = stmt Code_Symbol.Graph.T; | 
| 24219 | 302 | |
| 54889 
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
 haftmann parents: 
52801diff
changeset | 303 | fun unimplemented program = | 
| 55150 | 304 | Code_Symbol.Graph.fold (fn (Constant c, (NoStmt, _)) => cons c | _ => I) program []; | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 305 | |
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 306 | fun implemented_deps program = | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 307 | Code_Symbol.Graph.keys program | 
| 55150 | 308 | |> subtract (op =) (Code_Symbol.Graph.all_preds program (map Constant (unimplemented program))) | 
| 309 | |> map_filter (fn Constant c => SOME c | _ => NONE); | |
| 24219 | 310 | |
| 27711 | 311 | fun map_terms_bottom_up f (t as IConst _) = f t | 
| 312 | | map_terms_bottom_up f (t as IVar _) = f t | |
| 313 | | map_terms_bottom_up f (t1 `$ t2) = f | |
| 314 | (map_terms_bottom_up f t1 `$ map_terms_bottom_up f t2) | |
| 31724 | 315 | | map_terms_bottom_up f ((v, ty) `|=> t) = f | 
| 316 | ((v, ty) `|=> map_terms_bottom_up f t) | |
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 317 |   | map_terms_bottom_up f (ICase { term = t, typ = ty, clauses = clauses, primitive = t0 }) = f
 | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 318 |       (ICase { term = map_terms_bottom_up f t, typ = ty,
 | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58893diff
changeset | 319 | clauses = (map o apply2) (map_terms_bottom_up f) clauses, | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 320 | primitive = map_terms_bottom_up f t0 }); | 
| 27711 | 321 | |
| 37448 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 322 | fun map_classparam_instances_as_term f = | 
| 52519 
598addf65209
explicit hint for domain of class parameters in instance statements
 haftmann parents: 
52161diff
changeset | 323 | (map o apfst o apsnd o apfst) (fn const => case f (IConst const) of IConst const' => const') | 
| 37448 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 324 | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 325 | fun map_terms_stmt f NoStmt = NoStmt | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 326 | | map_terms_stmt f (Fun ((tysm, eqs), case_cong)) = Fun ((tysm, (map o apfst) | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 327 | (fn (ts, t) => (map f ts, f t)) eqs), case_cong) | 
| 27711 | 328 | | map_terms_stmt f (stmt as Datatype _) = stmt | 
| 329 | | map_terms_stmt f (stmt as Datatypecons _) = stmt | |
| 330 | | map_terms_stmt f (stmt as Class _) = stmt | |
| 331 | | map_terms_stmt f (stmt as Classrel _) = stmt | |
| 332 | | map_terms_stmt f (stmt as Classparam _) = stmt | |
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 333 |   | map_terms_stmt f (Classinst { class, tyco, vs, superinsts,
 | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 334 | inst_params, superinst_params }) = | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 335 |         Classinst { class = class, tyco = tyco, vs = vs, superinsts = superinsts,
 | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 336 | inst_params = map_classparam_instances_as_term f inst_params, | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 337 | superinst_params = map_classparam_instances_as_term f superinst_params }; | 
| 27711 | 338 | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 339 | fun is_constr program sym = case Code_Symbol.Graph.get_node program sym | 
| 24219 | 340 | of Datatypecons _ => true | 
| 341 | | _ => false; | |
| 342 | ||
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 343 | fun is_case (Fun (_, SOME _)) = true | 
| 37440 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 haftmann parents: 
37437diff
changeset | 344 | | is_case _ = false; | 
| 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 haftmann parents: 
37437diff
changeset | 345 | |
| 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 haftmann parents: 
37437diff
changeset | 346 | fun linear_stmts program = | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 347 | rev (Code_Symbol.Graph.strong_conn program) | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 348 | |> map (AList.make (Code_Symbol.Graph.get_node program)); | 
| 37440 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 haftmann parents: 
37437diff
changeset | 349 | |
| 55757 | 350 | fun group_stmts ctxt program = | 
| 32895 | 351 | let | 
| 352 | fun is_fun (_, Fun _) = true | is_fun _ = false; | |
| 353 | fun is_datatypecons (_, Datatypecons _) = true | is_datatypecons _ = false; | |
| 354 | fun is_datatype (_, Datatype _) = true | is_datatype _ = false; | |
| 355 | fun is_class (_, Class _) = true | is_class _ = false; | |
| 356 | fun is_classrel (_, Classrel _) = true | is_classrel _ = false; | |
| 357 | fun is_classparam (_, Classparam _) = true | is_classparam _ = false; | |
| 358 | fun is_classinst (_, Classinst _) = true | is_classinst _ = false; | |
| 359 | fun group stmts = | |
| 360 | if forall (is_datatypecons orf is_datatype) stmts | |
| 361 | then (filter is_datatype stmts, [], ([], [])) | |
| 362 | else if forall (is_class orf is_classrel orf is_classparam) stmts | |
| 363 | then ([], filter is_class stmts, ([], [])) | |
| 364 | else if forall (is_fun orf is_classinst) stmts | |
| 365 | then ([], [], List.partition is_fun stmts) | |
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51685diff
changeset | 366 |       else error ("Illegal mutual dependencies: " ^ (commas
 | 
| 55757 | 367 | o map (Code_Symbol.quote ctxt o fst)) stmts); | 
| 32895 | 368 | in | 
| 37440 
a5d44161ba2a
maintain cong rules for case combinators; more precise permissiveness
 haftmann parents: 
37437diff
changeset | 369 | linear_stmts program | 
| 32895 | 370 | |> map group | 
| 371 | end; | |
| 372 | ||
| 24219 | 373 | |
| 27103 | 374 | (** translation kernel **) | 
| 24219 | 375 | |
| 28724 | 376 | (* generic mechanisms *) | 
| 377 | ||
| 55190 | 378 | fun ensure_stmt symbolize generate x (deps, program) = | 
| 24219 | 379 | let | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 380 | val sym = symbolize x; | 
| 55190 | 381 | val add_dep = case deps of [] => I | 
| 382 | | dep :: _ => Code_Symbol.Graph.add_edge (dep, sym); | |
| 47576 | 383 | in | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 384 | if can (Code_Symbol.Graph.get_node program) sym | 
| 47576 | 385 | then | 
| 386 | program | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 387 | |> add_dep | 
| 55190 | 388 | |> pair deps | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 389 | |> pair x | 
| 47576 | 390 | else | 
| 391 | program | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 392 | |> Code_Symbol.Graph.default_node (sym, NoStmt) | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 393 | |> add_dep | 
| 55190 | 394 | |> curry generate (sym :: deps) | 
| 47576 | 395 | ||> snd | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 396 | |-> (fn stmt => (Code_Symbol.Graph.map_node sym) (K stmt)) | 
| 55190 | 397 | |> pair deps | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 398 | |> pair x | 
| 24219 | 399 | end; | 
| 400 | ||
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 401 | exception PERMISSIVE of unit; | 
| 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 402 | |
| 56920 
d651b944c67e
normalizing of type variables before evaluation with explicit resubstitution function: make nbe work with funny type variables like \<AA>;
 haftmann parents: 
56811diff
changeset | 403 | fun translation_error ctxt permissive some_thm deps msg sub_msg = | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 404 | if permissive | 
| 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 405 | then raise PERMISSIVE () | 
| 42385 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 406 | else | 
| 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 407 | let | 
| 55190 | 408 | val thm_msg = | 
| 61268 | 409 | Option.map (fn thm => "in code equation " ^ Thm.string_of_thm ctxt thm) some_thm; | 
| 55190 | 410 | val dep_msg = if null (tl deps) then NONE | 
| 411 |         else SOME ("with dependency "
 | |
| 412 | ^ space_implode " -> " (map (Code_Symbol.quote ctxt) (rev deps))); | |
| 413 | val thm_dep_msg = case (thm_msg, dep_msg) | |
| 414 |        of (SOME thm_msg, SOME dep_msg) => "\n(" ^ thm_msg ^ ",\n" ^ dep_msg ^ ")"
 | |
| 415 |         | (SOME thm_msg, NONE) => "\n(" ^ thm_msg ^ ")"
 | |
| 416 |         | (NONE, SOME dep_msg) => "\n(" ^ dep_msg ^ ")"
 | |
| 417 | | (NONE, NONE) => "" | |
| 418 | in error (msg ^ thm_dep_msg ^ ":\n" ^ sub_msg) end; | |
| 37698 | 419 | |
| 48074 | 420 | fun maybe_permissive f prgrm = | 
| 421 | f prgrm |>> SOME handle PERMISSIVE () => (NONE, prgrm); | |
| 422 | ||
| 56920 
d651b944c67e
normalizing of type variables before evaluation with explicit resubstitution function: make nbe work with funny type variables like \<AA>;
 haftmann parents: 
56811diff
changeset | 423 | fun not_wellsorted ctxt permissive some_thm deps ty sort e = | 
| 37698 | 424 | let | 
| 61262 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
60697diff
changeset | 425 | val err_class = Sorts.class_error (Context.Proof ctxt) e; | 
| 42385 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 426 | val err_typ = | 
| 55757 | 427 | "Type " ^ Syntax.string_of_typ ctxt ty ^ " not of sort " ^ | 
| 428 | Syntax.string_of_sort ctxt sort; | |
| 42385 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 429 | in | 
| 56920 
d651b944c67e
normalizing of type variables before evaluation with explicit resubstitution function: make nbe work with funny type variables like \<AA>;
 haftmann parents: 
56811diff
changeset | 430 | translation_error ctxt permissive some_thm deps | 
| 55190 | 431 | "Wellsortedness error" (err_typ ^ "\n" ^ err_class) | 
| 42385 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 432 | end; | 
| 26972 | 433 | |
| 47555 | 434 | |
| 44790 
c13fdf710a40
adding type inference for disambiguation annotations in code equation
 bulwahn parents: 
44789diff
changeset | 435 | (* inference of type annotations for disambiguation with type classes *) | 
| 
c13fdf710a40
adding type inference for disambiguation annotations in code equation
 bulwahn parents: 
44789diff
changeset | 436 | |
| 45000 
0febe2089425
adding abstraction layer; more precise function names
 bulwahn parents: 
44999diff
changeset | 437 | fun mk_tagged_type (true, T) = Type ("", [T])
 | 
| 47555 | 438 | | mk_tagged_type (false, T) = T; | 
| 45000 
0febe2089425
adding abstraction layer; more precise function names
 bulwahn parents: 
44999diff
changeset | 439 | |
| 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 | 440 | fun dest_tagged_type (Type ("", [T])) = (true, T)
 | 
| 47555 | 441 | | dest_tagged_type T = (false, T); | 
| 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 | 442 | |
| 47555 | 443 | val untag_term = map_types (snd o dest_tagged_type); | 
| 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 | 444 | |
| 45000 
0febe2089425
adding abstraction layer; more precise function names
 bulwahn parents: 
44999diff
changeset | 445 | fun tag_term (proj_sort, _) eqngr = | 
| 44997 | 446 | let | 
| 47555 | 447 | val has_sort_constraints = exists (not o null) o map proj_sort o Code_Preproc.sortargs eqngr; | 
| 47576 | 448 | fun tag (Const (_, T')) (Const (c, T)) = | 
| 45000 
0febe2089425
adding abstraction layer; more precise function names
 bulwahn parents: 
44999diff
changeset | 449 | Const (c, | 
| 
0febe2089425
adding abstraction layer; more precise function names
 bulwahn parents: 
44999diff
changeset | 450 | 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 | 451 | | tag (t1 $ u1) (t $ u) = tag t1 t $ tag u1 u | 
| 
0febe2089425
adding abstraction layer; more precise function names
 bulwahn parents: 
44999diff
changeset | 452 | | 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 | 453 | | tag (Free _) (t as Free _) = t | 
| 
0febe2089425
adding abstraction layer; more precise function names
 bulwahn parents: 
44999diff
changeset | 454 | | tag (Var _) (t as Var _) = t | 
| 47555 | 455 | | tag (Bound _) (t as Bound _) = t; | 
| 55757 | 456 | in tag end | 
| 44790 
c13fdf710a40
adding type inference for disambiguation annotations in code equation
 bulwahn parents: 
44789diff
changeset | 457 | |
| 55757 | 458 | fun annotate ctxt algbr eqngr (c, ty) args rhs = | 
| 44790 
c13fdf710a40
adding type inference for disambiguation annotations in code equation
 bulwahn parents: 
44789diff
changeset | 459 | let | 
| 55757 | 460 | val erase = map_types (fn _ => Type_Infer.anyT []); | 
| 461 | val reinfer = singleton (Type_Infer_Context.infer_types ctxt); | |
| 462 | val lhs = list_comb (Const (c, ty), map (map_types Type.strip_sorts o fst) args); | |
| 463 | val reinferred_rhs = snd (Logic.dest_equals (reinfer (Logic.mk_equals (lhs, erase rhs)))); | |
| 464 | in tag_term algbr eqngr reinferred_rhs rhs end | |
| 465 | ||
| 466 | fun annotate_eqns ctxt algbr eqngr (c, ty) eqns = | |
| 467 | let | |
| 468 | val ctxt' = ctxt |> Proof_Context.theory_of |> Proof_Context.init_global | |
| 469 | |> Config.put Type_Infer_Context.const_sorts false; | |
| 470 | (*avoid spurious fixed variables: there is no eigen context for equations*) | |
| 44790 
c13fdf710a40
adding type inference for disambiguation annotations in code equation
 bulwahn parents: 
44789diff
changeset | 471 | in | 
| 55757 | 472 | map (apfst (fn (args, (rhs, some_abs)) => (args, | 
| 473 | (annotate ctxt' algbr eqngr (c, ty) args rhs, some_abs)))) eqns | |
| 474 | end; | |
| 47555 | 475 | |
| 55189 | 476 | (* abstract dictionary construction *) | 
| 477 | ||
| 478 | datatype typarg_witness = | |
| 479 | Weakening of (class * class) list * plain_typarg_witness | |
| 480 | and plain_typarg_witness = | |
| 481 | Global of (string * class) * typarg_witness list list | |
| 63303 | 482 |   | Local of { var: string, index: int, sort: sort, unique: bool };
 | 
| 483 | ||
| 484 | fun brand_unique unique (w as Global _) = w | |
| 485 |   | brand_unique unique (Local { var, index, sort, unique = _ }) =
 | |
| 486 |       Local { var = var, index = index, sort = sort, unique = unique };
 | |
| 55189 | 487 | |
| 55757 | 488 | fun construct_dictionaries ctxt (proj_sort, algebra) permissive some_thm (ty, sort) (deps, program) = | 
| 55189 | 489 | let | 
| 63303 | 490 | fun class_relation unique (Weakening (classrels, x), sub_class) super_class = | 
| 491 | Weakening ((sub_class, super_class) :: classrels, brand_unique unique x); | |
| 55189 | 492 | fun type_constructor (tyco, _) dss class = | 
| 493 | Weakening ([], Global ((tyco, class), (map o map) fst dss)); | |
| 494 | fun type_variable (TFree (v, sort)) = | |
| 495 | let | |
| 496 | val sort' = proj_sort sort; | |
| 63303 | 497 | in map_index (fn (n, class) => (Weakening ([], Local | 
| 498 |         { var = v, index = n, sort = sort', unique = true }), class)) sort'
 | |
| 499 | end; | |
| 55189 | 500 | val typarg_witnesses = Sorts.of_sort_derivation algebra | 
| 63303 | 501 |       {class_relation = fn _ => fn unique =>
 | 
| 502 | Sorts.classrel_derivation algebra (class_relation unique), | |
| 55189 | 503 | type_constructor = type_constructor, | 
| 504 | type_variable = type_variable} (ty, proj_sort sort) | |
| 56920 
d651b944c67e
normalizing of type variables before evaluation with explicit resubstitution function: make nbe work with funny type variables like \<AA>;
 haftmann parents: 
56811diff
changeset | 505 | handle Sorts.CLASS_ERROR e => not_wellsorted ctxt permissive some_thm deps ty sort e; | 
| 55190 | 506 | in (typarg_witnesses, (deps, program)) end; | 
| 55189 | 507 | |
| 508 | ||
| 28724 | 509 | (* translation *) | 
| 510 | ||
| 55757 | 511 | fun ensure_tyco ctxt algbr eqngr permissive tyco = | 
| 30932 | 512 | let | 
| 55757 | 513 | val thy = Proof_Context.theory_of ctxt; | 
| 40726 
16dcfedc4eb7
keep type variable arguments of datatype constructors in bookkeeping
 haftmann parents: 
40711diff
changeset | 514 | val ((vs, cos), _) = Code.get_type thy tyco; | 
| 30932 | 515 | val stmt_datatype = | 
| 55757 | 516 | fold_map (translate_tyvar_sort ctxt algbr eqngr permissive) vs | 
| 48003 
1d11af40b106
dropped sort constraints on datatype specifications
 haftmann parents: 
47576diff
changeset | 517 | #>> map fst | 
| 40726 
16dcfedc4eb7
keep type variable arguments of datatype constructors in bookkeeping
 haftmann parents: 
40711diff
changeset | 518 | ##>> fold_map (fn (c, (vs, tys)) => | 
| 55757 | 519 | ensure_const ctxt algbr eqngr permissive c | 
| 40726 
16dcfedc4eb7
keep type variable arguments of datatype constructors in bookkeeping
 haftmann parents: 
40711diff
changeset | 520 | ##>> pair (map (unprefix "'" o fst) vs) | 
| 55757 | 521 | ##>> fold_map (translate_typ ctxt algbr eqngr permissive) tys) cos | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 522 | #>> Datatype; | 
| 55150 | 523 | in ensure_stmt Type_Constructor stmt_datatype tyco end | 
| 55757 | 524 | and ensure_const ctxt algbr eqngr permissive c = | 
| 30932 | 525 | let | 
| 55757 | 526 | val thy = Proof_Context.theory_of ctxt; | 
| 30932 | 527 | fun stmt_datatypecons tyco = | 
| 55757 | 528 | ensure_tyco ctxt algbr eqngr permissive tyco | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 529 | #>> Datatypecons; | 
| 30932 | 530 | fun stmt_classparam class = | 
| 55757 | 531 | ensure_class ctxt algbr eqngr permissive class | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 532 | #>> Classparam; | 
| 54889 
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
 haftmann parents: 
52801diff
changeset | 533 | fun stmt_fun cert = case Code.equations_of_cert thy cert | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 534 | of (_, NONE) => pair NoStmt | 
| 54889 
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
 haftmann parents: 
52801diff
changeset | 535 | | ((vs, ty), SOME eqns) => | 
| 
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
 haftmann parents: 
52801diff
changeset | 536 | let | 
| 55757 | 537 | val eqns' = annotate_eqns ctxt algbr eqngr (c, ty) eqns | 
| 54889 
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
 haftmann parents: 
52801diff
changeset | 538 | val some_case_cong = Code.get_case_cong thy c; | 
| 
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
 haftmann parents: 
52801diff
changeset | 539 | in | 
| 55757 | 540 | fold_map (translate_tyvar_sort ctxt algbr eqngr permissive) vs | 
| 541 | ##>> translate_typ ctxt algbr eqngr permissive ty | |
| 542 | ##>> translate_eqns ctxt algbr eqngr permissive eqns' | |
| 54889 
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
 haftmann parents: 
52801diff
changeset | 543 | #>> | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 544 | (fn (_, NONE) => NoStmt | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 545 | | (tyscm, SOME eqns) => Fun ((tyscm, eqns), some_case_cong)) | 
| 54889 
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
 haftmann parents: 
52801diff
changeset | 546 | end; | 
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 547 | val stmt_const = case Code.get_type_of_constr_or_abstr thy c | 
| 35226 | 548 | of SOME (tyco, _) => stmt_datatypecons tyco | 
| 51685 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
51658diff
changeset | 549 | | NONE => (case Axclass.class_of_param thy c | 
| 30932 | 550 | of SOME class => stmt_classparam class | 
| 34891 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34251diff
changeset | 551 | | NONE => stmt_fun (Code_Preproc.cert eqngr c)) | 
| 55150 | 552 | in ensure_stmt Constant stmt_const c end | 
| 55757 | 553 | and ensure_class ctxt (algbr as (_, algebra)) eqngr permissive class = | 
| 24918 | 554 | let | 
| 55757 | 555 | val thy = Proof_Context.theory_of ctxt; | 
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 556 | val super_classes = (Sorts.minimize_sort algebra o Sorts.super_classes algebra) class; | 
| 51685 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
51658diff
changeset | 557 | val cs = #params (Axclass.get_info thy class); | 
| 24918 | 558 | val stmt_class = | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 559 | fold_map (fn super_class => | 
| 55757 | 560 | ensure_classrel ctxt algbr eqngr permissive (class, super_class)) super_classes | 
| 561 | ##>> fold_map (fn (c, ty) => ensure_const ctxt algbr eqngr permissive c | |
| 562 | ##>> translate_typ ctxt algbr eqngr permissive ty) cs | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 563 | #>> (fn info => Class (unprefix "'" Name.aT, info)) | 
| 55150 | 564 | in ensure_stmt Type_Class stmt_class class end | 
| 55757 | 565 | and ensure_classrel ctxt algbr eqngr permissive (sub_class, super_class) = | 
| 24918 | 566 | let | 
| 567 | val stmt_classrel = | |
| 55757 | 568 | ensure_class ctxt algbr eqngr permissive sub_class | 
| 569 | ##>> ensure_class ctxt algbr eqngr permissive super_class | |
| 24918 | 570 | #>> Classrel; | 
| 55150 | 571 | in ensure_stmt Class_Relation stmt_classrel (sub_class, super_class) end | 
| 55757 | 572 | and ensure_inst ctxt (algbr as (_, algebra)) eqngr permissive (tyco, class) = | 
| 24918 | 573 | let | 
| 55757 | 574 | val thy = Proof_Context.theory_of ctxt; | 
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 575 | val super_classes = (Sorts.minimize_sort algebra o Sorts.super_classes algebra) class; | 
| 51685 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
51658diff
changeset | 576 | val these_class_params = these o try (#params o Axclass.get_info thy); | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 577 | val class_params = these_class_params class; | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 578 | val superclass_params = maps these_class_params | 
| 37448 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37447diff
changeset | 579 | ((Sorts.complete_sort algebra o Sorts.super_classes algebra) class); | 
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43326diff
changeset | 580 | val vs = Name.invent_names Name.context "'a" (Sorts.mg_domain algebra tyco [class]); | 
| 24918 | 581 | val sorts' = Sorts.mg_domain (Sign.classes_of thy) tyco [class]; | 
| 582 | val vs' = map2 (fn (v, sort1) => fn sort2 => (v, | |
| 583 | Sorts.inter_sort (Sign.classes_of thy) (sort1, sort2))) vs sorts'; | |
| 584 | val arity_typ = Type (tyco, map TFree vs); | |
| 585 | 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 | 586 | fun translate_super_instance super_class = | 
| 55757 | 587 | ensure_class ctxt algbr eqngr permissive super_class | 
| 588 | ##>> translate_dicts ctxt algbr eqngr permissive NONE (arity_typ, [super_class]) | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 589 | #>> (fn (super_class, [Dict ([], Dict_Const (_, dss))]) => (super_class, dss)); | 
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 590 | fun translate_classparam_instance (c, ty) = | 
| 24918 | 591 | let | 
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 592 | val raw_const = Const (c, map_type_tfree (K arity_typ') ty); | 
| 52519 
598addf65209
explicit hint for domain of class parameters in instance statements
 haftmann parents: 
52161diff
changeset | 593 | val dom_length = length (fst (strip_type ty)) | 
| 63073 
413184c7a2a2
clarified context, notably for internal use of Simplifier;
 wenzelm parents: 
62539diff
changeset | 594 | val thm = Axclass.unoverload_conv ctxt (Thm.cterm_of ctxt raw_const); | 
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 595 | val const = (apsnd Logic.unvarifyT_global o dest_Const o snd | 
| 24918 | 596 | o Logic.dest_equals o Thm.prop_of) thm; | 
| 597 | in | |
| 55757 | 598 | ensure_const ctxt algbr eqngr permissive c | 
| 599 | ##>> translate_const ctxt algbr eqngr permissive (SOME thm) (const, NONE) | |
| 52519 
598addf65209
explicit hint for domain of class parameters in instance statements
 haftmann parents: 
52161diff
changeset | 600 | #>> (fn (c, IConst const') => ((c, (const', dom_length)), (thm, true))) | 
| 24918 | 601 | end; | 
| 602 | val stmt_inst = | |
| 55757 | 603 | ensure_class ctxt algbr eqngr permissive class | 
| 604 | ##>> ensure_tyco ctxt algbr eqngr permissive tyco | |
| 605 | ##>> fold_map (translate_tyvar_sort ctxt algbr eqngr permissive) vs | |
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37216diff
changeset | 606 | ##>> fold_map translate_super_instance super_classes | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 607 | ##>> fold_map translate_classparam_instance class_params | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 608 | ##>> fold_map translate_classparam_instance superclass_params | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 609 | #>> (fn (((((class, tyco), vs), superinsts), inst_params), superinst_params) => | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 610 |           Classinst { class = class, tyco = tyco, vs = vs,
 | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 611 | superinsts = superinsts, inst_params = inst_params, superinst_params = superinst_params }); | 
| 55150 | 612 | in ensure_stmt Class_Instance stmt_inst (tyco, class) end | 
| 55757 | 613 | and translate_typ ctxt algbr eqngr permissive (TFree (v, _)) = | 
| 30932 | 614 | pair (ITyVar (unprefix "'" v)) | 
| 55757 | 615 | | translate_typ ctxt algbr eqngr permissive (Type (tyco, tys)) = | 
| 616 | ensure_tyco ctxt algbr eqngr permissive tyco | |
| 617 | ##>> fold_map (translate_typ ctxt algbr eqngr permissive) tys | |
| 30932 | 618 | #>> (fn (tyco, tys) => tyco `%% tys) | 
| 55757 | 619 | and translate_term ctxt algbr eqngr permissive some_thm (Const (c, ty), some_abs) = | 
| 620 | translate_app ctxt algbr eqngr permissive some_thm (((c, ty), []), some_abs) | |
| 621 | | translate_term ctxt algbr eqngr permissive some_thm (Free (v, _), some_abs) = | |
| 31889 | 622 | pair (IVar (SOME v)) | 
| 55757 | 623 | | translate_term ctxt algbr eqngr permissive some_thm (Abs (v, ty, t), some_abs) = | 
| 24918 | 624 | let | 
| 56811 | 625 | val (v', t') = Syntax_Trans.variant_abs (Name.desymbolize (SOME false) v, ty, t); | 
| 32273 | 626 | val v'' = if member (op =) (Term.add_free_names t' []) v' | 
| 627 | then SOME v' else NONE | |
| 24918 | 628 | in | 
| 55757 | 629 | translate_typ ctxt algbr eqngr permissive ty | 
| 630 | ##>> translate_term ctxt algbr eqngr permissive some_thm (t', some_abs) | |
| 32273 | 631 | #>> (fn (ty, t) => (v'', ty) `|=> t) | 
| 24918 | 632 | end | 
| 55757 | 633 | | translate_term ctxt algbr eqngr permissive some_thm (t as _ $ _, some_abs) = | 
| 24918 | 634 | case strip_comb t | 
| 635 | of (Const (c, ty), ts) => | |
| 55757 | 636 | translate_app ctxt algbr eqngr permissive some_thm (((c, ty), ts), some_abs) | 
| 24918 | 637 | | (t', ts) => | 
| 55757 | 638 | translate_term ctxt algbr eqngr permissive some_thm (t', some_abs) | 
| 639 | ##>> fold_map (translate_term ctxt algbr eqngr permissive some_thm o rpair NONE) ts | |
| 24918 | 640 | #>> (fn (t, ts) => t `$$ ts) | 
| 55757 | 641 | and translate_eqn ctxt algbr eqngr permissive ((args, (rhs, some_abs)), (some_thm, proper)) = | 
| 642 | fold_map (translate_term ctxt algbr eqngr permissive some_thm) args | |
| 643 | ##>> translate_term ctxt algbr eqngr permissive some_thm (rhs, some_abs) | |
| 35226 | 644 | #>> rpair (some_thm, proper) | 
| 55757 | 645 | and translate_eqns ctxt algbr eqngr permissive eqns = | 
| 646 | maybe_permissive (fold_map (translate_eqn ctxt algbr eqngr permissive) eqns) | |
| 647 | and translate_const ctxt algbr eqngr permissive some_thm ((c, ty), some_abs) (deps, program) = | |
| 30932 | 648 | let | 
| 55757 | 649 | val thy = Proof_Context.theory_of ctxt; | 
| 37698 | 650 | val _ = if (case some_abs of NONE => true | SOME abs => not (c = abs)) | 
| 35226 | 651 | andalso Code.is_abstr thy c | 
| 56920 
d651b944c67e
normalizing of type variables before evaluation with explicit resubstitution function: make nbe work with funny type variables like \<AA>;
 haftmann parents: 
56811diff
changeset | 652 | then translation_error ctxt permissive some_thm deps | 
| 37698 | 653 |           "Abstraction violation" ("constant " ^ Code.string_of_const thy c)
 | 
| 654 | else () | |
| 55757 | 655 | in translate_const_proper ctxt algbr eqngr permissive some_thm (c, ty) (deps, program) end | 
| 656 | and translate_const_proper ctxt algbr eqngr permissive some_thm (c, ty) = | |
| 55190 | 657 | let | 
| 55757 | 658 | val thy = Proof_Context.theory_of ctxt; | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 659 | val (annotate, ty') = dest_tagged_type ty; | 
| 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 660 | val typargs = Sign.const_typargs thy (c, ty'); | 
| 32873 | 661 | val sorts = Code_Preproc.sortargs eqngr c; | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 662 | val (dom, range) = Term.strip_type ty'; | 
| 26972 | 663 | in | 
| 55757 | 664 | ensure_const ctxt algbr eqngr permissive c | 
| 665 | ##>> fold_map (translate_typ ctxt algbr eqngr permissive) typargs | |
| 666 | ##>> fold_map (translate_dicts ctxt algbr eqngr permissive some_thm) (typargs ~~ sorts) | |
| 58397 | 667 | ##>> fold_map (translate_typ ctxt algbr eqngr permissive) (ty' :: dom) | 
| 668 | #>> (fn (((c, typargs), dss), annotation :: dom) => | |
| 55150 | 669 |       IConst { sym = Constant c, typargs = typargs, dicts = dss,
 | 
| 58397 | 670 | dom = dom, annotation = | 
| 671 | if annotate then SOME annotation else NONE }) | |
| 26972 | 672 | end | 
| 55757 | 673 | and translate_app_const ctxt algbr eqngr permissive some_thm ((c_ty, ts), some_abs) = | 
| 674 | translate_const ctxt algbr eqngr permissive some_thm (c_ty, some_abs) | |
| 675 | ##>> fold_map (translate_term ctxt algbr eqngr permissive some_thm o rpair NONE) ts | |
| 24918 | 676 | #>> (fn (t, ts) => t `$$ ts) | 
| 55757 | 677 | and translate_case ctxt algbr eqngr permissive some_thm (num_args, (t_pos, case_pats)) (c_ty, ts) = | 
| 24918 | 678 | let | 
| 55757 | 679 | val thy = Proof_Context.theory_of ctxt; | 
| 40844 | 680 | fun arg_types num_args ty = fst (chop num_args (binder_types ty)); | 
| 31892 | 681 | 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 | 682 | val ty = nth tys t_pos; | 
| 47437 
4625ee486ff6
generalise case certificates to allow ignored parameters
 Andreas Lochbihler parents: 
47005diff
changeset | 683 | fun mk_constr NONE t = NONE | 
| 47555 | 684 | | mk_constr (SOME c) t = | 
| 685 | let | |
| 686 | val n = Code.args_number thy c; | |
| 687 | in SOME ((c, arg_types n (fastype_of (untag_term t)) ---> ty), n) end; | |
| 688 | val constrs = | |
| 689 | if null case_pats then [] | |
| 690 | else map_filter I (map2 mk_constr case_pats (nth_drop t_pos ts)); | |
| 59544 | 691 | fun disjunctive_varnames ts = | 
| 692 | let | |
| 693 | val vs = (fold o fold_varnames) (insert (op =)) ts []; | |
| 694 | in fn pat => null (inter (op =) vs (fold_varnames (insert (op =)) pat [])) end; | |
| 695 | fun purge_unused_vars_in t = | |
| 696 | let | |
| 697 | val vs = fold_varnames (insert (op =)) t []; | |
| 698 | in | |
| 699 | map_terms_bottom_up (fn IVar (SOME v) => | |
| 700 | IVar (if member (op =) vs v then SOME v else NONE) | t => t) | |
| 701 | end; | |
| 702 | fun collapse_clause vs_map ts body = | |
| 703 | case body | |
| 66189 
23917e861eaa
treat "undefined" constants internally as special form of case combinators
 haftmann parents: 
65483diff
changeset | 704 |        of IConst { sym = Constant c, ... } => if Code.is_undefined thy c
 | 
| 59544 | 705 | then [] | 
| 706 | else [(ts, body)] | |
| 707 |         | ICase { term = IVar (SOME v), clauses = clauses, ... } =>
 | |
| 708 | if forall (fn (pat', body') => exists_var pat' v | |
| 709 | orelse not (exists_var body' v)) clauses | |
| 710 | andalso forall (disjunctive_varnames ts o fst) clauses | |
| 711 | then case AList.lookup (op =) vs_map v | |
| 712 | of SOME i => maps (fn (pat', body') => | |
| 713 | collapse_clause (AList.delete (op =) v vs_map) | |
| 714 | (nth_map i (K pat') ts |> map (purge_unused_vars_in body')) body') clauses | |
| 715 | | NONE => [(ts, body)] | |
| 716 | else [(ts, body)] | |
| 717 | | _ => [(ts, body)]; | |
| 718 | fun mk_clause mk tys t = | |
| 719 | let | |
| 720 | val (vs, body) = unfold_abs_eta tys t; | |
| 721 | val vs_map = fold_index (fn (i, (SOME v, _)) => cons (v, i) | _ => I) vs []; | |
| 722 | val ts = map (IVar o fst) vs; | |
| 723 | in map mk (collapse_clause vs_map ts body) end; | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 724 | fun casify constrs ty t_app ts = | 
| 24918 | 725 | let | 
| 31892 | 726 | val t = nth ts t_pos; | 
| 727 | val ts_clause = nth_drop t_pos ts; | |
| 31935 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 728 | val clauses = if null case_pats | 
| 
3896169e6ff9
cleaned up fundamental iml term functions; nested patterns
 haftmann parents: 
31892diff
changeset | 729 | then mk_clause (fn ([t], body) => (t, body)) [ty] (the_single ts_clause) | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 730 |           else maps (fn ((constr as IConst { dom = tys, ... }, n), t) =>
 | 
| 33957 | 731 | mk_clause (fn (ts, body) => (constr `$$ ts, body)) (take n tys) t) | 
| 47555 | 732 | (constrs ~~ (map_filter (fn (NONE, _) => NONE | (SOME _, t) => SOME t) | 
| 733 | (case_pats ~~ ts_clause))); | |
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
48003diff
changeset | 734 |       in ICase { term = t, typ = ty, clauses = clauses, primitive = t_app `$$ ts } end;
 | 
| 24918 | 735 | in | 
| 55757 | 736 | translate_const ctxt algbr eqngr permissive some_thm (c_ty, NONE) | 
| 737 | ##>> fold_map (fn (constr, n) => translate_const ctxt algbr eqngr permissive some_thm (constr, NONE) | |
| 47555 | 738 | #>> rpair n) constrs | 
| 55757 | 739 | ##>> translate_typ ctxt algbr eqngr permissive ty | 
| 740 | ##>> fold_map (translate_term ctxt algbr eqngr permissive some_thm o rpair NONE) ts | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 741 | #>> (fn (((t, constrs), ty), ts) => | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 742 | casify constrs ty t ts) | 
| 24918 | 743 | end | 
| 66189 
23917e861eaa
treat "undefined" constants internally as special form of case combinators
 haftmann parents: 
65483diff
changeset | 744 | and translate_app_case ctxt algbr eqngr permissive some_thm (case_schema as (num_args, _)) ((c, ty), ts) = | 
| 29973 | 745 | if length ts < num_args then | 
| 746 | let | |
| 747 | val k = length ts; | |
| 33957 | 748 | val tys = (take (num_args - k) o drop k o fst o strip_type) ty; | 
| 55757 | 749 | val names = (fold o fold_aterms) Term.declare_term_frees ts Name.context; | 
| 750 | val vs = Name.invent_names names "a" tys; | |
| 29973 | 751 | in | 
| 55757 | 752 | fold_map (translate_typ ctxt algbr eqngr permissive) tys | 
| 66189 
23917e861eaa
treat "undefined" constants internally as special form of case combinators
 haftmann parents: 
65483diff
changeset | 753 | ##>> translate_case ctxt algbr eqngr permissive some_thm case_schema ((c, ty), ts @ map Free vs) | 
| 31888 | 754 | #>> (fn (tys, t) => map2 (fn (v, _) => pair (SOME v)) vs tys `|==> t) | 
| 29973 | 755 | end | 
| 756 | else if length ts > num_args then | |
| 66189 
23917e861eaa
treat "undefined" constants internally as special form of case combinators
 haftmann parents: 
65483diff
changeset | 757 | translate_case ctxt algbr eqngr permissive some_thm case_schema ((c, ty), take num_args ts) | 
| 55757 | 758 | ##>> fold_map (translate_term ctxt algbr eqngr permissive some_thm o rpair NONE) (drop num_args ts) | 
| 29973 | 759 | #>> (fn (t, ts) => t `$$ ts) | 
| 760 | else | |
| 66189 
23917e861eaa
treat "undefined" constants internally as special form of case combinators
 haftmann parents: 
65483diff
changeset | 761 | translate_case ctxt algbr eqngr permissive some_thm case_schema ((c, ty), ts) | 
| 55757 | 762 | and translate_app ctxt algbr eqngr permissive some_thm (c_ty_ts as ((c, _), _), some_abs) = | 
| 66189 
23917e861eaa
treat "undefined" constants internally as special form of case combinators
 haftmann parents: 
65483diff
changeset | 763 | case Code.get_case_schema (Proof_Context.theory_of ctxt) c | 
| 
23917e861eaa
treat "undefined" constants internally as special form of case combinators
 haftmann parents: 
65483diff
changeset | 764 | of SOME case_schema => translate_app_case ctxt algbr eqngr permissive some_thm case_schema c_ty_ts | 
| 55757 | 765 | | NONE => translate_app_const ctxt algbr eqngr permissive some_thm (c_ty_ts, some_abs) | 
| 766 | and translate_tyvar_sort ctxt (algbr as (proj_sort, _)) eqngr permissive (v, sort) = | |
| 767 | fold_map (ensure_class ctxt algbr eqngr permissive) (proj_sort sort) | |
| 30932 | 768 | #>> (fn sort => (unprefix "'" v, sort)) | 
| 55757 | 769 | and translate_dicts ctxt algbr eqngr permissive some_thm (ty, sort) = | 
| 30932 | 770 | let | 
| 63303 | 771 | fun mk_dict (Weakening (classrels, d)) = | 
| 55757 | 772 | fold_map (ensure_classrel ctxt algbr eqngr permissive) classrels | 
| 63303 | 773 | ##>> mk_plain_dict d | 
| 41118 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 774 | #>> Dict | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 775 | and mk_plain_dict (Global (inst, dss)) = | 
| 55757 | 776 | ensure_inst ctxt algbr eqngr permissive inst | 
| 41100 
6c0940392fb4
dictionary constants must permit explicit weakening of classes;
 haftmann parents: 
40844diff
changeset | 777 | ##>> (fold_map o fold_map) mk_dict dss | 
| 63303 | 778 | #>> Dict_Const | 
| 779 |       | mk_plain_dict (Local { var, index, sort, unique }) =
 | |
| 780 | ensure_class ctxt algbr eqngr permissive (nth sort index) | |
| 781 |           #>> (fn class => Dict_Var { var = unprefix "'" var, index = index,
 | |
| 782 | length = length sort, class = class, unique = unique }) | |
| 55189 | 783 | in | 
| 55757 | 784 | construct_dictionaries ctxt algbr permissive some_thm (ty, sort) | 
| 55189 | 785 | #-> (fn typarg_witnesses => fold_map mk_dict typarg_witnesses) | 
| 786 | end; | |
| 24918 | 787 | |
| 25969 | 788 | |
| 27103 | 789 | (* store *) | 
| 790 | ||
| 34173 
458ced35abb8
reduced code generator cache to the baremost minimum
 haftmann parents: 
34084diff
changeset | 791 | structure Program = Code_Data | 
| 27103 | 792 | ( | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 793 | type T = program; | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 794 | val empty = Code_Symbol.Graph.empty; | 
| 27103 | 795 | ); | 
| 796 | ||
| 63160 | 797 | fun invoke_generation ignore_cache ctxt generate thing = | 
| 63159 | 798 | Program.change_yield | 
| 799 | (if ignore_cache then NONE else SOME (Proof_Context.theory_of ctxt)) | |
| 55190 | 800 | (fn program => ([], program) | 
| 63160 | 801 | |> generate thing | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 802 | |-> (fn thing => fn (_, program) => (thing, program))); | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 803 | |
| 27103 | 804 | |
| 805 | (* program generation *) | |
| 806 | ||
| 63175 
d191892b1c23
explicit check that abstract constructors cannot be part of official interface
 haftmann parents: 
63164diff
changeset | 807 | fun check_abstract_constructors thy consts = | 
| 
d191892b1c23
explicit check that abstract constructors cannot be part of official interface
 haftmann parents: 
63164diff
changeset | 808 | case filter (Code.is_abstr thy) consts of | 
| 
d191892b1c23
explicit check that abstract constructors cannot be part of official interface
 haftmann parents: 
63164diff
changeset | 809 | [] => () | 
| 
d191892b1c23
explicit check that abstract constructors cannot be part of official interface
 haftmann parents: 
63164diff
changeset | 810 |   | abstrs => error ("Cannot export abstract constructor(s): "
 | 
| 
d191892b1c23
explicit check that abstract constructors cannot be part of official interface
 haftmann parents: 
63164diff
changeset | 811 | ^ commas (map (Code.string_of_const thy) abstrs)); | 
| 
d191892b1c23
explicit check that abstract constructors cannot be part of official interface
 haftmann parents: 
63164diff
changeset | 812 | |
| 63164 | 813 | fun invoke_generation_for_consts ctxt { ignore_cache, permissive } { algebra, eqngr } consts =
 | 
| 63175 
d191892b1c23
explicit check that abstract constructors cannot be part of official interface
 haftmann parents: 
63164diff
changeset | 814 | let | 
| 
d191892b1c23
explicit check that abstract constructors cannot be part of official interface
 haftmann parents: 
63164diff
changeset | 815 | val thy = Proof_Context.theory_of ctxt; | 
| 
d191892b1c23
explicit check that abstract constructors cannot be part of official interface
 haftmann parents: 
63164diff
changeset | 816 | val _ = if permissive then () | 
| 
d191892b1c23
explicit check that abstract constructors cannot be part of official interface
 haftmann parents: 
63164diff
changeset | 817 | else check_abstract_constructors thy consts; | 
| 
d191892b1c23
explicit check that abstract constructors cannot be part of official interface
 haftmann parents: 
63164diff
changeset | 818 | in | 
| 
d191892b1c23
explicit check that abstract constructors cannot be part of official interface
 haftmann parents: 
63164diff
changeset | 819 | Code_Preproc.timed "translating program" #ctxt | 
| 
d191892b1c23
explicit check that abstract constructors cannot be part of official interface
 haftmann parents: 
63164diff
changeset | 820 |     (fn { ctxt, algebra, eqngr, consts } => invoke_generation ignore_cache ctxt
 | 
| 
d191892b1c23
explicit check that abstract constructors cannot be part of official interface
 haftmann parents: 
63164diff
changeset | 821 | (fold_map (ensure_const ctxt algebra eqngr permissive)) consts) | 
| 
d191892b1c23
explicit check that abstract constructors cannot be part of official interface
 haftmann parents: 
63164diff
changeset | 822 |       { ctxt = ctxt, algebra = algebra, eqngr = eqngr, consts = consts }
 | 
| 
d191892b1c23
explicit check that abstract constructors cannot be part of official interface
 haftmann parents: 
63164diff
changeset | 823 | end; | 
| 63160 | 824 | |
| 825 | fun invoke_generation_for_consts' ctxt ignore_cache_and_permissive consts = | |
| 826 | invoke_generation_for_consts ctxt | |
| 827 |     { ignore_cache = ignore_cache_and_permissive, permissive = ignore_cache_and_permissive }
 | |
| 63164 | 828 | (Code_Preproc.obtain ignore_cache_and_permissive | 
| 829 |       { ctxt = ctxt, consts = consts, terms = []}) consts
 | |
| 63160 | 830 | |> snd; | 
| 55188 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55150diff
changeset | 831 | |
| 63160 | 832 | fun invoke_generation_for_consts'' ctxt algebra_eqngr = | 
| 833 | invoke_generation_for_consts ctxt | |
| 834 |     { ignore_cache = true, permissive = false }
 | |
| 835 | algebra_eqngr | |
| 836 |   #> (fn (deps, program) => { deps = deps, program = program });
 | |
| 55188 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55150diff
changeset | 837 | |
| 63160 | 838 | fun consts_program_permissive ctxt = | 
| 839 | invoke_generation_for_consts' ctxt true; | |
| 840 | ||
| 841 | fun consts_program ctxt consts = | |
| 55188 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55150diff
changeset | 842 | let | 
| 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55150diff
changeset | 843 | fun project program = Code_Symbol.Graph.restrict | 
| 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55150diff
changeset | 844 | (member (op =) (Code_Symbol.Graph.all_succs program | 
| 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55150diff
changeset | 845 | (map Constant consts))) program; | 
| 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55150diff
changeset | 846 | in | 
| 63160 | 847 | invoke_generation_for_consts' ctxt false consts | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 848 | |> project | 
| 27103 | 849 | end; | 
| 850 | ||
| 851 | ||
| 852 | (* value evaluation *) | |
| 25969 | 853 | |
| 56969 
7491932da574
dropped obsolete hand-waving adjustion of type variables: safely done in preprocessor
 haftmann parents: 
56920diff
changeset | 854 | fun ensure_value ctxt algbr eqngr t = | 
| 24918 | 855 | let | 
| 856 | val ty = fastype_of t; | |
| 56969 
7491932da574
dropped obsolete hand-waving adjustion of type variables: safely done in preprocessor
 haftmann parents: 
56920diff
changeset | 857 | val vs = fold_term_types (K (fold_atyps (insert (eq_fst op =) | 
| 24918 | 858 | o dest_TFree))) t []; | 
| 69593 | 859 | val t' = annotate ctxt algbr eqngr (\<^const_name>\<open>Pure.dummy_pattern\<close>, ty) [] t; | 
| 860 | val dummy_constant = Constant \<^const_name>\<open>Pure.dummy_pattern\<close>; | |
| 24918 | 861 | val stmt_value = | 
| 55757 | 862 | fold_map (translate_tyvar_sort ctxt algbr eqngr false) vs | 
| 863 | ##>> translate_typ ctxt algbr eqngr false ty | |
| 864 | ##>> translate_term ctxt algbr eqngr false NONE (t', NONE) | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 865 | #>> (fn ((vs, ty), t) => Fun | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 866 | (((vs, ty), [(([], t), (NONE, true))]), NONE)); | 
| 56920 
d651b944c67e
normalizing of type variables before evaluation with explicit resubstitution function: make nbe work with funny type variables like \<AA>;
 haftmann parents: 
56811diff
changeset | 867 | fun term_value (_, program1) = | 
| 25969 | 868 | let | 
| 56969 
7491932da574
dropped obsolete hand-waving adjustion of type variables: safely done in preprocessor
 haftmann parents: 
56920diff
changeset | 869 | val Fun ((vs_ty, [(([], t), _)]), _) = | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 870 | Code_Symbol.Graph.get_node program1 dummy_constant; | 
| 55190 | 871 | val deps' = Code_Symbol.Graph.immediate_succs program1 dummy_constant; | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 872 | val program2 = Code_Symbol.Graph.del_node dummy_constant program1; | 
| 55190 | 873 | val deps_all = Code_Symbol.Graph.all_succs program2 deps'; | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54932diff
changeset | 874 | val program3 = Code_Symbol.Graph.restrict (member (op =) deps_all) program2; | 
| 56969 
7491932da574
dropped obsolete hand-waving adjustion of type variables: safely done in preprocessor
 haftmann parents: 
56920diff
changeset | 875 | in ((program3, ((vs_ty, t), deps')), (deps', program2)) end; | 
| 26011 | 876 | in | 
| 69593 | 877 | ensure_stmt Constant stmt_value \<^const_name>\<open>Pure.dummy_pattern\<close> | 
| 26011 | 878 | #> snd | 
| 31063 
88aaab83b6fc
dropped explicit suppport for frees in evaluation conversion stack
 haftmann parents: 
31054diff
changeset | 879 | #> term_value | 
| 26011 | 880 | end; | 
| 24219 | 881 | |
| 64957 | 882 | fun dynamic_evaluation comp ctxt algebra eqngr t = | 
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 883 | let | 
| 56969 
7491932da574
dropped obsolete hand-waving adjustion of type variables: safely done in preprocessor
 haftmann parents: 
56920diff
changeset | 884 | val ((program, (vs_ty_t', deps)), _) = | 
| 63164 | 885 | Code_Preproc.timed "translating term" #ctxt | 
| 886 |       (fn { ctxt, algebra, eqngr, t } =>
 | |
| 887 | invoke_generation false ctxt (ensure_value ctxt algebra eqngr) t) | |
| 888 |         { ctxt = ctxt, algebra = algebra, eqngr = eqngr, t = t };
 | |
| 64957 | 889 | in comp program t vs_ty_t' deps end; | 
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 890 | |
| 55757 | 891 | fun dynamic_conv ctxt conv = | 
| 63157 
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
 haftmann parents: 
63156diff
changeset | 892 | Code_Preproc.dynamic_conv ctxt | 
| 
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
 haftmann parents: 
63156diff
changeset | 893 | (dynamic_evaluation (fn program => fn _ => conv program) ctxt); | 
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 894 | |
| 63157 
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
 haftmann parents: 
63156diff
changeset | 895 | fun dynamic_value ctxt postproc comp = | 
| 
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
 haftmann parents: 
63156diff
changeset | 896 | Code_Preproc.dynamic_value ctxt postproc | 
| 
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
 haftmann parents: 
63156diff
changeset | 897 | (dynamic_evaluation comp ctxt); | 
| 41365 | 898 | |
| 63160 | 899 | fun static_evaluation ctxt consts algebra_eqngr static_eval = | 
| 900 | static_eval (invoke_generation_for_consts'' ctxt algebra_eqngr consts); | |
| 901 | ||
| 902 | fun static_evaluation_thingol ctxt consts (algebra_eqngr as { algebra, eqngr }) static_eval =
 | |
| 41365 | 903 | let | 
| 63160 | 904 | fun evaluation program dynamic_eval ctxt t = | 
| 63157 
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
 haftmann parents: 
63156diff
changeset | 905 | let | 
| 
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
 haftmann parents: 
63156diff
changeset | 906 | val ((_, ((vs_ty', t'), deps)), _) = | 
| 63164 | 907 | Code_Preproc.timed "translating term" #ctxt | 
| 908 |           (fn { ctxt, t } =>
 | |
| 909 | ensure_value ctxt algebra eqngr t ([], program)) | |
| 910 |             { ctxt = ctxt, t = t };
 | |
| 63157 
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
 haftmann parents: 
63156diff
changeset | 911 | in dynamic_eval ctxt t (vs_ty', t') deps end; | 
| 
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
 haftmann parents: 
63156diff
changeset | 912 | in | 
| 63160 | 913 | static_evaluation ctxt consts algebra_eqngr (fn program_deps => | 
| 914 | evaluation (#program program_deps) (static_eval program_deps)) | |
| 63157 
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
 haftmann parents: 
63156diff
changeset | 915 | end; | 
| 
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
 haftmann parents: 
63156diff
changeset | 916 | |
| 63160 | 917 | fun static_evaluation_isa ctxt consts algebra_eqngr static_eval = | 
| 918 | static_evaluation ctxt consts algebra_eqngr (fn program_deps => | |
| 919 | (static_eval (#program program_deps))); | |
| 39487 
80b2cf3b0779
proper closures for static evaluation; no need for FIXMEs any longer
 haftmann parents: 
39475diff
changeset | 920 | |
| 63156 | 921 | fun static_conv_thingol (ctxt_consts as { ctxt, consts }) conv =
 | 
| 63160 | 922 | Code_Preproc.static_conv ctxt_consts (fn algebra_eqngr => | 
| 923 | static_evaluation_thingol ctxt consts algebra_eqngr | |
| 924 | (fn program_deps => | |
| 925 | let | |
| 926 | val static_conv = conv program_deps; | |
| 927 | in | |
| 928 | fn ctxt => fn _ => fn vs_ty => fn deps => static_conv ctxt vs_ty deps | |
| 929 | end)); | |
| 38672 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 haftmann parents: 
38669diff
changeset | 930 | |
| 63156 | 931 | fun static_conv_isa (ctxt_consts as { ctxt, consts }) conv =
 | 
| 63160 | 932 | Code_Preproc.static_conv ctxt_consts (fn algebra_eqngr => | 
| 933 | static_evaluation_isa ctxt consts algebra_eqngr conv); | |
| 38672 
f1f64375f662
preliminary versions of static_eval_conv(_simple)
 haftmann parents: 
38669diff
changeset | 934 | |
| 63157 
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
 haftmann parents: 
63156diff
changeset | 935 | fun static_value (ctxt_postproc_consts as { ctxt, consts, ... }) comp =
 | 
| 63160 | 936 | Code_Preproc.static_value ctxt_postproc_consts (fn algebra_eqngr => | 
| 937 | static_evaluation_thingol ctxt consts algebra_eqngr comp); | |
| 39487 
80b2cf3b0779
proper closures for static evaluation; no need for FIXMEs any longer
 haftmann parents: 
39475diff
changeset | 938 | |
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 939 | |
| 55188 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55150diff
changeset | 940 | (** constant expressions **) | 
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 941 | |
| 55757 | 942 | fun read_const_exprs_internal ctxt = | 
| 31036 | 943 | let | 
| 55757 | 944 | val thy = Proof_Context.theory_of ctxt; | 
| 68482 | 945 | fun this_theory name = | 
| 946 | if Context.theory_name thy = name then thy | |
| 947 |       else Context.get_theory {long = false} thy name;
 | |
| 948 | ||
| 56062 | 949 | fun consts_of thy' = | 
| 950 | fold (fn (c, (_, NONE)) => cons c | _ => I) | |
| 63175 
d191892b1c23
explicit check that abstract constructors cannot be part of official interface
 haftmann parents: 
63164diff
changeset | 951 | (#constants (Consts.dest (Sign.consts_of thy'))) [] | 
| 
d191892b1c23
explicit check that abstract constructors cannot be part of official interface
 haftmann parents: 
63164diff
changeset | 952 | |> filter_out (Code.is_abstr thy); | 
| 36272 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 953 | fun belongs_here thy' c = forall | 
| 
4d358c582ffb
optionally ignore errors during translation of equations; tuned representation of abstraction points
 haftmann parents: 
36210diff
changeset | 954 | (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 | 955 | fun consts_of_select thy' = filter (belongs_here thy') (consts_of thy'); | 
| 52801 | 956 | fun read_const_expr str = | 
| 59795 | 957 | (case Syntax.parse_input ctxt (K NONE) (K Markup.empty) (SOME o Symbol_Pos.implode o #1) str of | 
| 52801 | 958 | SOME "_" => ([], consts_of thy) | 
| 959 | | SOME s => | |
| 960 | if String.isSuffix "._" s | |
| 68482 | 961 | then ([], consts_of_select (this_theory (unsuffix "._" s))) | 
| 52801 | 962 | else ([Code.read_const thy str], []) | 
| 963 | | NONE => ([Code.read_const thy str], [])); | |
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58893diff
changeset | 964 | in apply2 flat o split_list o map read_const_expr end; | 
| 31036 | 965 | |
| 55757 | 966 | fun read_const_exprs_all ctxt = op @ o read_const_exprs_internal ctxt; | 
| 55188 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55150diff
changeset | 967 | |
| 55757 | 968 | fun read_const_exprs ctxt const_exprs = | 
| 55188 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55150diff
changeset | 969 | let | 
| 63159 | 970 | val (consts, consts_permissive) = | 
| 971 | read_const_exprs_internal ctxt const_exprs; | |
| 63175 
d191892b1c23
explicit check that abstract constructors cannot be part of official interface
 haftmann parents: 
63164diff
changeset | 972 | val consts' = | 
| 
d191892b1c23
explicit check that abstract constructors cannot be part of official interface
 haftmann parents: 
63164diff
changeset | 973 | consts_program_permissive ctxt consts_permissive | 
| 
d191892b1c23
explicit check that abstract constructors cannot be part of official interface
 haftmann parents: 
63164diff
changeset | 974 | |> implemented_deps | 
| 
d191892b1c23
explicit check that abstract constructors cannot be part of official interface
 haftmann parents: 
63164diff
changeset | 975 | |> filter_out (Code.is_abstr (Proof_Context.theory_of ctxt)); | 
| 55188 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55150diff
changeset | 976 | in union (op =) consts' consts end; | 
| 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55150diff
changeset | 977 | |
| 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55150diff
changeset | 978 | |
| 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55150diff
changeset | 979 | (** diagnostic commands **) | 
| 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55150diff
changeset | 980 | |
| 55757 | 981 | fun code_depgr ctxt consts = | 
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 982 | let | 
| 63164 | 983 |     val { eqngr, ... } = Code_Preproc.obtain true
 | 
| 984 |       { ctxt = ctxt, consts = consts, terms = [] };
 | |
| 34173 
458ced35abb8
reduced code generator cache to the baremost minimum
 haftmann parents: 
34084diff
changeset | 985 | 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 | 986 | 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 | 987 | |
| 55757 | 988 | fun code_thms ctxt = Pretty.writeln o Code_Preproc.pretty ctxt o code_depgr ctxt; | 
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 989 | |
| 60203 
add72fdadd0b
filtering of reflexive dependencies avoids problems with state-of-the-art graph browser;
 wenzelm parents: 
60097diff
changeset | 990 | fun coalesce_strong_conn gr = | 
| 59208 
2486d625878b
for graph display, prefer graph data structure over list with dependencies;
 wenzelm parents: 
59207diff
changeset | 991 | let | 
| 
2486d625878b
for graph display, prefer graph data structure over list with dependencies;
 wenzelm parents: 
59207diff
changeset | 992 | val xss = Graph.strong_conn gr; | 
| 60203 
add72fdadd0b
filtering of reflexive dependencies avoids problems with state-of-the-art graph browser;
 wenzelm parents: 
60097diff
changeset | 993 | val xss_ys = map (fn xs => (xs, commas xs)) xss; | 
| 
add72fdadd0b
filtering of reflexive dependencies avoids problems with state-of-the-art graph browser;
 wenzelm parents: 
60097diff
changeset | 994 | val y_for = the o AList.lookup (op =) (maps (fn (xs, y) => map (fn x => (x, y)) xs) xss_ys); | 
| 
add72fdadd0b
filtering of reflexive dependencies avoids problems with state-of-the-art graph browser;
 wenzelm parents: 
60097diff
changeset | 995 | fun coalesced_succs_for xs = maps (Graph.immediate_succs gr) xs | 
| 
add72fdadd0b
filtering of reflexive dependencies avoids problems with state-of-the-art graph browser;
 wenzelm parents: 
60097diff
changeset | 996 | |> subtract (op =) xs | 
| 
add72fdadd0b
filtering of reflexive dependencies avoids problems with state-of-the-art graph browser;
 wenzelm parents: 
60097diff
changeset | 997 | |> map y_for | 
| 
add72fdadd0b
filtering of reflexive dependencies avoids problems with state-of-the-art graph browser;
 wenzelm parents: 
60097diff
changeset | 998 | |> distinct (op =); | 
| 
add72fdadd0b
filtering of reflexive dependencies avoids problems with state-of-the-art graph browser;
 wenzelm parents: 
60097diff
changeset | 999 | val succs = map (fn (xs, _) => (xs, coalesced_succs_for xs)) xss_ys; | 
| 59208 
2486d625878b
for graph display, prefer graph data structure over list with dependencies;
 wenzelm parents: 
59207diff
changeset | 1000 | in | 
| 60204 
137b3fc46bb3
code equations as displayable content in code dependency graph
 wenzelm parents: 
60203diff
changeset | 1001 | map (fn (xs, y) => ((y, xs), (maps (Graph.get_node gr) xs, (the o AList.lookup (op =) succs) xs))) xss_ys | 
| 59208 
2486d625878b
for graph display, prefer graph data structure over list with dependencies;
 wenzelm parents: 
59207diff
changeset | 1002 | end; | 
| 
2486d625878b
for graph display, prefer graph data structure over list with dependencies;
 wenzelm parents: 
59207diff
changeset | 1003 | |
| 55757 | 1004 | fun code_deps ctxt consts = | 
| 27103 | 1005 | let | 
| 55757 | 1006 | val thy = Proof_Context.theory_of ctxt; | 
| 60204 
137b3fc46bb3
code equations as displayable content in code dependency graph
 wenzelm parents: 
60203diff
changeset | 1007 | fun mk_entry ((name, consts), (ps, deps)) = | 
| 
137b3fc46bb3
code equations as displayable content in code dependency graph
 wenzelm parents: 
60203diff
changeset | 1008 | let | 
| 
137b3fc46bb3
code equations as displayable content in code dependency graph
 wenzelm parents: 
60203diff
changeset | 1009 | val label = commas (map (Code.string_of_const thy) consts); | 
| 
137b3fc46bb3
code equations as displayable content in code dependency graph
 wenzelm parents: 
60203diff
changeset | 1010 | in ((name, Graph_Display.content_node label (Pretty.str label :: ps)), deps) end; | 
| 59210 
8658b4290aed
clarified Graph_Display.graph etc.: sort_graph determines order from structure (and names);
 wenzelm parents: 
59208diff
changeset | 1011 | in | 
| 
8658b4290aed
clarified Graph_Display.graph etc.: sort_graph determines order from structure (and names);
 wenzelm parents: 
59208diff
changeset | 1012 | code_depgr ctxt consts | 
| 60204 
137b3fc46bb3
code equations as displayable content in code dependency graph
 wenzelm parents: 
60203diff
changeset | 1013 | |> Graph.map (K (Code.pretty_cert thy o snd)) | 
| 60203 
add72fdadd0b
filtering of reflexive dependencies avoids problems with state-of-the-art graph browser;
 wenzelm parents: 
60097diff
changeset | 1014 | |> coalesce_strong_conn | 
| 60204 
137b3fc46bb3
code equations as displayable content in code dependency graph
 wenzelm parents: 
60203diff
changeset | 1015 | |> map mk_entry | 
| 60203 
add72fdadd0b
filtering of reflexive dependencies avoids problems with state-of-the-art graph browser;
 wenzelm parents: 
60097diff
changeset | 1016 | |> Graph_Display.display_graph | 
| 59210 
8658b4290aed
clarified Graph_Display.graph etc.: sort_graph determines order from structure (and names);
 wenzelm parents: 
59208diff
changeset | 1017 | end; | 
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1018 | |
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1019 | local | 
| 27103 | 1020 | |
| 55757 | 1021 | fun code_thms_cmd ctxt = code_thms ctxt o read_const_exprs_all ctxt; | 
| 1022 | fun code_deps_cmd ctxt = code_deps ctxt o read_const_exprs_all ctxt; | |
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1023 | |
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1024 | in | 
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1025 | |
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1026 | val _ = | 
| 69593 | 1027 | Outer_Syntax.command \<^command_keyword>\<open>code_thms\<close> | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46665diff
changeset | 1028 | "print system of code equations for code" | 
| 52801 | 1029 | (Scan.repeat1 Parse.term >> (fn cs => | 
| 60097 
d20ca79d50e4
discontinued pointless warnings: commands are only defined inside a theory context;
 wenzelm parents: 
60089diff
changeset | 1030 | Toplevel.keep (fn st => code_thms_cmd (Toplevel.context_of st) cs))); | 
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1031 | |
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1032 | val _ = | 
| 69593 | 1033 | Outer_Syntax.command \<^command_keyword>\<open>code_deps\<close> | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46665diff
changeset | 1034 | "visualize dependencies of code equations for code" | 
| 52801 | 1035 | (Scan.repeat1 Parse.term >> (fn cs => | 
| 59210 
8658b4290aed
clarified Graph_Display.graph etc.: sort_graph determines order from structure (and names);
 wenzelm parents: 
59208diff
changeset | 1036 | Toplevel.keep (fn st => code_deps_cmd (Toplevel.context_of st) cs))); | 
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1037 | |
| 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30932diff
changeset | 1038 | end; | 
| 27103 | 1039 | |
| 24219 | 1040 | end; (*struct*) | 
| 1041 | ||
| 1042 | ||
| 28054 | 1043 | structure Basic_Code_Thingol: BASIC_CODE_THINGOL = Code_Thingol; |