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