src/HOL/Tools/reification.ML
author hoelzl
Wed, 02 Apr 2014 18:35:07 +0200
changeset 56371 fb9ae0727548
parent 52288 ca4932dad084
child 59058 a78612c67ec0
permissions -rw-r--r--
extend continuous_intros; remove continuous_on_intros and isCont_intros
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
52286
8170e5327c02 make reification part of HOL
haftmann
parents: 52276
diff changeset
     1
(*  Title:      HOL/Tools/reification.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
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
52286
8170e5327c02 make reification part of HOL
haftmann
parents: 52276
diff changeset
     7
signature REIFICATION =
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
     8
sig
52286
8170e5327c02 make reification part of HOL
haftmann
parents: 52276
diff changeset
     9
  val conv: Proof.context -> thm list -> conv
8170e5327c02 make reification part of HOL
haftmann
parents: 52276
diff changeset
    10
  val tac: Proof.context -> thm list -> term option -> int -> tactic
8170e5327c02 make reification part of HOL
haftmann
parents: 52276
diff changeset
    11
  val lift_conv: Proof.context -> conv -> term option -> int -> tactic
8170e5327c02 make reification part of HOL
haftmann
parents: 52276
diff changeset
    12
  val dereify: Proof.context -> thm list -> conv
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    13
end;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    14
52286
8170e5327c02 make reification part of HOL
haftmann
parents: 52276
diff changeset
    15
structure Reification : REIFICATION =
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    16
struct
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    17
52275
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
    18
fun dest_listT (Type (@{type_name "list"}, [T])) = T;
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
    19
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    20
val FWD = curry (op OF);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    21
52275
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
    22
fun rewrite_with ctxt eqs = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps eqs);
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
    23
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
    24
val pure_subst = @{lemma "x == y ==> PROP P y ==> PROP P x" by simp}
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    25
52276
329c41438154 reflection without evaluation
haftmann
parents: 52275
diff changeset
    26
