src/Tools/Code/code_simp.ML
author wenzelm
Mon, 24 Oct 2016 15:00:13 +0200
changeset 64375 74a2af7c5145
parent 63164 72aaf69328fc
child 69593 3dda49e08b9d
permissions -rw-r--r--
remove old target: it is on the classpath and may break scalac;
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
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55147
diff changeset
    10
  val dynamic_conv: Proof.context -> conv
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55147
diff changeset
    11
  val dynamic_tac: Proof.context -> int -> tactic
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55147
diff changeset
    12
  val dynamic_value: Proof.context -> term -> term
56973
62da80041afd syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents: 56969
diff changeset
    13
  val static_conv: { ctxt: Proof.context, simpset: simpset option, consts: string list }
62da80041afd syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents: 56969
diff changeset
    14
    -> Proof.context -> conv
62da80041afd syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents: 56969
diff changeset
    15
  val static_tac: { ctxt: Proof.context, simpset: simpset option, consts: string list }
62da80041afd syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents: 56969
diff changeset
    16
    -> Proof.context -> int -> tactic
53147
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    17
  val trace: bool Config.T
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    18
end;
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    19
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    20
structure Code_Simp : CODE_SIMP =
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    21
struct
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    22
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    23
(* dedicated simpset *)
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    24
38759
37a9092de102 simplification/standardization of some theory data;
wenzelm
parents: 38671
diff changeset
    25
structure Simpset = Theory_Data
37a9092de102 simplification/standardization of some theory data;
wenzelm
parents: 38671
diff changeset
    26
(
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    27
  type T = simpset;
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    28
  val empty = empty_ss;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48072
diff changeset
    29
  val extend = I;
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    30
  val merge = merge_ss;
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    31
);
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    32
63164
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63160
diff changeset
    33
fun map_ss f thy =
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63160
diff changeset
    34
  Simpset.map (simpset_map (Proof_Context.init_global thy) f) thy;
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    35
63164
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63160
diff changeset
    36
fun simpset_default ctxt some_ss =
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63160
diff changeset
    37
  the_default (Simpset.get (Proof_Context.theory_of ctxt)) some_ss;
54929
f1ded3cea58d proper context for simplifier invocations in code generation stack
haftmann
parents: 54928
diff changeset
    38
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    39
54928
cb077b02c9a4 spelling
haftmann
parents: 54895
diff changeset
    40
(* diagnostic *)
53147
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    41
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    42
val trace = Attrib.setup_config_bool @{binding "code_simp_trace"} (K false);
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    43
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    44
fun set_trace ctxt =
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    45
  let
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    46
    val global = Config.get ctxt trace;
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    47
  in
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    48
    ctxt |> Config.map Simplifier.simp_trace (fn b => b orelse global)
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    49
  end;
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    50
8e8941fea278 separate tracing option for code_simp
haftmann
parents: 52736
diff changeset
    51
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55147
diff changeset
    52
(* build simpset context and conversion from program *)
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    53
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 54929
diff changeset
    54
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
    55
      ss
317663b422bb more correct context for dynamic invocations of static code conversions etc.
haftmann
parents: 51717
diff changeset
    56
      |> 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
    57
      |> fold Simplifier.add_cong (the_list some_cong)
48072
ace701efe203 prefer records with speaking labels over deeply nested tuples
haftmann
parents: 47576
diff changeset
    58
  | 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
    59
      ss
317663b422bb more correct context for dynamic invocations of static code conversions etc.
haftmann
parents: 51717
diff changeset
    60
      |> fold Simplifier.add_simp (map (fst o snd) inst_params)
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    61
  | add_stmt _ ss = ss;
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    62
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 54929
diff changeset
    63
val add_program = Code_Symbol.Graph.fold (add_stmt o fst o snd);
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    64
63164
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63160
diff changeset
    65
val simpset_program =
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63160
diff changeset
    66
  Code_Preproc.timed "building simpset" #ctxt
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63160
diff changeset
    67
  (fn { ctxt, some_ss, program } => simpset_map ctxt (add_program program) (simpset_default ctxt some_ss));
37461
3489cea0e0e9 conclude simplification with default simpset
haftmann
parents: 37449
diff changeset
    68
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55147
diff changeset
    69
