src/Tools/Code/code_preproc.ML
author wenzelm
Sun, 18 Mar 2012 13:04:22 +0100
changeset 47005 421760a1efe7
parent 46961 5c6955f487e5
child 48075 ec5e62b868eb
permissions -rw-r--r--
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming); more explicit Context.generic for Name_Space.declare/define and derivatives (NB: naming changed after Proof_Context.init_global); prefer Context.pretty in low-level operations of structure Sorts/Type (defer full Syntax.init_pretty until error output); simplified signatures;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37744
3daaf23b9ab4 tuned titles
haftmann
parents: 37442
diff changeset
     1
(*  Title:      Tools/Code/code_preproc.ML
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
     2
    Author:     Florian Haftmann, TU Muenchen
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
     3
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
     4
Preprocessing code equations into a well-sorted system
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
     5
in a graph with explicit dependencies.
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
     6
*)
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
     7
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
     8
signature CODE_PREPROC =
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
     9
sig
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    10
  val map_pre: (simpset -> simpset) -> theory -> theory
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    11
  val map_post: (simpset -> simpset) -> theory -> theory
32072
d4bff63bcbf1 added code_unfold_post attribute
haftmann
parents: 31998
diff changeset
    12
  val add_unfold: thm -> theory -> theory
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    13
  val add_functrans: string * (theory -> (thm * bool) list -> (thm * bool) list option) -> theory -> theory
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    14
  val del_functrans: string -> theory -> theory
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    15
  val simple_functrans: (theory -> thm list -> thm list option)
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    16
    -> theory -> (thm * bool) list -> (thm * bool) list option
34893
ecdc526af73a function transformer preprocessor applies to both code generators
haftmann
parents: 34891
diff changeset
    17
  val preprocess_functrans: theory -> (thm * bool) list -> (thm * bool) list
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    18
  val print_codeproc: theory -> unit
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    19
30947
dd551284a300 re-engineering of evaluation conversions
haftmann
parents: 30942
diff changeset
    20
  type code_algebra
dd551284a300 re-engineering of evaluation conversions
haftmann
parents: 30942
diff changeset
    21
  type code_graph
34891
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34251
diff changeset
    22
  val cert: code_graph -> string -> Code.cert
32873
333945c9ac6a tuned handling of type variable names further
haftmann
parents: 32872
diff changeset
    23
  val sortargs: code_graph -> string -> sort list
30947
dd551284a300 re-engineering of evaluation conversions
haftmann
parents: 30942
diff changeset
    24
  val all: code_graph -> string list
dd551284a300 re-engineering of evaluation conversions
haftmann
parents: 30942
diff changeset
    25
  val pretty: theory -> code_graph -> Pretty.T
39398
2e30660a2e21 ignore code cache optionally; corrected scope of term value in static_eval_conv
haftmann
parents: 39133
diff changeset
    26
  val obtain: bool -> theory -> string list -> term list -> code_algebra * code_graph
41184
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 40176
diff changeset
    27
  val dynamic_conv: theory
38670
3c7db0192db9 added static_eval_conv
haftmann
parents: 38669
diff changeset
    28
    -> (code_algebra -> code_graph -> (string * sort) list -> term -> conv) -> conv
41184
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 40176
diff changeset
    29
  val dynamic_value: theory -> ((term -> term) -> 'a -> 'a)
