src/Tools/Code/code_simp.ML
author wenzelm
Sun, 23 Feb 2014 21:11:59 +0100
changeset 55694 a1184dfb8e00
parent 55147 bce3dbc11f95
child 55757 9fc71814b8c1
permissions -rw-r--r--
clarified semantic completion: retain kind.full_name as official item name for history; misc tuning;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37744
3daaf23b9ab4 tuned titles
haftmann
parents: 37463
diff changeset
     1
(*  Title:      Tools/Code/code_simp.ML
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
     2
    Author:     Florian Haftmann, TU Muenchen
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
     3
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
     4
Connecting the simplifier and the code generator.
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
     5
*)
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
     6
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
     7
signature CODE_SIMP =
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
     8
sig
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48072
diff changeset
     9
  val map_ss: (Proof.context -> Proof.context) -> theory -> theory
41247
c5cb19ecbd41 avoid slightly odd Conv.tap_thy
haftmann
parents: 41188
diff changeset
    10
  val dynamic_conv: theory -> conv
41188
7cded8957e72 more uniform naming
haftmann
parents: 41184
diff changeset
    11
  val dynamic_tac: theory -> int -> tactic
41184
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 39606
diff changeset
    12
  val dynamic_value: theory -> term -> term
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 39606
diff changeset
    13
  val static_conv: theory -> simpset option -> string list -> conv
52736
317663b422bb more correct context for dynamic invocations of static code conversions etc.
haftmann
parents: 51717
diff changeset
    14
  val static_tac: theory -> simpset option -> string list -> Proof.context -> int -> tactic
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    15
  val setup: theory -> theory
53147
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    16
  val trace: bool Config.T
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    17
end;
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    18
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    19
structure Code_Simp : CODE_SIMP =
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    20
struct
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    21
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    22
(* dedicated simpset *)
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    23
38759
37a9092de102 simplification/standardization of some theory data;
wenzelm
parents: 38671
diff changeset
    24
structure Simpset = Theory_Data
37a9092de102 simplification/standardization of some theory data;
wenzelm
parents: 38671
diff changeset
    25
(
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    26
  type T = simpset;
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    27
  val empty = empty_ss;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48072
diff changeset
    28
  val extend = I;
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    29
  val merge = merge_ss;
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    30
);
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    31
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48072
diff changeset
    32
fun map_ss f thy = Simpset.map (simpset_map (Proof_Context.init_global thy) f) thy;
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    33
54929
f1ded3cea58d proper context for simplifier invocations in code generation stack
haftmann
parents: 54928
diff changeset
    34
fun simpset_default thy some_ss = the_default (Simpset.get thy) some_ss;
f1ded3cea58d proper context for simplifier invocations in code generation stack
haftmann
parents: 54928
diff changeset
    35
f1ded3cea58d proper context for simplifier invocations in code generation stack
haftmann
parents: 54928
diff changeset
    36
fun simp_ctxt_default thy some_ss =
f1ded3cea58d proper context for simplifier invocations in code generation stack
haftmann
parents: 54928
diff changeset
    37
  Proof_Context.init_global thy
f1ded3cea58d proper context for simplifier invocations in code generation stack
haftmann
parents: 54928
diff changeset
    38
  |> put_simpset (simpset_default thy some_ss);
37461
3489cea0e0e9 conclude simplification with default simpset
haftmann
parents: 37449
diff changeset
    39
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    40
54928
cb077b02c9a4 spelling
haftmann
parents: 54895
diff changeset
    41
(* diagnostic *)
53147
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    42
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    43
val trace = Attrib.setup_config_bool @{binding "code_simp_trace"} (K false);
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    44
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    45
fun set_trace ctxt =
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    46
  let
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    47
    val global = Config.get ctxt trace;
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    48
  in
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    49
    ctxt |> Config.map Simplifier.simp_trace (fn b => b orelse global)
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    50
  end;
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    51
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    52
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    53
(* build simpset and conversion from program *)
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    54
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 54929
diff changeset
    55
fun add_stmt (Code_Thingol.Fun ((_, eqs), some_cong)) ss =
52736
317663b422bb more correct context for dynamic invocations of static code conversions etc.
haftmann
parents: 51717
diff changeset
    56
      ss
317663b422bb more correct context for dynamic invocations of static code conversions etc.
haftmann
parents: 51717
diff changeset
    57
      |> fold Simplifier.add_simp ((map_filter (fst o snd)) eqs)
45620
f2a587696afb modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 43619
diff changeset
    58
      |> fold Simplifier.add_cong (the_list some_cong)
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 47576
diff changeset
    59
  | add_stmt (Code_Thingol.Classinst { inst_params, ... }) ss =
52736
317663b422bb more correct context for dynamic invocations of static code conversions etc.
haftmann
parents: 51717
diff changeset
    60
      ss
317663b422bb more correct context for dynamic invocations of static code conversions etc.
haftmann
parents: 51717
diff changeset
    61
      |> fold Simplifier.add_simp (map (fst o snd) inst_params)
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    62
  | add_stmt _ ss = ss;
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    63
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 54929
diff changeset
    64
