src/HOL/Tools/reflection.ML
author haftmann
Fri, 31 May 2013 09:30:32 +0200
changeset 52274 35a2668ac3b0
parent 52273 3d7472b6b720
child 52275 9b4c04da53b1
permissions -rw-r--r--
permit multiple variable arguments in reflect
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Tools/reflection.ML
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
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
     4
A trial for automatical reification.
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
     5
*)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
     6
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
     7
signature REFLECTION =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
     8
sig
52272
47d138cae708 dropped vacuous prefix
haftmann
parents: 51726
diff changeset
     9
  val reify: Proof.context -> thm list -> term -> thm
47d138cae708 dropped vacuous prefix
haftmann
parents: 51726
diff changeset
    10
  val reify_tac: Proof.context -> thm list -> term option -> int -> tactic
47d138cae708 dropped vacuous prefix
haftmann
parents: 51726
diff changeset
    11
  val reflect: Proof.context -> (cterm -> thm)
47d138cae708 dropped vacuous prefix
haftmann
parents: 51726
diff changeset
    12
    -> thm list -> thm list -> term -> thm
47d138cae708 dropped vacuous prefix
haftmann
parents: 51726
diff changeset
    13
  val reflection_tac: Proof.context -> (cterm -> thm)
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    14
    -> thm list -> thm list -> term option -> int -> tactic
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    15
  val get_default: Proof.context -> { reification_eqs: thm list, correctness_thms: thm list }
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    16
  val add_reification_eq: attribute
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    17
  val del_reification_eq: attribute
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    18
  val add_correctness_thm: attribute
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    19
  val del_correctness_thm: attribute
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    20
  val default_reify_tac: Proof.context -> thm list -> term option -> int -> tactic
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    21
  val default_reflection_tac: Proof.context -> thm list -> thm list -> term option -> int -> tactic
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    22
end;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    23
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    24
structure Reflection : REFLECTION =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    25
struct
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    26
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    27
val FWD = curry (op OF);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    28
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    29
fun dest_listT (Type (@{type_name "list"}, [T])) = T;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    30
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    31
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    32
(* Make a congruence rule out of a defining equation for the interpretation
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    33
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    34
   th is one defining equation of f,
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    35
     i.e. th is "f (Cp ?t1 ... ?tn) = P(f ?t1, .., f ?tn)" 
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    36
   Cp is a constructor pattern and P is a pattern 
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    37
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    38
   The result is:
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    39
     [|?A1 = f ?t1 ; .. ; ?An= f ?tn |] ==> P (?A1, .., ?An) = f (Cp ?t1 .. ?tn)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    40
       + the a list of names of the A1 .. An, Those are fresh in the ctxt *)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    41
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    42
fun mk_congeq ctxt fs th =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    43
  let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    44
    val Const (fN, _) = th |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    45
      |> fst |> strip_comb |> fst;
