src/Tools/Code/code_simp.ML
author haftmann
Thu, 01 Jul 2010 08:12:40 +0200
changeset 37661 f6b592f2aca4
parent 37463 7315100b916d
child 37744 3daaf23b9ab4
permissions -rw-r--r--
revert to plain for now mkdir
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
     1
(*  Title:      Tools/code/code_simp.ML
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
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
     9
  val no_frees_conv: conv -> conv
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    10
  val map_ss: (simpset -> simpset) -> theory -> theory
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    11
  val current_conv: theory -> conv
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    12
  val current_tac: theory -> int -> tactic
37444
2e7e7ff21e25 added simp evaluator
haftmann
parents: 37442
diff changeset
    13
  val current_value: theory -> term -> term
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    14
  val make_conv: theory -> simpset option -> string list -> conv
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    15
  val make_tac: theory -> simpset option -> string list -> int -> tactic
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    16
  val setup: theory -> theory
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
(* avoid free variables during conversion *)
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    23
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    24
fun no_frees_conv conv ct =
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    25
  let
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    26
    val frees = Thm.add_cterm_frees ct [];
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    27
    fun apply_beta free thm = Thm.combination thm (Thm.reflexive free)
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    28
      |> Conv.fconv_rule (Conv.arg_conv (Conv.try_conv (Thm.beta_conversion false)))
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    29
      |> Conv.fconv_rule (Conv.arg1_conv (Thm.beta_conversion false));
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    30
  in
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    31
    ct
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    32
    |> fold_rev Thm.cabs frees
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    33
    |> conv
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    34
    |> fold apply_beta frees
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    35
  end;
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    36
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    37
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    38
(* dedicated simpset *)
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    39
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    40
structure Simpset = Theory_Data (
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    41
  type T = simpset;
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    42
  val empty = empty_ss;
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    43
  fun extend ss = MetaSimplifier.inherit_context empty_ss ss;
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    44
  val merge = merge_ss;
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    45
);
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    46
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    47
val map_ss = Simpset.map;
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    48
37461
3489cea0e0e9 conclude simplification with default simpset
haftmann
parents: 37449
diff changeset
    49
fun simpset_default thy = Simplifier.global_context thy o the_default (Simpset.get thy);
3489cea0e0e9 conclude simplification with default simpset
haftmann
parents: 37449
diff changeset
    50
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    51
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    52
(* build simpset and conversion from program *)
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    53
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    54
fun add_stmt (Code_Thingol.Fun (_, ((_, eqs), some_cong))) ss =
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    55
      ss addsimps (map_filter (fst o snd)) eqs addcongs (the_list some_cong)
37449
034ebe92f090 more precise code
haftmann
parents: 37444
diff changeset
    56
  | add_stmt (Code_Thingol.Classinst (_, (_, (classparam_instances, _)))) ss =
034ebe92f090 more precise code
haftmann
parents: 37444
diff changeset
    57
      ss addsimps (map (fst o snd) classparam_instances)
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    58
  | add_stmt _ ss = ss;
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    59
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    60
val add_program = Graph.fold (add_stmt o fst o snd)
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    61
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    62
fun rewrite_modulo thy some_ss program = Simplifier.full_rewrite
37461
3489cea0e0e9 conclude simplification with default simpset
haftmann
parents: 37449
diff changeset
    63
  (add_program program (simpset_default thy some_ss));
3489cea0e0e9 conclude simplification with default simpset
haftmann
parents: 37449
diff changeset
    64
3489cea0e0e9 conclude simplification with default simpset
haftmann
parents: 37449
diff changeset
    65
fun conclude_tac thy some_ss = Simplifier.full_simp_tac (simpset_default thy some_ss);
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    66
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    67
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    68
(* evaluation with current code context *)
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    69
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    70
fun current_conv thy = no_frees_conv (Code_Thingol.eval_conv thy
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    71
  (fn naming => fn program => fn t => fn deps => rewrite_modulo thy NONE program));
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    72
37461
3489cea0e0e9 conclude simplification with default simpset
haftmann
parents: 37449
diff changeset
    73
fun current_tac thy = CONVERSION (current_conv thy) THEN' conclude_tac thy NONE;
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    74
37461
3489cea0e0e9 conclude simplification with default simpset
haftmann
parents: 37449
diff changeset
    75
fun current_value thy = snd o Logic.dest_equals o Thm.prop_of o current_conv thy o Thm.cterm_of thy;
37444
2e7e7ff21e25 added simp evaluator
haftmann
parents: 37442
diff changeset
    76
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    77
val setup = Method.setup (Binding.name "code_simp")
37463
7315100b916d code_simp: only succeed on real progress
haftmann
parents: 37461
diff changeset
    78
  (Scan.succeed (SIMPLE_METHOD' o (CHANGED_PROP oo current_tac o ProofContext.theory_of)))
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    79
  "simplification with code equations"
37444
2e7e7ff21e25 added simp evaluator
haftmann
parents: 37442
diff changeset
    80
  #> Value.add_evaluator ("simp", current_value o ProofContext.theory_of);
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    81
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    82
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    83
(* evaluation with freezed code context *)
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    84
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    85
fun make_conv thy some_ss consts = no_frees_conv (Code_Preproc.pre_post_conv thy
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    86
  ((rewrite_modulo thy some_ss o snd o snd o Code_Thingol.consts_program thy false) consts));
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    87
37461
3489cea0e0e9 conclude simplification with default simpset
haftmann
parents: 37449
diff changeset
    88
fun make_tac thy some_ss consts = CONVERSION (make_conv thy some_ss consts)
3489cea0e0e9 conclude simplification with default simpset
haftmann
parents: 37449
diff changeset
    89
  THEN' conclude_tac thy some_ss;
37442
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    90
037ee7b712b2 added code_simp infrastructure
haftmann
parents:
diff changeset
    91
end;