30947
dd551284a300 re-engineering of evaluation conversions
haftmann
parents: 30942
diff changeset
    30
    -> (code_algebra -> code_graph -> (string * sort) list -> term -> 'a) -> term -> 'a
41184
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 40176
diff changeset
    31
  val static_conv: theory -> string list
41346
6673f6fa94ca canonical handling of theory context argument
haftmann
parents: 41226
diff changeset
    32
    -> (code_algebra -> code_graph -> (string * sort) list -> term -> conv) -> conv
41184
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 40176
diff changeset
    33
  val static_value: theory -> ((term -> term) -> 'a -> 'a) -> string list
41346
6673f6fa94ca canonical handling of theory context argument
haftmann
parents: 41226
diff changeset
    34
    -> (code_algebra -> code_graph -> (string * sort) list -> term -> 'a) -> term -> 'a
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    35
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    36
  val setup: theory -> theory
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
    37
end
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
    38
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    39
structure Code_Preproc : CODE_PREPROC =
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
    40
struct
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
    41
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    42
(** preprocessor administration **)
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    43
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    44
(* theory data *)
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    45
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    46
datatype thmproc = Thmproc of {
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    47
  pre: simpset,
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    48
  post: simpset,
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    49
  functrans: (string * (serial * (theory -> (thm * bool) list -> (thm * bool) list option))) list
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    50
};
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    51
31599
97b4d289c646 tuned make/map/merge combinators
haftmann
parents: 31156
diff changeset
    52
fun make_thmproc ((pre, post), functrans) =
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    53
  Thmproc { pre = pre, post = post, functrans = functrans };
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    54
fun map_thmproc f (Thmproc { pre, post, functrans }) =
31599
97b4d289c646 tuned make/map/merge combinators
haftmann
parents: 31156
diff changeset
    55
  make_thmproc (f ((pre, post), functrans));
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    56
fun merge_thmproc (Thmproc { pre = pre1, post = post1, functrans = functrans1 },
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    57
  Thmproc { pre = pre2, post = post2, functrans = functrans2 }) =
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    58
    let
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    59
      val pre = Simplifier.merge_ss (pre1, pre2);
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    60
      val post = Simplifier.merge_ss (post1, post2);
34891
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34251
diff changeset
    61
      val functrans = AList.merge (op =) (eq_fst (op =)) (functrans1, functrans2)
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34251
diff changeset
    62
        handle AList.DUP => error ("Duplicate function transformer");
31599
97b4d289c646 tuned make/map/merge combinators
haftmann
parents: 31156
diff changeset
    63
    in make_thmproc ((pre, post), functrans) end;
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    64
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33063
diff changeset
    65
structure Code_Preproc_Data = Theory_Data
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    66
(
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    67
  type T = thmproc;
31599
97b4d289c646 tuned make/map/merge combinators
haftmann
parents: 31156
diff changeset
    68
  val empty = make_thmproc ((Simplifier.empty_ss, Simplifier.empty_ss), []);
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33063
diff changeset
    69
  val extend = I;
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33063
diff changeset
    70
  val merge = merge_thmproc;
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    71
);
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    72
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    73
fun the_thmproc thy = case Code_Preproc_Data.get thy
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    74
 of Thmproc x => x;
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    75
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    76
fun delete_force msg key xs =
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    77
  if AList.defined (op =) xs key then AList.delete (op =) key xs
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    78
  else error ("No such " ^ msg ^ ": " ^ quote key);
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    79
38669
9ff76d0f0610 refined and unified naming convention for dynamic code evaluation techniques
haftmann
parents: 38291
diff changeset
    80
val map_data = Code_Preproc_Data.map o map_thmproc;
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    81
32072
d4bff63bcbf1 added code_unfold_post attribute
haftmann
parents: 31998
diff changeset
    82
val map_pre_post = map_data o apfst;
d4bff63bcbf1 added code_unfold_post attribute
haftmann
parents: 31998
diff changeset
    83
val map_pre = map_pre_post o apfst;
d4bff63bcbf1 added code_unfold_post attribute
haftmann
parents: 31998
diff changeset
    84
val map_post = map_pre_post o apsnd;
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
    85
41226
adcb9a1198e7 clarified exports of structure Simplifier;
wenzelm
parents: 41189
diff changeset
    86
val add_unfold = map_pre o Simplifier.add_simp;
adcb9a1198e7 clarified exports of structure Simplifier;
wenzelm
parents: 41189
diff changeset
    87
val del_unfold = map_pre o Simplifier.del_simp;
adcb9a1198e7 clarified exports of structure Simplifier;
wenzelm
parents: 41189
diff changeset
    88
val add_post = map_post o Simplifier.add_simp;
adcb9a1198e7 clarified exports of structure Simplifier;
wenzelm
parents: 41189
diff changeset
    89
val del_post = map_post o Simplifier.del_simp;
32072
d4bff63bcbf1 added code_unfold_post attribute
haftmann
parents: 31998
diff changeset
    90
46026
83caa4f4bd56 semiring_numeral_0_eq_0, semiring_numeral_1_eq_1 now [simp], superseeding corresponding simp rules on type nat
haftmann
parents: 45230
diff changeset
    91
fun add_code_abbrev raw_thm thy =
32072
d4bff63bcbf1 added code_unfold_post attribute
haftmann
parents: 31998
diff changeset
    92
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41346
diff changeset
    93
    val thm = Local_Defs.meta_rewrite_rule (Proof_Context.init_global thy) raw_thm;
32072
d4bff63bcbf1 added code_unfold_post attribute
haftmann
parents: 31998
diff changeset
    94
    val thm_sym = Thm.symmetric thm;
d4bff63bcbf1 added code_unfold_post attribute
haftmann
parents: 31998
diff changeset
    95
  in
d4bff63bcbf1 added code_unfold_post attribute
haftmann
parents: 31998
diff changeset
    96
    thy |> map_pre_post (fn (pre, post) =>
46026
83caa4f4bd56 semiring_numeral_0_eq_0, semiring_numeral_1_eq_1 now [simp], superseeding corresponding simp rules on type nat
haftmann
parents: 45230
diff changeset
    97
      (pre |> Simplifier.add_simp thm_sym, post |> Simplifier.add_simp thm))
32072
d4bff63bcbf1 added code_unfold_post attribute
haftmann
parents: 31998
diff changeset
    98
  end;
d4bff63bcbf1 added code_unfold_post attribute
haftmann
parents: 31998
diff changeset
    99
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   100
fun add_functrans (name, f) = (map_data o apsnd)
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   101
  (AList.update (op =) (name, (serial (), f)));
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   102
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   103
fun del_functrans name = (map_data o apsnd)
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   104
  (delete_force "function transformer" name);
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   105
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   106
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   107
(* post- and preprocessing *)
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   108
39604
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   109
fun no_variables_conv conv ct =
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   110
  let
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   111
    val cert = Thm.cterm_of (Thm.theory_of_cterm ct);
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   112
    val all_vars = fold_aterms (fn t as Free _ => insert (op aconvc) (cert t)
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   113
      | t as Var _ => insert (op aconvc) (cert t)
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   114
      | _ => I) (Thm.term_of ct) [];
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   115
    fun apply_beta var thm = Thm.combination thm (Thm.reflexive var)
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   116
      |> Conv.fconv_rule (Conv.arg_conv (Conv.try_conv (Thm.beta_conversion false)))
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   117
      |> Conv.fconv_rule (Conv.arg1_conv (Thm.beta_conversion false));
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   118
  in
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   119
    ct
46497
89ccf66aa73d renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
wenzelm
parents: 46026
diff changeset
   120
    |> fold_rev Thm.lambda all_vars
39604
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   121
    |> conv
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   122
    |> fold apply_beta all_vars
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   123
  end;
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   124
33942
haftmann
parents: 33699
diff changeset
   125
fun trans_conv_rule conv thm = Thm.transitive thm ((conv o Thm.rhs_of) thm);
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   126
34891
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34251
diff changeset
   127
fun eqn_conv conv ct =
31957
a9742afd403e tuned interface of structure Code
haftmann
parents: 31775
diff changeset
   128
  let
a9742afd403e tuned interface of structure Code
haftmann
parents: 31775
diff changeset
   129
    fun lhs_conv ct = if can Thm.dest_comb ct
a9742afd403e tuned interface of structure Code
haftmann
parents: 31775
diff changeset
   130
      then Conv.combination_conv lhs_conv conv ct
a9742afd403e tuned interface of structure Code
haftmann
parents: 31775
diff changeset
   131
      else Conv.all_conv ct;
34891
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34251
diff changeset
   132
  in Conv.combination_conv (Conv.arg_conv lhs_conv) conv ct end;
31957
a9742afd403e tuned interface of structure Code
haftmann
parents: 31775
diff changeset
   133
a9742afd403e tuned interface of structure Code
haftmann
parents: 31775
diff changeset
   134
val rewrite_eqn = Conv.fconv_rule o eqn_conv o Simplifier.rewrite;
a9742afd403e tuned interface of structure Code
haftmann
parents: 31775
diff changeset
   135
41189
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   136
fun term_of_conv thy conv =
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   137
  Thm.cterm_of thy
41189
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   138
  #> conv
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   139
  #> Thm.prop_of
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   140
  #> Logic.dest_equals
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   141
  #> snd;
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   142
41189
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   143
fun term_of_conv_resubst thy conv t =
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   144
  let
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   145
    val all_vars = fold_aterms (fn t as Free _ => insert (op aconv) t
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   146
      | t as Var _ => insert (op aconv) t
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   147
      | _ => I) t [];
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   148
    val resubst = curry (Term.betapplys o swap) all_vars;
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   149
  in (resubst, term_of_conv thy conv (fold_rev lambda all_vars t)) end;
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   150
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   151
34893
ecdc526af73a function transformer preprocessor applies to both code generators
haftmann
parents: 34891
diff changeset
   152
fun preprocess_functrans thy = 
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   153
  let
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   154
    val functrans = (map (fn (_, (_, f)) => f thy) o #functrans
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   155
      o the_thmproc) thy;
34893
ecdc526af73a function transformer preprocessor applies to both code generators
haftmann
parents: 34891
diff changeset
   156
  in perhaps (perhaps_loop (perhaps_apply functrans)) end;
ecdc526af73a function transformer preprocessor applies to both code generators
haftmann
parents: 34891
diff changeset
   157
ecdc526af73a function transformer preprocessor applies to both code generators
haftmann
parents: 34891
diff changeset
   158
fun preprocess thy =
ecdc526af73a function transformer preprocessor applies to both code generators
haftmann
parents: 34891
diff changeset
   159
  let
43122
027ed67f5d98 code preprocessor applies simplifier when schematic variables are fixed to free variables to allow rewriting with congruence rules in the preprocessing steps
bulwahn
parents: 42387
diff changeset
   160
    val ctxt = Proof_Context.init_global thy;
35232
f588e1169c8b renamed Simplifier.theory_context to Simplifier.global_context to emphasize that this is not the real thing;
wenzelm
parents: 34895
diff changeset
   161
    val pre = (Simplifier.global_context thy o #pre o the_thmproc) thy;
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   162
  in
34893
ecdc526af73a function transformer preprocessor applies to both code generators
haftmann
parents: 34891
diff changeset
   163
    preprocess_functrans thy
43122
027ed67f5d98 code preprocessor applies simplifier when schematic variables are fixed to free variables to allow rewriting with congruence rules in the preprocessing steps
bulwahn
parents: 42387
diff changeset
   164
    #> (map o apfst) (singleton (Variable.trade (K (map (rewrite_eqn pre))) ctxt)) 
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   165
  end;
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   166
41189
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   167
fun preprocess_conv thy =
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   168
  let
35232
f588e1169c8b renamed Simplifier.theory_context to Simplifier.global_context to emphasize that this is not the real thing;
wenzelm
parents: 34895
diff changeset
   169
    val pre = (Simplifier.global_context thy o #pre o the_thmproc) thy;
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   170
  in
41189
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   171
    Simplifier.rewrite pre
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   172
    #> trans_conv_rule (AxClass.unoverload_conv thy)
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   173
  end;
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   174
41189
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   175
fun preprocess_term thy = term_of_conv_resubst thy (preprocess_conv thy);
39475
9cc1ba3c5706 separation of static and dynamic thy context
haftmann
parents: 39398
diff changeset
   176
41189
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   177
fun postprocess_conv thy =
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   178
  let
35232
f588e1169c8b renamed Simplifier.theory_context to Simplifier.global_context to emphasize that this is not the real thing;
wenzelm
parents: 34895
diff changeset
   179
    val post = (Simplifier.global_context thy o #post o the_thmproc) thy;
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   180
  in
41189
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   181
    AxClass.overload_conv thy
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   182
    #> trans_conv_rule (Simplifier.rewrite post)
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   183
  end;
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   184
32544
e129333b9df0 moved eq handling in nbe into separate oracle
haftmann
parents: 32353
diff changeset
   185
fun postprocess_term thy = term_of_conv thy (postprocess_conv thy);
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   186
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   187
fun print_codeproc thy =
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   188
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41346
diff changeset
   189
    val ctxt = Proof_Context.init_global thy;
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   190
    val pre = (#pre o the_thmproc) thy;
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   191
    val post = (#post o the_thmproc) thy;
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   192
    val functrans = (map fst o #functrans o the_thmproc) thy;
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   193
  in
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   194
    (Pretty.writeln o Pretty.chunks) [
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   195
      Pretty.block [
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   196
        Pretty.str "preprocessing simpset:",
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   197
        Pretty.fbrk,
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   198
        Simplifier.pretty_ss ctxt pre
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   199
      ],
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   200
      Pretty.block [
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   201
        Pretty.str "postprocessing simpset:",
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   202
        Pretty.fbrk,
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   203
        Simplifier.pretty_ss ctxt post
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   204
      ],
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   205
      Pretty.block (
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   206
        Pretty.str "function transformers:"
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   207
        :: Pretty.fbrk
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   208
        :: (Pretty.fbreaks o map Pretty.str) functrans
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   209
      )
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   210
    ]
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   211
  end;
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   212
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   213
fun simple_functrans f thy eqns = case f thy (map fst eqns)
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   214
 of SOME thms' => SOME (map (rpair (forall snd eqns)) thms')
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   215
  | NONE => NONE;
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   216
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   217
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   218
(** sort algebra and code equation graph types **)
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   219
30947
dd551284a300 re-engineering of evaluation conversions
haftmann
parents: 30942
diff changeset
   220
type code_algebra = (sort -> sort) * Sorts.algebra;
34891
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34251
diff changeset
   221
type code_graph = ((string * sort) list * Code.cert) Graph.T;
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   222
39604
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   223
fun get_node eqngr const = Graph.get_node eqngr const
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   224
  handle Graph.UNDEF _ => error ("No such constant in code equation graph: " ^ quote const);
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   225
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   226
fun cert eqngr = snd o get_node eqngr;
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   227
fun sortargs eqngr = map snd o fst o get_node eqngr;
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   228
fun all eqngr = Graph.keys eqngr;
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   229
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   230
fun pretty thy eqngr =
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   231
  AList.make (snd o Graph.get_node eqngr) (Graph.keys eqngr)
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31125
diff changeset
   232
  |> (map o apfst) (Code.string_of_const thy)
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   233
  |> sort (string_ord o pairself fst)
34895
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34893
diff changeset
   234
  |> map (fn (s, cert) => (Pretty.block o Pretty.fbreaks) (Pretty.str s :: Code.pretty_cert thy cert))
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   235
  |> Pretty.chunks;
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   236
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   237
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   238
(** the Waisenhaus algorithm **)
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   239
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   240
(* auxiliary *)
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   241
30942
1e246776f876 diagnostic commands now in code_thingol; tuned code of funny continuations
haftmann
parents: 30876
diff changeset
   242
fun is_proper_class thy = can (AxClass.get_info thy);
1e246776f876 diagnostic commands now in code_thingol; tuned code of funny continuations
haftmann
parents: 30876
diff changeset
   243
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   244
fun complete_proper_sort thy =
30942
1e246776f876 diagnostic commands now in code_thingol; tuned code of funny continuations
haftmann
parents: 30876
diff changeset
   245
  Sign.complete_sort thy #> filter (is_proper_class thy);
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   246
30029
d14d0b4bf5b4 also consider superclasses properly
haftmann
parents: 30024
diff changeset
   247
fun inst_params thy tyco =
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   248
  map (fn (c, _) => AxClass.param_of_inst thy (c, tyco))
30029
d14d0b4bf5b4 also consider superclasses properly
haftmann
parents: 30024
diff changeset
   249
    o maps (#params o AxClass.get_info thy);
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   250
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   251
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   252
(* data structures *)
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   253
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   254
datatype const = Fun of string | Inst of class * string;
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   255
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   256
fun const_ord (Fun c1, Fun c2) = fast_string_ord (c1, c2)
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   257
  | const_ord (Inst class_tyco1, Inst class_tyco2) =
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   258
      prod_ord fast_string_ord fast_string_ord (class_tyco1, class_tyco2)
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   259
  | const_ord (Fun _, Inst _) = LESS
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   260
  | const_ord (Inst _, Fun _) = GREATER;
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   261
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   262
type var = const * int;
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   263
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   264
structure Vargraph =
31971
8c1b845ed105 renamed functor TableFun to Table, and GraphFun to Graph;
wenzelm
parents: 31962
diff changeset
   265
  Graph(type key = var val ord = prod_ord const_ord int_ord);
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   266
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   267
datatype styp = Tyco of string * styp list | Var of var | Free;
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   268
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   269
fun styp_of c_lhs (Type (tyco, tys)) = Tyco (tyco, map (styp_of c_lhs) tys)
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   270
  | styp_of c_lhs (TFree (v, _)) = case c_lhs
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   271
     of SOME (c, lhs) => Var (Fun c, find_index (fn (v', _) => v = v') lhs)
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   272
      | NONE => Free;
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   273
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   274
type vardeps_data = ((string * styp list) list * class list) Vargraph.T
34891
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34251
diff changeset
   275
  * (((string * sort) list * Code.cert) Symtab.table
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   276
    * (class * string) list);
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   277
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   278
val empty_vardeps_data : vardeps_data =
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   279
  (Vargraph.empty, (Symtab.empty, []));
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   280
30876
613c2eb8aef6 tuned whitespace
haftmann
parents: 30769
diff changeset
   281
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   282
(* retrieving equations and instances from the background context *)
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   283
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   284
fun obtain_eqns thy eqngr c =
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   285
  case try (Graph.get_node eqngr) c
34891
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34251
diff changeset
   286
   of SOME (lhs, cert) => ((lhs, []), cert)
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   287
    | NONE => let
34891
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34251
diff changeset
   288
        val cert = Code.get_cert thy (preprocess thy) c;
35224
1c9866c5f6fb simplified
haftmann
parents: 34895
diff changeset
   289
        val (lhs, rhss) = Code.typargs_deps_of_cert thy cert;
34891
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34251
diff changeset
   290
      in ((lhs, rhss), cert) end;
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   291
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   292
fun obtain_instance thy arities (inst as (class, tyco)) =
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   293
  case AList.lookup (op =) arities inst
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   294
   of SOME classess => (classess, ([], []))
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   295
    | NONE => let
30029
d14d0b4bf5b4 also consider superclasses properly
haftmann
parents: 30024
diff changeset
   296
        val all_classes = complete_proper_sort thy [class];
37384
5aba26803073 more consistent naming aroud type classes and instances
haftmann
parents: 36960
diff changeset
   297
        val super_classes = remove (op =) class all_classes;
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   298
        val classess = map (complete_proper_sort thy)
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   299
          (Sign.arity_sorts thy tyco [class]);
30029
d14d0b4bf5b4 also consider superclasses properly
haftmann
parents: 30024
diff changeset
   300
        val inst_params = inst_params thy tyco all_classes;
37384
5aba26803073 more consistent naming aroud type classes and instances
haftmann
parents: 36960
diff changeset
   301
      in (classess, (super_classes, inst_params)) end;
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   302
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   303
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   304
(* computing instantiations *)
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   305
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   306
fun add_classes thy arities eqngr c_k new_classes vardeps_data =
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   307
  let
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   308
    val (styps, old_classes) = Vargraph.get_node (fst vardeps_data) c_k;
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   309
    val diff_classes = new_classes |> subtract (op =) old_classes;
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   310
  in if null diff_classes then vardeps_data
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   311
  else let
44338
700008399ee5 refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents: 43122
diff changeset
   312
    val c_ks = Vargraph.immediate_succs (fst vardeps_data) c_k |> insert (op =) c_k;
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   313
  in
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   314
    vardeps_data
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   315
    |> (apfst o Vargraph.map_node c_k o apsnd) (append diff_classes)
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   316
    |> fold (fn styp => fold (ensure_typmatch_inst thy arities eqngr styp) new_classes) styps
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   317
    |> fold (fn c_k => add_classes thy arities eqngr c_k diff_classes) c_ks
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   318
  end end
37384
5aba26803073 more consistent naming aroud type classes and instances
haftmann
parents: 36960
diff changeset
   319
and add_styp thy arities eqngr c_k new_tyco_styps vardeps_data =
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   320
  let
37384
5aba26803073 more consistent naming aroud type classes and instances
haftmann
parents: 36960
diff changeset
   321
    val (old_tyco_stypss, classes) = Vargraph.get_node (fst vardeps_data) c_k;
5aba26803073 more consistent naming aroud type classes and instances
haftmann
parents: 36960
diff changeset
   322
  in if member (op =) old_tyco_stypss new_tyco_styps then vardeps_data
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   323
  else
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   324
    vardeps_data
37384
5aba26803073 more consistent naming aroud type classes and instances
haftmann
parents: 36960
diff changeset
   325
    |> (apfst o Vargraph.map_node c_k o apfst) (cons new_tyco_styps)
5aba26803073 more consistent naming aroud type classes and instances
haftmann
parents: 36960
diff changeset
   326
    |> fold (ensure_typmatch_inst thy arities eqngr new_tyco_styps) classes
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   327
  end
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   328
and add_dep thy arities eqngr c_k c_k' vardeps_data =
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   329
  let
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   330
    val (_, classes) = Vargraph.get_node (fst vardeps_data) c_k;
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   331
  in
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   332
    vardeps_data
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   333
    |> add_classes thy arities eqngr c_k' classes
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   334
    |> apfst (Vargraph.add_edge (c_k, c_k'))
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   335
  end
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   336
and ensure_typmatch_inst thy arities eqngr (tyco, styps) class vardeps_data =
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   337
  if can (Sign.arity_sorts thy tyco) [class]
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   338
  then vardeps_data
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   339
    |> ensure_inst thy arities eqngr (class, tyco)
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   340
    |> fold_index (fn (k, styp) =>
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   341
         ensure_typmatch thy arities eqngr styp (Inst (class, tyco), k)) styps
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   342
  else vardeps_data (*permissive!*)
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   343
and ensure_inst thy arities eqngr (inst as (class, tyco)) (vardeps_data as (_, (_, insts))) =
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   344
  if member (op =) insts inst then vardeps_data
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   345
  else let
37384
5aba26803073 more consistent naming aroud type classes and instances
haftmann
parents: 36960
diff changeset
   346
    val (classess, (super_classes, inst_params)) =
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   347
      obtain_instance thy arities inst;
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   348
  in
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   349
    vardeps_data
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   350
    |> (apsnd o apsnd) (insert (op =) inst)
30083
41a20af1fb77 robustified
haftmann
parents: 30075
diff changeset
   351
    |> fold_index (fn (k, _) =>
41a20af1fb77 robustified
haftmann
parents: 30075
diff changeset
   352
         apfst (Vargraph.new_node ((Inst (class, tyco), k), ([] ,[])))) classess
37384
5aba26803073 more consistent naming aroud type classes and instances
haftmann
parents: 36960
diff changeset
   353
    |> fold (fn super_class => ensure_inst thy arities eqngr (super_class, tyco)) super_classes
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   354
    |> fold (ensure_fun thy arities eqngr) inst_params
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   355
    |> fold_index (fn (k, classes) =>
30075
ff5b4900d9a5 repaired order of variable node allocation
haftmann
parents: 30064
diff changeset
   356
         add_classes thy arities eqngr (Inst (class, tyco), k) classes
37384
5aba26803073 more consistent naming aroud type classes and instances
haftmann
parents: 36960
diff changeset
   357
         #> fold (fn super_class =>
5aba26803073 more consistent naming aroud type classes and instances
haftmann
parents: 36960
diff changeset
   358
             add_dep thy arities eqngr (Inst (super_class, tyco), k)
5aba26803073 more consistent naming aroud type classes and instances
haftmann
parents: 36960
diff changeset
   359
             (Inst (class, tyco), k)) super_classes
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   360
         #> fold (fn inst_param =>
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   361
             add_dep thy arities eqngr (Fun inst_param, k)
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   362
             (Inst (class, tyco), k)
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   363
             ) inst_params
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   364
         ) classess
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   365
  end
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   366
and ensure_typmatch thy arities eqngr (Tyco tyco_styps) c_k vardeps_data =
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   367
      vardeps_data
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   368
      |> add_styp thy arities eqngr c_k tyco_styps
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   369
  | ensure_typmatch thy arities eqngr (Var c_k') c_k vardeps_data =
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   370
      vardeps_data
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   371
      |> add_dep thy arities eqngr c_k c_k'
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   372
  | ensure_typmatch thy arities eqngr Free c_k vardeps_data =
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   373
      vardeps_data
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   374
and ensure_rhs thy arities eqngr (c', styps) vardeps_data =
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   375
  vardeps_data
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   376
  |> ensure_fun thy arities eqngr c'
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   377
  |> fold_index (fn (k, styp) =>
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   378
       ensure_typmatch thy arities eqngr styp (Fun c', k)) styps
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   379
and ensure_fun thy arities eqngr c (vardeps_data as (_, (eqntab, _))) =
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   380
  if Symtab.defined eqntab c then vardeps_data
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   381
  else let
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   382
    val ((lhs, rhss), eqns) = obtain_eqns thy eqngr c;
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   383
    val rhss' = (map o apsnd o map) (styp_of (SOME (c, lhs))) rhss;
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   384
  in
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   385
    vardeps_data
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   386
    |> (apsnd o apfst) (Symtab.update_new (c, (lhs, eqns)))
30083
41a20af1fb77 robustified
haftmann
parents: 30075
diff changeset
   387
    |> fold_index (fn (k, _) =>
41a20af1fb77 robustified
haftmann
parents: 30075
diff changeset
   388
         apfst (Vargraph.new_node ((Fun c, k), ([] ,[])))) lhs
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   389
    |> fold_index (fn (k, (_, sort)) =>
30083
41a20af1fb77 robustified
haftmann
parents: 30075
diff changeset
   390
         add_classes thy arities eqngr (Fun c, k) (complete_proper_sort thy sort)) lhs
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   391
    |> fold (ensure_rhs thy arities eqngr) rhss'
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   392
  end;
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   393
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   394
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   395
(* applying instantiations *)
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   396
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   397
fun dicts_of thy (proj_sort, algebra) (T, sort) =
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   398
  let
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   399
    fun class_relation (x, _) _ = x;
36102
a51d1d154c71 of_sort_derivation: pass-through full type information -- following the version by krauss/schropp;
wenzelm
parents: 35624
diff changeset
   400
    fun type_constructor (tyco, _) xs class =
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   401
      inst_params thy tyco (Sorts.complete_sort algebra [class])
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   402
        @ (maps o maps) fst xs;
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   403
    fun type_variable (TFree (_, sort)) = map (pair []) (proj_sort sort);
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   404
  in
32795
a0f38d8d633a Sorts.of_sort_derivation: no pp here;
wenzelm
parents: 32544
diff changeset
   405
    flat (Sorts.of_sort_derivation algebra
36102
a51d1d154c71 of_sort_derivation: pass-through full type information -- following the version by krauss/schropp;
wenzelm
parents: 35624
diff changeset
   406
      { class_relation = K class_relation, type_constructor = type_constructor,
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   407
        type_variable = type_variable } (T, proj_sort sort)
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   408
       handle Sorts.CLASS_ERROR _ => [] (*permissive!*))
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   409
  end;
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   410
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   411
fun add_arity thy vardeps (class, tyco) =
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 32873
diff changeset
   412
  AList.default (op =) ((class, tyco),
34891
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34251
diff changeset
   413
    map_range (fn k => (snd o Vargraph.get_node vardeps) (Inst (class, tyco), k))
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34251
diff changeset
   414
      (Sign.arity_number thy tyco));
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   415
34891
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34251
diff changeset
   416
fun add_cert thy vardeps (c, (proto_lhs, proto_cert)) (rhss, eqngr) =
30058
f84c2412e870 more liberality needed
haftmann
parents: 30054
diff changeset
   417
  if can (Graph.get_node eqngr) c then (rhss, eqngr)
f84c2412e870 more liberality needed
haftmann
parents: 30054
diff changeset
   418
  else let
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   419
    val lhs = map_index (fn (k, (v, _)) =>
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   420
      (v, snd (Vargraph.get_node vardeps (Fun c, k)))) proto_lhs;
38291
62abd53f37fa minimize sorts in certificate instantiation
haftmann
parents: 37744
diff changeset
   421
    val cert = Code.constrain_cert thy (map (Sign.minimize_sort thy o snd) lhs) proto_cert;
35224
1c9866c5f6fb simplified
haftmann
parents: 34895
diff changeset
   422
    val (vs, rhss') = Code.typargs_deps_of_cert thy cert;
34891
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34251
diff changeset
   423
    val eqngr' = Graph.new_node (c, (vs, cert)) eqngr;
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   424
  in (map (pair c) rhss' @ rhss, eqngr') end;
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   425
32873
333945c9ac6a tuned handling of type variable names further
haftmann
parents: 32872
diff changeset
   426
fun extend_arities_eqngr thy cs ts (arities, (eqngr : code_graph)) =
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   427
  let
30942
1e246776f876 diagnostic commands now in code_thingol; tuned code of funny continuations
haftmann
parents: 30876
diff changeset
   428
    val cs_rhss = (fold o fold_aterms) (fn Const (c_ty as (c, _)) =>
1e246776f876 diagnostic commands now in code_thingol; tuned code of funny continuations
haftmann
parents: 30876
diff changeset
   429
      insert (op =) (c, (map (styp_of NONE) o Sign.const_typargs thy) c_ty) | _ => I) ts [];
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   430
    val (vardeps, (eqntab, insts)) = empty_vardeps_data
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   431
      |> fold (ensure_fun thy arities eqngr) cs
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   432
      |> fold (ensure_rhs thy arities eqngr) cs_rhss;
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   433
    val arities' = fold (add_arity thy vardeps) insts arities;
47005
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 46961
diff changeset
   434
    val algebra = Sorts.subalgebra (Context.pretty_global thy) (is_proper_class thy)
30064
3cd19b113854 use canonical subalgebra projection
haftmann
parents: 30061
diff changeset
   435
      (AList.lookup (op =) arities') (Sign.classes_of thy);
34891
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34251
diff changeset
   436
    val (rhss, eqngr') = Symtab.fold (add_cert thy vardeps) eqntab ([], eqngr);
30942
1e246776f876 diagnostic commands now in code_thingol; tuned code of funny continuations
haftmann
parents: 30876
diff changeset
   437
    fun deps_of (c, rhs) = c :: maps (dicts_of thy algebra)
32873
333945c9ac6a tuned handling of type variable names further
haftmann
parents: 32872
diff changeset
   438
      (rhs ~~ sortargs eqngr' c);
30054
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   439
    val eqngr'' = fold (fn (c, rhs) => fold
36d7d337510e simplified evaluation
haftmann
parents: 30050
diff changeset
   440
      (curry Graph.add_edge c) (deps_of rhs)) rhss eqngr';
30942
1e246776f876 diagnostic commands now in code_thingol; tuned code of funny continuations
haftmann
parents: 30876
diff changeset
   441
  in (algebra, (arities', eqngr'')) end;
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   442
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   443
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   444
(** store for preprocessed arities and code equations **)
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   445
34173
458ced35abb8 reduced code generator cache to the baremost minimum
haftmann
parents: 33942
diff changeset
   446
structure Wellsorted = Code_Data
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   447
(
30947
dd551284a300 re-engineering of evaluation conversions
haftmann
parents: 30942
diff changeset
   448
  type T = ((string * class) * sort list) list * code_graph;
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   449
  val empty = ([], Graph.empty);
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   450
);
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   451
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   452
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   453
(** retrieval and evaluation interfaces **)
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   454
39398
2e30660a2e21 ignore code cache optionally; corrected scope of term value in static_eval_conv
haftmann
parents: 39133
diff changeset
   455
fun obtain ignore_cache thy consts ts = apsnd snd
2e30660a2e21 ignore code cache optionally; corrected scope of term value in static_eval_conv
haftmann
parents: 39133
diff changeset
   456
  (Wellsorted.change_yield (if ignore_cache then NONE else SOME thy) (extend_arities_eqngr thy consts ts));
38670
3c7db0192db9 added static_eval_conv
haftmann
parents: 38669
diff changeset
   457
3c7db0192db9 added static_eval_conv
haftmann
parents: 38669
diff changeset
   458
fun dest_cterm ct = let val t = Thm.term_of ct in (Term.add_tfrees t [], t) end;
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   459
41184
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 40176
diff changeset
   460
fun dynamic_conv thy conv = no_variables_conv (fn ct =>
30942
1e246776f876 diagnostic commands now in code_thingol; tuned code of funny continuations
haftmann
parents: 30876
diff changeset
   461
  let
39475
9cc1ba3c5706 separation of static and dynamic thy context
haftmann
parents: 39398
diff changeset
   462
    val thm1 = preprocess_conv thy ct;
9cc1ba3c5706 separation of static and dynamic thy context
haftmann
parents: 39398
diff changeset
   463
    val ct' = Thm.rhs_of thm1;
38670
3c7db0192db9 added static_eval_conv
haftmann
parents: 38669
diff changeset
   464
    val (vs', t') = dest_cterm ct';
30947
dd551284a300 re-engineering of evaluation conversions
haftmann
parents: 30942
diff changeset
   465
    val consts = fold_aterms
dd551284a300 re-engineering of evaluation conversions
haftmann
parents: 30942
diff changeset
   466
      (fn Const (c, _) => insert (op =) c | _ => I) t' [];
39398
2e30660a2e21 ignore code cache optionally; corrected scope of term value in static_eval_conv
haftmann
parents: 39133
diff changeset
   467
    val (algebra', eqngr') = obtain false thy consts [t'];
39475
9cc1ba3c5706 separation of static and dynamic thy context
haftmann
parents: 39398
diff changeset
   468
    val thm2 = conv algebra' eqngr' vs' t' ct';
9cc1ba3c5706 separation of static and dynamic thy context
haftmann
parents: 39398
diff changeset
   469
    val thm3 = postprocess_conv thy (Thm.rhs_of thm2);
9cc1ba3c5706 separation of static and dynamic thy context
haftmann
parents: 39398
diff changeset
   470
  in
9cc1ba3c5706 separation of static and dynamic thy context
haftmann
parents: 39398
diff changeset
   471
    Thm.transitive thm1 (Thm.transitive thm2 thm3) handle THM _ =>
9cc1ba3c5706 separation of static and dynamic thy context
haftmann
parents: 39398
diff changeset
   472
      error ("could not construct evaluation proof:\n"
9cc1ba3c5706 separation of static and dynamic thy context
haftmann
parents: 39398
diff changeset
   473
      ^ (cat_lines o map (Display.string_of_thm_global thy)) [thm1, thm2, thm3])
39604
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   474
  end);
30947
dd551284a300 re-engineering of evaluation conversions
haftmann
parents: 30942
diff changeset
   475
41184
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 40176
diff changeset
   476
fun dynamic_value thy postproc evaluator t =
30942
1e246776f876 diagnostic commands now in code_thingol; tuned code of funny continuations
haftmann
parents: 30876
diff changeset
   477
  let
39604
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   478
    val (resubst, t') = preprocess_term thy t;
39475
9cc1ba3c5706 separation of static and dynamic thy context
haftmann
parents: 39398
diff changeset
   479
    val vs' = Term.add_tfrees t' [];
9cc1ba3c5706 separation of static and dynamic thy context
haftmann
parents: 39398
diff changeset
   480
    val consts = fold_aterms
9cc1ba3c5706 separation of static and dynamic thy context
haftmann
parents: 39398
diff changeset
   481
      (fn Const (c, _) => insert (op =) c | _ => I) t' [];
9cc1ba3c5706 separation of static and dynamic thy context
haftmann
parents: 39398
diff changeset
   482
    val (algebra', eqngr') = obtain false thy consts [t'];
9cc1ba3c5706 separation of static and dynamic thy context
haftmann
parents: 39398
diff changeset
   483
  in
40176
haftmann
parents: 40167
diff changeset
   484
    t'
haftmann
parents: 40167
diff changeset
   485
    |> evaluator algebra' eqngr' vs'
39604
f17fb9ccb836 avoid frees and vars in terms to be evaluated by abstracting and applying
haftmann
parents: 39475
diff changeset
   486
    |> postproc (postprocess_term thy o resubst)
39475
9cc1ba3c5706 separation of static and dynamic thy context
haftmann
parents: 39398
diff changeset
   487
  end;
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   488
41184
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 40176
diff changeset
   489
fun static_conv thy consts conv =
38670
3c7db0192db9 added static_eval_conv
haftmann
parents: 38669
diff changeset
   490
  let
39398
2e30660a2e21 ignore code cache optionally; corrected scope of term value in static_eval_conv
haftmann
parents: 39133
diff changeset
   491
    val (algebra, eqngr) = obtain true thy consts [];
41346
6673f6fa94ca canonical handling of theory context argument
haftmann
parents: 41226
diff changeset
   492
    val conv' = conv algebra eqngr;
39398
2e30660a2e21 ignore code cache optionally; corrected scope of term value in static_eval_conv
haftmann
parents: 39133
diff changeset
   493
  in
41189
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   494
    no_variables_conv ((preprocess_conv thy)
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   495
      then_conv (fn ct => uncurry conv' (dest_cterm ct) ct)
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   496
      then_conv (postprocess_conv thy))
39398
2e30660a2e21 ignore code cache optionally; corrected scope of term value in static_eval_conv
haftmann
parents: 39133
diff changeset
   497
  end;
38670
3c7db0192db9 added static_eval_conv
haftmann
parents: 38669
diff changeset
   498
41184
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 40176
diff changeset
   499
fun static_value thy postproc consts evaluator =
39475
9cc1ba3c5706 separation of static and dynamic thy context
haftmann
parents: 39398
diff changeset
   500
  let
9cc1ba3c5706 separation of static and dynamic thy context
haftmann
parents: 39398
diff changeset
   501
    val (algebra, eqngr) = obtain true thy consts [];
9cc1ba3c5706 separation of static and dynamic thy context
haftmann
parents: 39398
diff changeset
   502
    val evaluator' = evaluator algebra eqngr;
41189
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   503
  in 
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   504
    preprocess_term thy
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   505
    #-> (fn resubst => fn t => t
41346
6673f6fa94ca canonical handling of theory context argument
haftmann
parents: 41226
diff changeset
   506
      |> evaluator' (Term.add_tfrees t [])
41189
ba1eac745c5a more appropriate closures for static evaluation
haftmann
parents: 41184
diff changeset
   507
      |> postproc (postprocess_term thy o resubst))
39475
9cc1ba3c5706 separation of static and dynamic thy context
haftmann
parents: 39398
diff changeset
   508
  end;
9cc1ba3c5706 separation of static and dynamic thy context
haftmann
parents: 39398
diff changeset
   509
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   510
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   511
(** setup **)
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   512
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   513
val setup = 
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   514
  let
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   515
    fun mk_attribute f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I);
32072
d4bff63bcbf1 added code_unfold_post attribute
haftmann
parents: 31998
diff changeset
   516
    fun add_del_attribute_parser add del =
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31977
diff changeset
   517
      Attrib.add_del (mk_attribute add) (mk_attribute del);
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   518
  in
45189
80cb73210612 removing declaration of code_unfold to address the old code generator
bulwahn
parents: 44338
diff changeset
   519
    Attrib.setup @{binding code_unfold} (add_del_attribute_parser add_unfold del_unfold)
80cb73210612 removing declaration of code_unfold to address the old code generator
bulwahn
parents: 44338
diff changeset
   520
        "preprocessing equations for code generator"
32072
d4bff63bcbf1 added code_unfold_post attribute
haftmann
parents: 31998
diff changeset
   521
    #> Attrib.setup @{binding code_post} (add_del_attribute_parser add_post del_post)
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31977
diff changeset
   522
        "postprocessing equations for code generator"
46026
83caa4f4bd56 semiring_numeral_0_eq_0, semiring_numeral_1_eq_1 now [simp], superseeding corresponding simp rules on type nat
haftmann
parents: 45230
diff changeset
   523
    #> Attrib.setup @{binding code_abbrev} (Scan.succeed (mk_attribute add_code_abbrev))
83caa4f4bd56 semiring_numeral_0_eq_0, semiring_numeral_1_eq_1 now [simp], superseeding corresponding simp rules on type nat
haftmann
parents: 45230
diff changeset
   524
        "post- and preprocessing equations for code generator"
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   525
  end;
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   526
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   527
val _ =
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46497
diff changeset
   528
  Outer_Syntax.improper_command @{command_spec "print_codeproc"} "print code preprocessor setup"
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46497
diff changeset
   529
    (Scan.succeed
31125
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   530
      (Toplevel.no_timing o Toplevel.unknown_theory o Toplevel.keep
80218ee73167 transferred code generator preprocessor into separate module
haftmann
parents: 31089
diff changeset
   531
        (print_codeproc o Toplevel.theory_of)));
30010
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   532
862fc7751a15 tuned and incremental version of wellsorting algorithm
haftmann
parents:
diff changeset
   533
end; (*struct*)