52273
haftmann
parents: 52272
diff changeset
    46
    val cert = Thm.cterm_of (Proof_Context.theory_of ctxt);
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    47
    val ((_, [th']), ctxt') = Variable.import true [th] ctxt;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    48
    val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of th'));
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    49
    fun add_fterms (t as t1 $ t2) =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    50
          if exists (fn f => Term.could_unify (t |> strip_comb |> fst, f)) fs
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    51
          then insert (op aconv) t
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    52
          else add_fterms t1 #> add_fterms t2
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    53
      | add_fterms (t as Abs _) =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    54
          if exists_Const (fn (c, _) => c = fN) t
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    55
          then K [t]
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    56
          else K []
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    57
      | add_fterms _ = I;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    58
    val fterms = add_fterms rhs [];
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    59
    val (xs, ctxt'') = Variable.variant_fixes (replicate (length fterms) "x") ctxt';
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    60
    val tys = map fastype_of fterms;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    61
    val vs = map Free (xs ~~ tys);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    62
    val env = fterms ~~ vs; (*FIXME*)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    63
    fun replace_fterms (t as t1 $ t2) =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    64
        (case AList.lookup (op aconv) env t of
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    65
            SOME v => v
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    66
          | NONE => replace_fterms t1 $ replace_fterms t2)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    67
      | replace_fterms t =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    68
        (case AList.lookup (op aconv) env t of
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    69
            SOME v => v
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    70
          | NONE => t);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    71
    fun mk_def (Abs (x, xT, t), v) =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    72
          HOLogic.mk_Trueprop (HOLogic.all_const xT $ Abs (x, xT, HOLogic.mk_eq (v $ Bound 0, t)))
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    73
      | mk_def (t, v) = HOLogic.mk_Trueprop (HOLogic.mk_eq (v, t));
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    74
    fun tryext x =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    75
      (x RS @{lemma "(\<forall>x. f x = g x) \<Longrightarrow> f = g" by blast} handle THM _ => x);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    76
    val cong =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    77
      (Goal.prove ctxt'' [] (map mk_def env)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    78
        (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, replace_fterms rhs)))
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    79
        (fn {context, prems, ...} =>
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    80
          Local_Defs.unfold_tac context (map tryext prems) THEN rtac th' 1)) RS sym;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    81
    val (cong' :: vars') =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    82
      Variable.export ctxt'' ctxt (cong :: map (Drule.mk_term o cert) vs);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    83
    val vs' = map (fst o fst o Term.dest_Var o Thm.term_of o Drule.dest_term) vars';
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    84
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    85
  in (vs', cong') end;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    86
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    87
(* congs is a list of pairs (P,th) where th is a theorem for
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    88
     [| f p1 = A1; ...; f pn = An|] ==> f (C p1 .. pn) = P *)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    89
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    90
fun rearrange congs =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    91
  let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    92
    fun P (_, th) =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    93
      let val @{term "Trueprop"} $ (Const (@{const_name HOL.eq}, _) $ l $ _) = concl_of th
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    94
      in can dest_Var l end;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    95
    val (yes, no) = List.partition P congs;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    96
  in no @ yes end;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    97
52272
47d138cae708 dropped vacuous prefix
haftmann
parents: 51726
diff changeset
    98
fun reify ctxt eqs t =
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    99
  let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   100
    fun index_of t bds =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   101
      let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   102
        val tt = HOLogic.listT (fastype_of t);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   103
      in
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   104
        (case AList.lookup Type.could_unify bds tt of
52273
haftmann
parents: 52272
diff changeset
   105
            NONE => error "index_of: type not found in environements!"
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   106
          | SOME (tbs, tats) =>
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   107
              let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   108
                val i = find_index (fn t' => t' = t) tats;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   109
                val j = find_index (fn t' => t' = t) tbs;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   110
              in
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   111
                if j = ~1 then
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   112
                  if i = ~1
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   113
                  then (length tbs + length tats, AList.update Type.could_unify (tt, (tbs, tats @ [t])) bds)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   114
                  else (i, bds)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   115
                else (j, bds)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   116
              end)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   117
      end;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   118
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   119
    (* Generic decomp for reification : matches the actual term with the
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   120
       rhs of one cong rule. The result of the matching guides the
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   121
       proof synthesis: The matches of the introduced Variables A1 .. An are
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   122
       processed recursively
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   123
       The rest is instantiated in the cong rule,i.e. no reification is needed *)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   124
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   125
    (* da is the decomposition for atoms, ie. it returns ([],g) where g
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   126
       returns the right instance f (AtC n) = t , where AtC is the Atoms
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   127
       constructor and n is the number of the atom corresponding to t *)
52272
47d138cae708 dropped vacuous prefix
haftmann
parents: 51726
diff changeset
   128
    fun decomp_reify da cgns (t, ctxt) bds =
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   129
      let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   130
        val thy = Proof_Context.theory_of ctxt;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   131
        val cert = cterm_of thy;
52273
haftmann
parents: 52272
diff changeset
   132
        val certT = ctyp_of thy;
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   133
        fun tryabsdecomp (s, ctxt) bds =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   134
          (case s of
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   135
            Abs (_, xT, ta) =>
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   136
              let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   137
                val ([raw_xn], ctxt') = Variable.variant_fixes ["x"] ctxt;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   138
                val (xn, ta) = Syntax_Trans.variant_abs (raw_xn, xT, ta);  (* FIXME !? *)
52273
haftmann
parents: 52272
diff changeset
   139
                val x = Free (xn, xT);
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   140
                val bds = (case AList.lookup Type.could_unify bds (HOLogic.listT xT) of
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   141
                    NONE => error "tryabsdecomp: Type not found in the Environement"
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   142
                  | SOME (bsT, atsT) => AList.update Type.could_unify (HOLogic.listT xT, (x :: bsT, atsT)) bds);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   143
               in (([(ta, ctxt')],
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   144
                    fn ([th], bds) =>
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   145
                      (hd (Variable.export ctxt' ctxt [(Thm.forall_intr (cert x) th) COMP allI]),
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   146
                       let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   147
                         val (bsT, asT) = the (AList.lookup Type.could_unify bds (HOLogic.listT xT));
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   148
                       in
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   149
                         AList.update Type.could_unify (HOLogic.listT xT,(tl bsT, asT)) bds
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   150
                       end)),
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   151
                   bds)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   152
               end
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   153
           | _ => da (s, ctxt) bds)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   154
      in
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   155
        (case cgns of
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   156
          [] => tryabsdecomp (t, ctxt) bds
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   157
        | ((vns, cong) :: congs) =>
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   158
            (let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   159
              val (tyenv, tmenv) =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   160
                Pattern.match thy
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   161
                  ((fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) (concl_of cong), t)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   162
                  (Vartab.empty, Vartab.empty);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   163
              val (fnvs, invs) = List.partition (fn ((vn, _),_) => member (op =) vns vn) (Vartab.dest tmenv);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   164
              val (fts, its) =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   165
                (map (snd o snd) fnvs,
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   166
                 map (fn ((vn, vi), (tT, t)) => (cert (Var ((vn, vi), tT)), cert t)) invs);
52273
haftmann
parents: 52272
diff changeset
   167
              val ctyenv = map (fn ((vn, vi), (s, ty)) => (certT (TVar((vn, vi), s)), certT ty)) (Vartab.dest tyenv);
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   168
            in
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   169
              ((fts ~~ replicate (length fts) ctxt,
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   170
                 apfst (FWD (Drule.instantiate_normalize (ctyenv, its) cong))), bds)
52272
47d138cae708 dropped vacuous prefix
haftmann
parents: 51726
diff changeset
   171
            end handle Pattern.MATCH => decomp_reify da congs (t,ctxt) bds))
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   172
      end;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   173
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   174
 (* looks for the atoms equation and instantiates it with the right number *)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   175
    fun mk_decompatom eqs (t, ctxt) bds = (([], fn (_, bds) =>
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   176
      let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   177
        val tT = fastype_of t;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   178
        fun isat eq =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   179
          let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   180
            val rhs = eq |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   181
          in exists_Const
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   182
            (fn (n, ty) => n = @{const_name "List.nth"}
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   183
              andalso AList.defined Type.could_unify bds (domain_type ty)) rhs
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   184
              andalso Type.could_unify (fastype_of rhs, tT)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   185
          end;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   186
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   187
        fun get_nths t acc =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   188
          case t of
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   189
            Const(@{const_name "List.nth"}, _) $ vs $ n => insert (fn ((a, _), (b, _)) => a aconv b) (t, (vs, n)) acc
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   190
          | t1 $ t2 => get_nths t1 (get_nths t2 acc)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   191
          | Abs (_ ,_ ,t') => get_nths t' acc
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   192
          | _ => acc;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   193
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   194
        fun tryeqs [] bds = error "Can not find the atoms equation"
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   195
          | tryeqs (eq :: eqs) bds = ((
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   196
              let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   197
                val rhs = eq |> prop_of |> HOLogic.dest_Trueprop  |> HOLogic.dest_eq |> snd;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   198
                val nths = get_nths rhs [];
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   199
                val (vss, _) = fold_rev (fn (_, (vs, n)) => fn (vss, ns) =>
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   200
                  (insert (op aconv) vs vss, insert (op aconv) n ns)) nths ([], []);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   201
                val (vsns, ctxt') = Variable.variant_fixes (replicate (length vss) "vs") ctxt;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   202
                val (xns, ctxt'') = Variable.variant_fixes (replicate (length nths) "x") ctxt';
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   203
                val thy = Proof_Context.theory_of ctxt'';
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   204
                val cert = cterm_of thy;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   205
                val certT = ctyp_of thy;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   206
                val vsns_map = vss ~~ vsns;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   207
                val xns_map = fst (split_list nths) ~~ xns;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   208
                val subst = map (fn (nt, xn) => (nt, Var ((xn, 0), fastype_of nt))) xns_map;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   209
                val rhs_P = subst_free subst rhs;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   210
                val (tyenv, tmenv) = Pattern.match thy (rhs_P, t) (Vartab.empty, Vartab.empty);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   211
                val sbst = Envir.subst_term (tyenv, tmenv);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   212
                val sbsT = Envir.subst_type tyenv;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   213
                val subst_ty = map (fn (n, (s, t)) =>
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   214
                  (certT (TVar (n, s)), certT t)) (Vartab.dest tyenv)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   215
                val tml = Vartab.dest tmenv;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   216
                val (subst_ns, bds) = fold_map
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   217
                  (fn (Const _ $ _ $ n, Var (xn0, _)) => fn bds =>
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   218
                    let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   219
                      val name = snd (the (AList.lookup (op =) tml xn0));
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   220
                      val (idx, bds) = index_of name bds;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   221
                    in ((cert n, idx |> (HOLogic.mk_nat #> cert)), bds) end) subst bds;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   222
                val subst_vs =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   223
                  let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   224
                    fun h (Const _ $ (vs as Var (_, lT)) $ _, Var (_, T)) =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   225
                      let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   226
                        val cns = sbst (Const (@{const_name "List.Cons"}, T --> lT --> lT));
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   227
                        val lT' = sbsT lT;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   228
                        val (bsT, _) = the (AList.lookup Type.could_unify bds lT);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   229
                        val vsn = the (AList.lookup (op =) vsns_map vs);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   230
                        val cvs = cert (fold_rev (fn x => fn xs => cns $ x $xs) bsT (Free (vsn, lT')));
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   231
                      in (cert vs, cvs) end;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   232
                  in map h subst end;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   233
                val cts = map (fn ((vn, vi), (tT, t)) => (cert (Var ((vn, vi), tT)), cert t))
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   234
                  (fold (AList.delete (fn (((a : string), _), (b, _)) => a = b))
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   235
                    (map (fn n => (n, 0)) xns) tml);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   236
                val substt =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   237
                  let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   238
                    val ih = Drule.cterm_rule (Thm.instantiate (subst_ty, []));
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   239
                  in map (fn (v, t) => (ih v, ih t)) (subst_ns @ subst_vs @ cts) end;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   240
                val th = (Drule.instantiate_normalize (subst_ty, substt) eq) RS sym;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   241
              in (hd (Variable.export ctxt'' ctxt [th]), bds) end)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   242
              handle Pattern.MATCH => tryeqs eqs bds)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   243
          in tryeqs (filter isat eqs) bds end), bds);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   244
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   245
  (* Generic reification procedure: *)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   246
  (* creates all needed cong rules and then just uses the theorem synthesis *)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   247
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   248
    fun mk_congs ctxt eqs =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   249
      let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   250
        val fs = fold_rev (fn eq => insert (op =) (eq |> prop_of |> HOLogic.dest_Trueprop
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   251
          |> HOLogic.dest_eq |> fst |> strip_comb
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   252
          |> fst)) eqs [];
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   253
        val tys = fold_rev (fn f => fold (insert (op =)) (f |> fastype_of |> binder_types |> tl)) fs [];
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   254
        val (vs, ctxt') = Variable.variant_fixes (replicate (length tys) "vs") ctxt;
52273
haftmann
parents: 52272
diff changeset
   255
        val cert = cterm_of (Proof_Context.theory_of ctxt');
haftmann
parents: 52272
diff changeset
   256
        val subst =
haftmann
parents: 52272
diff changeset
   257
          the o AList.lookup (op =) (map2 (fn T => fn v => (T, cert (Free (v, T)))) tys vs);
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   258
        fun prep_eq eq =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   259
          let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   260
            val (_, _ :: vs) = eq |> prop_of |> HOLogic.dest_Trueprop
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   261
              |> HOLogic.dest_eq |> fst |> strip_comb;
52273
haftmann
parents: 52272
diff changeset
   262
            val subst = map_filter (fn (v as Var (_, T)) => SOME (cert v, subst T)
haftmann
parents: 52272
diff changeset
   263
              | _ => NONE) vs;
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   264
          in Thm.instantiate ([], subst) eq end;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   265
        val (ps, congs) = map_split (mk_congeq ctxt' fs o prep_eq) eqs;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   266
        val bds = AList.make (K ([], [])) tys;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   267
      in (ps ~~ Variable.export ctxt' ctxt congs, bds) end
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   268
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   269
    val (congs, bds) = mk_congs ctxt eqs;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   270
    val congs = rearrange congs;
52272
47d138cae708 dropped vacuous prefix
haftmann
parents: 51726
diff changeset
   271
    val (th, bds) = divide_and_conquer' (decomp_reify (mk_decompatom eqs) congs) (t,ctxt) bds;
52273
haftmann
parents: 52272
diff changeset
   272
    fun is_list_var (Var (_, t)) = can dest_listT t
haftmann
parents: 52272
diff changeset
   273
      | is_list_var _ = false;
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   274
    val vars = th |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd
52273
haftmann
parents: 52272
diff changeset
   275
      |> strip_comb |> snd |> filter is_list_var;
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   276
    val cert = cterm_of (Proof_Context.theory_of ctxt);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   277
    val cvs = map (fn (v as Var(_, t)) => (cert v,
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   278
      the (AList.lookup Type.could_unify bds t) |> snd |> HOLogic.mk_list (dest_listT t) |> cert)) vars;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   279
    val th' = Drule.instantiate_normalize ([], cvs) th;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   280
    val t' = (fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) th';
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   281
    val th'' = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, t')))
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   282
      (fn _ => simp_tac ctxt 1)
52273
haftmann
parents: 52272
diff changeset
   283
  in FWD trans [th'', th'] end;
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   284
52272
47d138cae708 dropped vacuous prefix
haftmann
parents: 51726
diff changeset
   285
fun reflect ctxt conv corr_thms eqs t =
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   286
  let
52272
47d138cae708 dropped vacuous prefix
haftmann
parents: 51726
diff changeset
   287
    val reify_thm = reify ctxt eqs t;
52273
haftmann
parents: 52272
diff changeset
   288
    fun try_corr corr_thm =
haftmann
parents: 52272
diff changeset
   289
      SOME (FWD trans [reify_thm, corr_thm RS sym]) handle THM _ => NONE;
52274
35a2668ac3b0 permit multiple variable arguments in reflect
haftmann
parents: 52273
diff changeset
   290
    val refl_thm = case get_first try_corr corr_thms
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   291
     of NONE => error "No suitable correctness theorem found"
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   292
      | SOME thm => thm;
52274
35a2668ac3b0 permit multiple variable arguments in reflect
haftmann
parents: 52273
diff changeset
   293
    val rhs = (Thm.dest_arg o Thm.dest_arg o cprop_of) refl_thm;
35a2668ac3b0 permit multiple variable arguments in reflect
haftmann
parents: 52273
diff changeset
   294
    val number_of_args = (length o snd o strip_comb o term_of) rhs;
35a2668ac3b0 permit multiple variable arguments in reflect
haftmann
parents: 52273
diff changeset
   295
    val reified = Thm.dest_arg (fold_range (K Thm.dest_fun) (number_of_args - 1) rhs);
35a2668ac3b0 permit multiple variable arguments in reflect
haftmann
parents: 52273
diff changeset
   296
    val evaluated = conv reified;
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   297
  in
52274
35a2668ac3b0 permit multiple variable arguments in reflect
haftmann
parents: 52273
diff changeset
   298
    refl_thm
35a2668ac3b0 permit multiple variable arguments in reflect
haftmann
parents: 52273
diff changeset
   299
    |> simplify (put_simpset HOL_basic_ss ctxt addsimps [evaluated])
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   300
    |> simplify (put_simpset HOL_basic_ss ctxt addsimps eqs addsimps @{thms nth_Cons_0 nth_Cons_Suc})
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   301
  end;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   302
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   303
fun tac_of_thm mk_thm to = SUBGOAL (fn (goal, i) =>
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   304
  let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   305
    val t = (case to of NONE => HOLogic.dest_Trueprop goal | SOME t => t)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   306
    val thm = mk_thm t RS ssubst;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   307
  in rtac thm i end);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   308
 
52272
47d138cae708 dropped vacuous prefix
haftmann
parents: 51726
diff changeset
   309
fun reify_tac ctxt eqs = tac_of_thm (reify ctxt eqs);
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   310
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   311
(*Reflection calls reification and uses the correctness theorem assumed to be the head of the list*)
52272
47d138cae708 dropped vacuous prefix
haftmann
parents: 51726
diff changeset
   312
fun reflection_tac ctxt conv corr_thms eqs =
47d138cae708 dropped vacuous prefix
haftmann
parents: 51726
diff changeset
   313
  tac_of_thm (reflect ctxt conv corr_thms eqs);
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   314
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   315
structure Data = Generic_Data
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   316
(
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   317
  type T = thm list * thm list;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   318
  val empty = ([], []);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   319
  val extend = I;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   320
  fun merge ((ths1, rths1), (ths2, rths2)) =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   321
    (Thm.merge_thms (ths1, ths2), Thm.merge_thms (rths1, rths2));
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   322
);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   323
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   324
fun get_default ctxt =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   325
  let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   326
    val (reification_eqs, correctness_thms) = Data.get (Context.Proof ctxt);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   327
  in { reification_eqs = reification_eqs, correctness_thms = correctness_thms } end;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   328
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   329
val add_reification_eq = Thm.declaration_attribute (Data.map o apfst o Thm.add_thm);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   330
val del_reification_eq = Thm.declaration_attribute (Data.map o apfst o Thm.del_thm);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   331
val add_correctness_thm = Thm.declaration_attribute (Data.map o apsnd o Thm.add_thm);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   332
val del_correctness_thm = Thm.declaration_attribute (Data.map o apsnd o Thm.del_thm);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   333
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   334
val _ = Context.>> (Context.map_theory
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   335
  (Attrib.setup @{binding reify}
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   336
    (Attrib.add_del add_reification_eq del_reification_eq) "declare reification equations" #>
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   337
  Attrib.setup @{binding reflection}
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   338
    (Attrib.add_del add_correctness_thm del_correctness_thm) "declare reflection correctness theorems"));
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   339
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   340
fun default_reify_tac ctxt user_eqs =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   341
  let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   342
    val { reification_eqs = default_eqs, correctness_thms = _ } =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   343
      get_default ctxt;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   344
    val eqs = fold Thm.add_thm user_eqs default_eqs;
52272
47d138cae708 dropped vacuous prefix
haftmann
parents: 51726
diff changeset
   345
  in reify_tac ctxt eqs end;
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   346
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   347
fun default_reflection_tac ctxt user_thms user_eqs =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   348
  let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   349
    val { reification_eqs = default_eqs, correctness_thms = default_thms } =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   350
      get_default ctxt;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   351
    val corr_thms = fold Thm.add_thm user_thms default_thms;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   352
    val eqs = fold Thm.add_thm user_eqs default_eqs; 
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   353
    val conv = Code_Evaluation.dynamic_conv (Proof_Context.theory_of ctxt);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   354
      (*FIXME why Code_Evaluation.dynamic_conv? very specific*)
52272
47d138cae708 dropped vacuous prefix
haftmann
parents: 51726
diff changeset
   355
  in reflection_tac ctxt conv corr_thms eqs end;
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   356
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   357
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   358
end