fun lift_conv ctxt conv some_t = Subgoal.FOCUS (fn { concl, ... } =>
52275
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
    27
  let
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
    28
    val ct = case some_t
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
    29
     of NONE => Thm.dest_arg concl
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
    30
      | SOME t => Thm.cterm_of (Proof_Context.theory_of ctxt) t
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
    31
    val thm = conv ct;
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
    32
  in
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
    33
    if Thm.is_reflexive thm then no_tac
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
    34
    else ALLGOALS (rtac (pure_subst OF [thm]))
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
    35
  end) ctxt;
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    36
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    37
(* Make a congruence rule out of a defining equation for the interpretation
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    38
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    39
   th is one defining equation of f,
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    40
     i.e. th is "f (Cp ?t1 ... ?tn) = P(f ?t1, .., f ?tn)" 
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    41
   Cp is a constructor pattern and P is a pattern 
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    42
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    43
   The result is:
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    44
     [|?A1 = f ?t1 ; .. ; ?An= f ?tn |] ==> P (?A1, .., ?An) = f (Cp ?t1 .. ?tn)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    45
       + 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
    46
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    47
fun mk_congeq ctxt fs th =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    48
  let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    49
    val Const (fN, _) = th |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    50
      |> fst |> strip_comb |> fst;
52273
haftmann
parents: 52272
diff changeset
    51
    val cert = Thm.cterm_of (Proof_Context.theory_of ctxt);
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    52
    val ((_, [th']), ctxt') = Variable.import true [th] ctxt;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    53
    val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of th'));
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    54
    fun add_fterms (t as t1 $ t2) =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    55
          if exists (fn f => Term.could_unify (t |> strip_comb |> fst, f)) fs
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    56
          then insert (op aconv) t
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    57
          else add_fterms t1 #> add_fterms t2
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    58
      | add_fterms (t as Abs _) =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    59
          if exists_Const (fn (c, _) => c = fN) t
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    60
          then K [t]
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    61
          else K []
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    62
      | add_fterms _ = I;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    63
    val fterms = add_fterms rhs [];
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    64
    val (xs, ctxt'') = Variable.variant_fixes (replicate (length fterms) "x") ctxt';
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    65
    val tys = map fastype_of fterms;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    66
    val vs = map Free (xs ~~ tys);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    67
    val env = fterms ~~ vs; (*FIXME*)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    68
    fun replace_fterms (t as t1 $ t2) =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    69
        (case AList.lookup (op aconv) env t of
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    70
            SOME v => v
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    71
          | NONE => replace_fterms t1 $ replace_fterms t2)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    72
      | replace_fterms t =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    73
        (case AList.lookup (op aconv) env t of
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    74
            SOME v => v
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    75
          | NONE => t);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    76
    fun mk_def (Abs (x, xT, t), v) =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    77
          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
    78
      | mk_def (t, v) = HOLogic.mk_Trueprop (HOLogic.mk_eq (v, t));
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    79
    fun tryext x =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    80
      (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
    81
    val cong =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    82
      (Goal.prove ctxt'' [] (map mk_def env)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    83
        (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, replace_fterms rhs)))
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    84
        (fn {context, prems, ...} =>
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    85
          Local_Defs.unfold_tac context (map tryext prems) THEN rtac th' 1)) RS sym;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    86
    val (cong' :: vars') =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    87
      Variable.export ctxt'' ctxt (cong :: map (Drule.mk_term o cert) vs);
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    88
    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
    89
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    90
  in (vs', cong') end;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    91
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    92
(* congs is a list of pairs (P,th) where th is a theorem for
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    93
     [| f p1 = A1; ...; f pn = An|] ==> f (C p1 .. pn) = P *)
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    94
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    95
fun rearrange congs =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    96
  let
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    97
    fun P (_, th) =
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    98
      let val @{term "Trueprop"} $ (Const (@{const_name HOL.eq}, _) $ l $ _) = concl_of th
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
    99
      in can dest_Var l end;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   100
    val (yes, no) = List.partition P congs;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   101
  in no @ yes end;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   102
52275
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
   103
fun dereify ctxt eqs =
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
   104
  rewrite_with ctxt (eqs @ @{thms nth_Cons_0 nth_Cons_Suc});
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
   105
52288
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   106
fun index_of t bds =
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   107
  let
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   108
    val tt = HOLogic.listT (fastype_of t);
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   109
  in
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   110
    (case AList.lookup Type.could_unify bds tt of
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   111
        NONE => error "index_of: type not found in environements!"
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   112
      | SOME (tbs, tats) =>
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   113
          let
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   114
            val i = find_index (fn t' => t' = t) tats;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   115
            val j = find_index (fn t' => t' = t) tbs;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   116
          in
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   117
            if j = ~1 then
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   118
              if i = ~1
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   119
              then (length tbs + length tats, AList.update Type.could_unify (tt, (tbs, tats @ [t])) bds)
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   120
              else (i, bds)
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   121
            else (j, bds)
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   122
          end)
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   123
  end;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   124
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   125
(* Generic decomp for reification : matches the actual term with the
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   126
   rhs of one cong rule. The result of the matching guides the
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   127
   proof synthesis: The matches of the introduced Variables A1 .. An are
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   128
   processed recursively
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   129
   The rest is instantiated in the cong rule,i.e. no reification is needed *)
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   130
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   131
(* da is the decomposition for atoms, ie. it returns ([],g) where g
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   132
   returns the right instance f (AtC n) = t , where AtC is the Atoms
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   133
   constructor and n is the number of the atom corresponding to t *)
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   134
fun decomp_reify da cgns (ct, ctxt) bds =
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   135
  let
52288
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   136
    val thy = Proof_Context.theory_of ctxt;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   137
    val cert = cterm_of thy;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   138
    val certT = ctyp_of thy;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   139
    fun tryabsdecomp (ct, ctxt) bds =
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   140
      (case Thm.term_of ct of
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   141
        Abs (_, xT, ta) =>
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   142
          let
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   143
            val ([raw_xn], ctxt') = Variable.variant_fixes ["x"] ctxt;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   144
            val (xn, ta) = Syntax_Trans.variant_abs (raw_xn, xT, ta);  (* FIXME !? *)
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   145
            val x = Free (xn, xT);
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   146
            val cx = cert x;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   147
            val cta = cert ta;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   148
            val bds = (case AList.lookup Type.could_unify bds (HOLogic.listT xT) of
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   149
                NONE => error "tryabsdecomp: Type not found in the Environement"
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   150
              | SOME (bsT, atsT) => AList.update Type.could_unify (HOLogic.listT xT,
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   151
                  (x :: bsT, atsT)) bds);
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   152
           in (([(cta, ctxt')],
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   153
                fn ([th], bds) =>
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   154
                  (hd (Variable.export ctxt' ctxt [(Thm.forall_intr cx th) COMP allI]),
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   155
                   let
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   156
                     val (bsT, asT) = the (AList.lookup Type.could_unify bds (HOLogic.listT xT));
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   157
                   in
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   158
                     AList.update Type.could_unify (HOLogic.listT xT, (tl bsT, asT)) bds
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   159
                   end)),
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   160
               bds)
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   161
           end
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   162
       | _ => da (ct, ctxt) bds)
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   163
  in
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   164
    (case cgns of
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   165
      [] => tryabsdecomp (ct, ctxt) bds
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   166
    | ((vns, cong) :: congs) =>
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   167
        (let
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   168
          val (tyenv, tmenv) =
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   169
            Pattern.match thy
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   170
              ((fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) (concl_of cong), Thm.term_of ct)
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   171
              (Vartab.empty, Vartab.empty);
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   172
          val (fnvs, invs) = List.partition (fn ((vn, _),_) => member (op =) vns vn) (Vartab.dest tmenv);
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   173
          val (fts, its) =
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   174
            (map (snd o snd) fnvs,
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   175
             map (fn ((vn, vi), (tT, t)) => (cert (Var ((vn, vi), tT)), cert t)) invs);
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   176
          val ctyenv = map (fn ((vn, vi), (s, ty)) => (certT (TVar((vn, vi), s)), certT ty)) (Vartab.dest tyenv);
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   177
        in
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   178
          ((map cert fts ~~ replicate (length fts) ctxt,
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   179
             apfst (FWD (Drule.instantiate_normalize (ctyenv, its) cong))), bds)
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   180
        end handle Pattern.MATCH => decomp_reify da congs (ct, ctxt) bds))
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   181
  end;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   182
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   183
fun get_nths (t as (Const (@{const_name "List.nth"}, _) $ vs $ n)) =
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   184
      AList.update (op aconv) (t, (vs, n))
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   185
  | get_nths (t1 $ t2) = get_nths t1 #> get_nths t2
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   186
  | get_nths (Abs (_, _, t')) = get_nths t'
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   187
  | get_nths _ = I;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   188
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   189
fun tryeqs [] (ct, ctxt) bds = error "Cannot find the atoms equation"
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   190
  | tryeqs (eq :: eqs) (ct, ctxt) bds = ((
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   191
      let
52288
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   192
        val rhs = eq |> prop_of |> HOLogic.dest_Trueprop  |> HOLogic.dest_eq |> snd;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   193
        val nths = get_nths rhs [];
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   194
        val (vss, _) = fold_rev (fn (_, (vs, n)) => fn (vss, ns) =>
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   195
          (insert (op aconv) vs vss, insert (op aconv) n ns)) nths ([], []);
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   196
        val (vsns, ctxt') = Variable.variant_fixes (replicate (length vss) "vs") ctxt;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   197
        val (xns, ctxt'') = Variable.variant_fixes (replicate (length nths) "x") ctxt';
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   198
        val thy = Proof_Context.theory_of ctxt'';
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   199
        val cert = cterm_of thy;
52273
haftmann
parents: 52272
diff changeset
   200
        val certT = ctyp_of thy;
52288
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   201
        val vsns_map = vss ~~ vsns;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   202
        val xns_map = fst (split_list nths) ~~ xns;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   203
        val subst = map (fn (nt, xn) => (nt, Var ((xn, 0), fastype_of nt))) xns_map;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   204
        val rhs_P = subst_free subst rhs;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   205
        val (tyenv, tmenv) = Pattern.match thy (rhs_P, Thm.term_of ct) (Vartab.empty, Vartab.empty);
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   206
        val sbst = Envir.subst_term (tyenv, tmenv);
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   207
        val sbsT = Envir.subst_type tyenv;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   208
        val subst_ty = map (fn (n, (s, t)) =>
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   209
          (certT (TVar (n, s)), certT t)) (Vartab.dest tyenv)
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   210
        val tml = Vartab.dest tmenv;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   211
        val (subst_ns, bds) = fold_map
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   212
          (fn (Const _ $ _ $ n, Var (xn0, _)) => fn bds =>
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   213
            let
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   214
              val name = snd (the (AList.lookup (op =) tml xn0));
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   215
              val (idx, bds) = index_of name bds;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   216
            in ((cert n, idx |> (HOLogic.mk_nat #> cert)), bds) end) subst bds;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   217
        val subst_vs =
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   218
          let
52288
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   219
            fun h (Const _ $ (vs as Var (_, lT)) $ _, Var (_, T)) =
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   220
              let
52288
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   221
                val cns = sbst (Const (@{const_name "List.Cons"}, T --> lT --> lT));
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   222
                val lT' = sbsT lT;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   223
                val (bsT, _) = the (AList.lookup Type.could_unify bds lT);
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   224
                val vsn = the (AList.lookup (op =) vsns_map vs);
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   225
                val cvs = cert (fold_rev (fn x => fn xs => cns $ x $xs) bsT (Free (vsn, lT')));
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   226
              in (cert vs, cvs) end;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   227
          in map h subst end;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   228
        val cts = map (fn ((vn, vi), (tT, t)) => (cert (Var ((vn, vi), tT)), cert t))
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   229
          (fold (AList.delete (fn (((a : string), _), (b, _)) => a = b))
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   230
            (map (fn n => (n, 0)) xns) tml);
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   231
        val substt =
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   232
          let
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   233
            val ih = Drule.cterm_rule (Thm.instantiate (subst_ty, []));
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   234
          in map (pairself ih) (subst_ns @ subst_vs @ cts) end;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   235
        val th = (Drule.instantiate_normalize (subst_ty, substt) eq) RS sym;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   236
      in (hd (Variable.export ctxt'' ctxt [th]), bds) end)
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   237
      handle Pattern.MATCH => tryeqs eqs (ct, ctxt) bds);
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   238
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   239
(* looks for the atoms equation and instantiates it with the right number *)
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   240
52288
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   241
fun mk_decompatom eqs (ct, ctxt) bds = (([], fn (_, bds) =>
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   242
  let
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   243
    val tT = fastype_of (Thm.term_of ct);
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   244
    fun isat eq =
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   245
      let
52288
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   246
        val rhs = eq |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   247
      in exists_Const
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   248
        (fn (n, ty) => n = @{const_name "List.nth"}
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   249
          andalso AList.defined Type.could_unify bds (domain_type ty)) rhs
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   250
          andalso Type.could_unify (fastype_of rhs, tT)
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   251
      end;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   252
  in tryeqs (filter isat eqs) (ct, ctxt) bds end), bds);
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   253
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   254
(* Generic reification procedure: *)
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   255
(* creates all needed cong rules and then just uses the theorem synthesis *)
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   256
52288
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   257
fun mk_congs ctxt eqs =
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   258
  let
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   259
    val fs = fold_rev (fn eq => insert (op =) (eq |> prop_of |> HOLogic.dest_Trueprop
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   260
      |> HOLogic.dest_eq |> fst |> strip_comb
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   261
      |> fst)) eqs [];
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   262
    val tys = fold_rev (fn f => fold (insert (op =)) (f |> fastype_of |> binder_types |> tl)) fs [];
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   263
    val (vs, ctxt') = Variable.variant_fixes (replicate (length tys) "vs") ctxt;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   264
    val cert = cterm_of (Proof_Context.theory_of ctxt');
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   265
    val subst =
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   266
      the o AList.lookup (op =) (map2 (fn T => fn v => (T, cert (Free (v, T)))) tys vs);
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   267
    fun prep_eq eq =
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   268
      let
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   269
        val (_, _ :: vs) = eq |> prop_of |> HOLogic.dest_Trueprop
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   270
          |> HOLogic.dest_eq |> fst |> strip_comb;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   271
        val subst = map_filter (fn (v as Var (_, T)) => SOME (cert v, subst T)
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   272
          | _ => NONE) vs;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   273
      in Thm.instantiate ([], subst) eq end;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   274
    val (ps, congs) = map_split (mk_congeq ctxt' fs o prep_eq) eqs;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   275
    val bds = AList.make (K ([], [])) tys;
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   276
  in (ps ~~ Variable.export ctxt' ctxt congs, bds) end
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   277
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   278
fun conv ctxt eqs ct =
ca4932dad084 denesting of functions
haftmann
parents: 52286
diff changeset
   279
  let
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   280
    val (congs, bds) = mk_congs ctxt eqs;
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   281
    val congs = rearrange congs;
52275
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
   282
    val (th, bds') = apfst mk_eq (divide_and_conquer' (decomp_reify (mk_decompatom eqs) congs) (ct, ctxt) bds);
52273
haftmann
parents: 52272
diff changeset
   283
    fun is_list_var (Var (_, t)) = can dest_listT t
haftmann
parents: 52272
diff changeset
   284
      | is_list_var _ = false;
52275
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
   285
    val vars = th |> prop_of |> Logic.dest_equals |> snd
52273
haftmann
parents: 52272
diff changeset
   286
      |> strip_comb |> snd |> filter is_list_var;
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   287
    val cert = cterm_of (Proof_Context.theory_of ctxt);
52275
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
   288
    val vs = map (fn v as Var (_, T) =>
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
   289
      (v, the (AList.lookup Type.could_unify bds' T) |> snd |> HOLogic.mk_list (dest_listT T))) vars;
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
   290
    val th' = Drule.instantiate_normalize ([], (map o pairself) cert vs) th;
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
   291
    val th'' = Thm.symmetric (dereify ctxt [] (Thm.lhs_of th'));
9b4c04da53b1 decomposed reflection steps into conversions;
haftmann
parents: 52274
diff changeset
   292
  in Thm.transitive th'' th' end;
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   293
52286
8170e5327c02 make reification part of HOL
haftmann
parents: 52276
diff changeset
   294
fun tac ctxt eqs =
8170e5327c02 make reification part of HOL
haftmann
parents: 52276
diff changeset
   295
  lift_conv ctxt (conv ctxt eqs);
51726
b3e599b5ecc8 reflection as official HOL tool
haftmann
parents:
diff changeset
   296
52276
329c41438154 reflection without evaluation
haftmann
parents: 52275
diff changeset
   297
end;