val add_program = Code_Symbol.Graph.fold (add_stmt o fst o snd);
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    65
54929
f1ded3cea58d proper context for simplifier invocations in code generation stack
haftmann
parents: 54928
diff changeset
    66
fun simp_ctxt_program thy some_ss program =
f1ded3cea58d proper context for simplifier invocations in code generation stack
haftmann
parents: 54928
diff changeset
    67
  simp_ctxt_default thy some_ss
f1ded3cea58d proper context for simplifier invocations in code generation stack
haftmann
parents: 54928
diff changeset
    68
  |> add_program program;
37461
3489cea0e0e9 conclude simplification with default simpset
haftmann
parents: 37449
diff changeset
    69
52736
317663b422bb more correct context for dynamic invocations of static code conversions etc.
haftmann
parents: 51717
diff changeset
    70
fun rewrite_modulo thy some_ss program =
54929
f1ded3cea58d proper context for simplifier invocations in code generation stack
haftmann
parents: 54928
diff changeset
    71
  Simplifier.full_rewrite (simp_ctxt_program thy some_ss program |> set_trace);
52736
317663b422bb more correct context for dynamic invocations of static code conversions etc.
haftmann
parents: 51717
diff changeset
    72
317663b422bb more correct context for dynamic invocations of static code conversions etc.
haftmann
parents: 51717
diff changeset
    73
fun conclude_tac thy some_ss =
54929
f1ded3cea58d proper context for simplifier invocations in code generation stack
haftmann
parents: 54928
diff changeset
    74
  Simplifier.full_simp_tac o Simplifier.put_simpset (simpset_default thy some_ss);
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    75
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    76
39486
e992bcc4440e refined static_eval_conv_simple; tuned comments
haftmann
parents: 39475
diff changeset
    77
(* evaluation with dynamic code context *)
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    78
41247
c5cb19ecbd41 avoid slightly odd Conv.tap_thy
haftmann
parents: 41188
diff changeset
    79
fun dynamic_conv thy = Code_Thingol.dynamic_conv thy
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 54929
diff changeset
    80
  (fn program => fn _ => fn _ => rewrite_modulo thy NONE program);
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    81
52736
317663b422bb more correct context for dynamic invocations of static code conversions etc.
haftmann
parents: 51717
diff changeset
    82
fun dynamic_tac thy = CONVERSION (dynamic_conv thy) THEN' conclude_tac thy NONE (Proof_Context.init_global thy);
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    83
41247
c5cb19ecbd41 avoid slightly odd Conv.tap_thy
haftmann
parents: 41188
diff changeset
    84
fun dynamic_value thy = snd o Logic.dest_equals o Thm.prop_of o dynamic_conv thy o Thm.cterm_of thy;
37444
2e7e7ff21e25 added simp evaluator
haftmann
parents: 37442
diff changeset
    85
43619
3803869014aa proper @{binding} antiquotations (relevant for formal references);
wenzelm
parents: 42361
diff changeset
    86
val setup =
3803869014aa proper @{binding} antiquotations (relevant for formal references);
wenzelm
parents: 42361
diff changeset
    87
  Method.setup @{binding code_simp}
3803869014aa proper @{binding} antiquotations (relevant for formal references);
wenzelm
parents: 42361
diff changeset
    88
    (Scan.succeed (SIMPLE_METHOD' o (CHANGED_PROP oo dynamic_tac o Proof_Context.theory_of)))
3803869014aa proper @{binding} antiquotations (relevant for formal references);
wenzelm
parents: 42361
diff changeset
    89
    "simplification with code equations"
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41346
diff changeset
    90
  #> Value.add_evaluator ("simp", dynamic_value o Proof_Context.theory_of);
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    91
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    92
39486
e992bcc4440e refined static_eval_conv_simple; tuned comments
haftmann
parents: 39475
diff changeset
    93
(* evaluation with static code context *)
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    94
41184
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 39606
diff changeset
    95
fun static_conv thy some_ss consts =
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 39606
diff changeset
    96
  Code_Thingol.static_conv_simple thy consts
41346
6673f6fa94ca canonical handling of theory context argument
haftmann
parents: 41251
diff changeset
    97
    (fn program => fn _ => fn _ => rewrite_modulo thy some_ss program);
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    98
52736
317663b422bb more correct context for dynamic invocations of static code conversions etc.
haftmann
parents: 51717
diff changeset
    99
fun static_tac thy some_ss consts =
317663b422bb more correct context for dynamic invocations of static code conversions etc.
haftmann
parents: 51717
diff changeset
   100
  let
317663b422bb more correct context for dynamic invocations of static code conversions etc.
haftmann
parents: 51717
diff changeset
   101
    val conv = static_conv thy some_ss consts;
317663b422bb more correct context for dynamic invocations of static code conversions etc.
haftmann
parents: 51717
diff changeset
   102
    val tac = conclude_tac thy some_ss;
317663b422bb more correct context for dynamic invocations of static code conversions etc.
haftmann
parents: 51717
diff changeset
   103
  in fn ctxt => CONVERSION conv THEN' tac ctxt end;
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
   104
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
   105
end;