src/HOL/Tools/reflection.ML
author wenzelm
Thu, 20 Jun 2013 11:27:43 +0200
changeset 52418 f00e4d8dde4c
parent 52406 1e57c3c4e05c
child 53171 a5e54d4d9081
permissions -rw-r--r--
updated to Isabelle2013;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
52406
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
     1
(*  Title:      HOL/Tools/reflection.ML
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
     2
    Author:     Amine Chaieb, TU Muenchen
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
     3
52406
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
     4
A trial for automatical reflection with user-space declarations.
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
     5
*)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
     6
52406
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
     7
signature REFLECTION =
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
     8
sig
52406
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
     9
  val reflect: Proof.context -> thm list -> thm list -> conv
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    10
  val reflection_tac: Proof.context -> thm list -> thm list -> term option -> int -> tactic
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    11
  val reflect_with_eval: Proof.context -> thm list -> thm list -> conv -> conv
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    12
  val reflection_with_eval_tac: Proof.context -> thm list -> thm list -> conv -> term option -> int -> tactic
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    13
  val get_default: Proof.context -> { reification_eqs: thm list, correctness_thms: thm list }
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    14
  val add_reification_eq: attribute
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    15
  val del_reification_eq: attribute
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    16
  val add_correctness_thm: attribute
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    17
  val del_correctness_thm: attribute
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    18
  val default_reify_tac: Proof.context -> thm list -> term option -> int -> tactic
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    19
  val default_reflection_tac: Proof.context -> thm list -> thm list -> term option -> int -> tactic
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    20
end;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    21
52406
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    22
structure Reflection : REFLECTION =
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    23
struct
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    24
52406
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    25
fun subst_correctness corr_thms ct =
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    26
  Conv.rewrs_conv (map (Thm.symmetric o mk_eq) corr_thms) ct
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    27
    handle CTERM _ => error "No suitable correctness theorem found";
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    28
52406
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    29
fun reflect ctxt corr_thms eqs =
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    30
  (Reification.conv ctxt eqs) then_conv (subst_correctness corr_thms)
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    31
52406
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    32
fun reflection_tac ctxt corr_thms eqs =
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    33
  Reification.lift_conv ctxt (reflect ctxt corr_thms eqs);
52275
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
    34
52406
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    35
fun first_arg_conv conv =
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    36
  let
52406
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    37
    fun conv' ct =
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    38
      if can Thm.dest_comb (fst (Thm.dest_comb ct))
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    39
      then Conv.combination_conv conv' Conv.all_conv ct
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    40
      else Conv.combination_conv Conv.all_conv conv ct;
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    41
  in conv' end;
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    42
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    43
fun reflect_with_eval ctxt corr_thms eqs conv =
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    44
  (reflect ctxt corr_thms eqs) then_conv (first_arg_conv conv) then_conv (Reification.dereify ctxt eqs);
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    45
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    46
fun reflection_with_eval_tac ctxt corr_thms eqs conv =
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    47
  Reification.lift_conv ctxt (reflect_with_eval ctxt corr_thms eqs conv);
52288
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
    48
52406
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    49
structure Data = Generic_Data
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    50
(
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    51
  type T = thm list * thm list;
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    52
  val empty = ([], []);
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    53
  val extend = I;
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    54
  fun merge ((ths1, rths1), (ths2, rths2)) =
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    55
    (Thm.merge_thms (ths1, ths2), Thm.merge_thms (rths1, rths2));
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    56
);
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    57
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    58
fun get_default ctxt =
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    59
  let
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    60
    val (reification_eqs, correctness_thms) = Data.get (Context.Proof ctxt);
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    61
  in { reification_eqs = reification_eqs, correctness_thms = correctness_thms } end;
52288
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
    62
52406
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    63
val add_reification_eq = Thm.declaration_attribute (Data.map o apfst o Thm.add_thm);
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    64
val del_reification_eq = Thm.declaration_attribute (Data.map o apfst o Thm.del_thm);
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    65
val add_correctness_thm = Thm.declaration_attribute (Data.map o apsnd o Thm.add_thm);
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    66
val del_correctness_thm = Thm.declaration_attribute (Data.map o apsnd o Thm.del_thm);
52288
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
    67
52406
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    68
val _ = Context.>> (Context.map_theory
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    69
  (Attrib.setup @{binding reify}
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    70
    (Attrib.add_del add_reification_eq del_reification_eq) "declare reification equations" #>
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    71
  Attrib.setup @{binding reflection}
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    72
    (Attrib.add_del add_correctness_thm del_correctness_thm) "declare reflection correctness theorems"));
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    73
52406
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    74
fun default_reify_tac ctxt user_eqs =
52288
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
    75
  let
52406
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    76
    val { reification_eqs = default_eqs, correctness_thms = _ } =
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    77
      get_default ctxt;
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    78
    val eqs = fold Thm.add_thm user_eqs default_eqs;
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    79
  in Reification.tac ctxt eqs end;
52288
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
    80
52406
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    81
fun default_reflection_tac ctxt user_thms user_eqs =
52288
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
    82
  let
52406
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    83
    val { reification_eqs = default_eqs, correctness_thms = default_thms } =
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    84
      get_default ctxt;
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    85
    val corr_thms = fold Thm.add_thm user_thms default_thms;
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    86
    val eqs = fold Thm.add_thm user_eqs default_eqs; 
1e57c3c4e05c updated documentation of sort hypotheses;
wenzelm
parents: 52288
diff changeset
    87
  in reflection_tac ctxt corr_thms eqs end;
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    88
52276
329c41438154 reflection without evaluation
haftmann
parents: 52275
diff changeset
    89
end;