fun rewrite_modulo ctxt some_ss program =
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55147
diff changeset
    70
  let
63164
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63160
diff changeset
    71
    val ss = simpset_program
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63160
diff changeset
    72
      { ctxt = ctxt, some_ss = some_ss, program = program };
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63160
diff changeset
    73
  in fn ctxt => 
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63160
diff changeset
    74
    Code_Preproc.timed_conv "simplifying"
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63160
diff changeset
    75
      Simplifier.full_rewrite (ctxt |> put_simpset ss |> set_trace)
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63160
diff changeset
    76
  end;
52736
317663b422bb more correct context for dynamic invocations of static code conversions etc.
haftmann
parents: 51717
diff changeset
    77
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55147
diff changeset
    78
fun conclude_tac ctxt some_ss =
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55147
diff changeset
    79
  let
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55147
diff changeset
    80
    val ss = simpset_default ctxt some_ss
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55147
diff changeset
    81
  in fn ctxt => Simplifier.full_simp_tac (ctxt |> put_simpset ss) end;
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    82
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    83
39486
e992bcc4440e refined static_eval_conv_simple; tuned comments
haftmann
parents: 39475
diff changeset
    84
(* evaluation with dynamic code context *)
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    85
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55147
diff changeset
    86
fun dynamic_conv ctxt = Code_Thingol.dynamic_conv ctxt
56969
7491932da574 dropped obsolete hand-waving adjustion of type variables: safely done in preprocessor
haftmann
parents: 56924
diff changeset
    87
  (fn program => fn _ => fn _ => rewrite_modulo ctxt NONE program ctxt);
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    88
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55147
diff changeset
    89
fun dynamic_tac ctxt = CONVERSION (dynamic_conv ctxt)
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55147
diff changeset
    90
  THEN' conclude_tac ctxt NONE ctxt;
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    91
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55147
diff changeset
    92
fun dynamic_value ctxt =
59633
a372513af1e2 clarified context;
wenzelm
parents: 59621
diff changeset
    93
  snd o Logic.dest_equals o Thm.prop_of o dynamic_conv ctxt o Thm.cterm_of ctxt;
37444
2e7e7ff21e25 added simp evaluator
haftmann
parents: 37442
diff changeset
    94
59323
468bd3aedfa1 modernized and more uniform style
haftmann
parents: 56973
diff changeset
    95
val _ = Theory.setup 
468bd3aedfa1 modernized and more uniform style
haftmann
parents: 56973
diff changeset
    96
  (Method.setup @{binding code_simp}
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55147
diff changeset
    97
    (Scan.succeed (SIMPLE_METHOD' o (CHANGED_PROP oo dynamic_tac)))
59323
468bd3aedfa1 modernized and more uniform style
haftmann
parents: 56973
diff changeset
    98
    "simplification with code equations");
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    99
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
   100
39486
e992bcc4440e refined static_eval_conv_simple; tuned comments
haftmann
parents: 39475
diff changeset
   101
(* evaluation with static code context *)
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
   102
56973
62da80041afd syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents: 56969
diff changeset
   103
fun static_conv { ctxt, simpset, consts } =
63156
3cb84e4469a7 clarified names of variants
haftmann
parents: 59633
diff changeset
   104
  Code_Thingol.static_conv_isa { ctxt = ctxt, consts = consts }
63164
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63160
diff changeset
   105
    (fn program => let
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63160
diff changeset
   106
       val conv = rewrite_modulo ctxt simpset program
72aaf69328fc optional timing for code generator conversions
haftmann
parents: 63160
diff changeset
   107
     in fn ctxt => fn _ => conv ctxt end);
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
   108
56973
62da80041afd syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents: 56969
diff changeset
   109
fun static_tac { ctxt, simpset, consts } =
52736
317663b422bb more correct context for dynamic invocations of static code conversions etc.
haftmann
parents: 51717
diff changeset
   110
  let
56973
62da80041afd syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents: 56969
diff changeset
   111
    val conv = static_conv { ctxt = ctxt, simpset = simpset, consts = consts };
62da80041afd syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents: 56969
diff changeset
   112
    val tac = conclude_tac ctxt simpset;
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55147
diff changeset
   113
  in fn ctxt' => CONVERSION (conv ctxt') THEN' (tac ctxt') end;
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
   114
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
   115
end;