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