src/HOL/Library/reflection.ML
author bulwahn
Mon, 25 Jul 2011 11:21:44 +0200
changeset 44830 285ffb18da30
parent 44215 2bdec7f430d3
child 44831 c2554cc82d34
permissions -rw-r--r--
fixed typo
     1 (*  Title:      HOL/Library/reflection.ML
     2     Author:     Amine Chaieb, TU Muenchen
     3 
     4 A trial for automatical reification.
     5 *)
     6 
     7 signature REFLECTION =
     8 sig
     9   val genreify_tac: Proof.context -> thm list -> term option -> int -> tactic
    10   val reflection_tac: Proof.context -> thm list -> thm list -> term option -> int -> tactic
    11   val gen_reflection_tac: Proof.context -> (cterm -> thm)
    12     -> thm list -> thm list -> term option -> int -> tactic
    13   val genreif : Proof.context -> thm list -> term -> thm
    14 end;
    15 
    16 structure Reflection : REFLECTION =
    17 struct
    18 
    19 val ext2 = @{thm ext2};
    20 val nth_Cons_0 = @{thm nth_Cons_0};
    21 val nth_Cons_Suc = @{thm nth_Cons_Suc};
    22 
    23   (* Make a congruence rule out of a defining equation for the interpretation *)
    24   (* th is one defining equation of f, i.e.
    25      th is "f (Cp ?t1 ... ?tn) = P(f ?t1, .., f ?tn)" *)
    26   (* Cp is a constructor pattern and P is a pattern *)
    27 
    28   (* The result is:
    29       [|?A1 = f ?t1 ; .. ; ?An= f ?tn |] ==> P (?A1, .., ?An) = f (Cp ?t1 .. ?tn) *)
    30   (*  + the a list of names of the A1 .. An, Those are fresh in the ctxt*)
    31 
    32 fun mk_congeq ctxt fs th =
    33   let
    34    val (f as Const(fN,fT)) = th |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq
    35      |> fst |> strip_comb |> fst
    36    val thy = Proof_Context.theory_of ctxt
    37    val cert = Thm.cterm_of thy
    38    val (((_,_),[th']), ctxt') = Variable.import true [th] ctxt
    39    val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of th'))
    40    fun add_fterms (t as t1 $ t2) =
    41        if exists (fn f => Term.could_unify (t |> strip_comb |> fst, f)) fs then insert (op aconv) t
    42        else add_fterms t1 #> add_fterms t2
    43      | add_fterms (t as Abs(xn,xT,t')) =
    44        if exists_Const (fn (c, _) => c = fN) t then (fn _ => [t]) else (fn _ => [])
    45      | add_fterms _ = I
    46    val fterms = add_fterms rhs []
    47    val (xs, ctxt'') = Variable.variant_fixes (replicate (length fterms) "x") ctxt'
    48    val tys = map fastype_of fterms
    49    val vs = map Free (xs ~~ tys)
    50    val env = fterms ~~ vs
    51                     (* FIXME!!!!*)
    52    fun replace_fterms (t as t1 $ t2) =
    53        (case AList.lookup (op aconv) env t of
    54             SOME v => v
    55           | NONE => replace_fterms t1 $ replace_fterms t2)
    56      | replace_fterms t = (case AList.lookup (op aconv) env t of
    57                                SOME v => v
    58                              | NONE => t)
    59 
    60    fun mk_def (Abs(x,xT,t),v) = HOLogic.mk_Trueprop ((HOLogic.all_const xT)$ Abs(x,xT,HOLogic.mk_eq(v$(Bound 0), t)))
    61      | mk_def (t, v) = HOLogic.mk_Trueprop (HOLogic.mk_eq (v, t))
    62    fun tryext x = (x RS ext2 handle THM _ =>  x)
    63    val cong = (Goal.prove ctxt'' [] (map mk_def env)
    64                           (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, replace_fterms rhs)))
    65                           (fn x => Local_Defs.unfold_tac (#context x) (map tryext (#prems x))
    66                                                         THEN rtac th' 1)) RS sym
    67 
    68    val (cong' :: vars') =
    69        Variable.export ctxt'' ctxt (cong :: map (Drule.mk_term o cert) vs)
    70    val vs' = map (fst o fst o Term.dest_Var o Thm.term_of o Drule.dest_term) vars'
    71 
    72   in  (vs', cong') end;
    73  (* congs is a list of pairs (P,th) where th is a theorem for *)
    74         (* [| f p1 = A1; ...; f pn = An|] ==> f (C p1 .. pn) = P *)
    75 val FWD = curry (op OF);
    76 
    77 
    78 exception REIF of string;
    79 
    80 fun dest_listT (Type (@{type_name "list"}, [T])) = T;
    81 
    82 fun rearrange congs =
    83   let
    84     fun P (_, th) =
    85       let val @{term "Trueprop"}$(Const (@{const_name HOL.eq},_) $l$_) = concl_of th
    86       in can dest_Var l end
    87     val (yes,no) = List.partition P congs
    88   in no @ yes end
    89 
    90 fun genreif ctxt raw_eqs t =
    91   let
    92     fun index_of t bds =
    93       let
    94         val tt = HOLogic.listT (fastype_of t)
    95       in
    96        (case AList.lookup Type.could_unify bds tt of
    97           NONE => error "index_of : type not found in environements!"
    98         | SOME (tbs,tats) =>
    99           let
   100             val i = find_index (fn t' => t' = t) tats
   101             val j = find_index (fn t' => t' = t) tbs
   102           in (if j = ~1 then
   103               if i = ~1
   104               then (length tbs + length tats,
   105                     AList.update Type.could_unify (tt,(tbs,tats@[t])) bds)
   106               else (i, bds) else (j, bds))
   107           end)
   108       end;
   109 
   110     (* Generic decomp for reification : matches the actual term with the
   111        rhs of one cong rule. The result of the matching guides the
   112        proof synthesis: The matches of the introduced Variables A1 .. An are
   113        processed recursively
   114        The rest is instantiated in the cong rule,i.e. no reification is needed *)
   115 
   116     (* da is the decomposition for atoms, ie. it returns ([],g) where g
   117        returns the right instance f (AtC n) = t , where AtC is the Atoms
   118        constructor and n is the number of the atom corresponding to t *)
   119     fun decomp_genreif da cgns (t,ctxt) bds =
   120       let
   121         val thy = Proof_Context.theory_of ctxt
   122         val cert = cterm_of thy
   123         fun tryabsdecomp (s,ctxt) bds =
   124           (case s of
   125              Abs(xn,xT,ta) => (
   126                let
   127                  val ([xn],ctxt') = Variable.variant_fixes ["x"] ctxt
   128                  val (xn,ta) = Syntax_Trans.variant_abs (xn,xT,ta)  (* FIXME !? *)
   129                  val x = Free(xn,xT)
   130                  val bds = (case AList.lookup Type.could_unify bds (HOLogic.listT xT)
   131                           of NONE => error "tryabsdecomp: Type not found in the Environement"
   132                            | SOME (bsT,atsT) =>
   133                              (AList.update Type.could_unify (HOLogic.listT xT, ((x::bsT), atsT)) bds))
   134                in (([(ta, ctxt')],
   135                     fn ([th], bds) =>
   136                       (hd (Variable.export ctxt' ctxt [(Thm.forall_intr (cert x) th) COMP allI]),
   137                        let val (bsT,asT) = the(AList.lookup Type.could_unify bds (HOLogic.listT xT))
   138                        in AList.update Type.could_unify (HOLogic.listT xT,(tl bsT,asT)) bds
   139                        end)),
   140                    bds)
   141                end)
   142            | _ => da (s,ctxt) bds)
   143       in
   144         (case cgns of
   145           [] => tryabsdecomp (t,ctxt) bds
   146         | ((vns,cong)::congs) =>
   147             (let
   148               val cert = cterm_of thy
   149               val certy = ctyp_of thy
   150               val (tyenv, tmenv) =
   151                 Pattern.match thy
   152                   ((fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) (concl_of cong), t)
   153                   (Vartab.empty, Vartab.empty)
   154               val (fnvs,invs) = List.partition (fn ((vn,_),_) => member (op =) vns vn) (Vartab.dest tmenv)
   155               val (fts,its) =
   156                 (map (snd o snd) fnvs,
   157                  map (fn ((vn,vi),(tT,t)) => (cert(Var ((vn,vi),tT)), cert t)) invs)
   158               val ctyenv = map (fn ((vn,vi),(s,ty)) => (certy (TVar((vn,vi),s)), certy ty)) (Vartab.dest tyenv)
   159             in ((fts ~~ (replicate (length fts) ctxt),
   160                  Library.apfst (FWD (Drule.instantiate_normalize (ctyenv, its) cong))), bds)
   161             end handle Pattern.MATCH => decomp_genreif da congs (t,ctxt) bds))
   162       end;
   163 
   164  (* looks for the atoms equation and instantiates it with the right number *)
   165     fun mk_decompatom eqs (t,ctxt) bds = (([], fn (_, bds) =>
   166       let
   167         val tT = fastype_of t
   168         fun isat eq =
   169           let
   170             val rhs = eq |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd
   171           in exists_Const
   172             (fn (n,ty) => n = @{const_name "List.nth"}
   173                           andalso
   174                           AList.defined Type.could_unify bds (domain_type ty)) rhs
   175             andalso Type.could_unify (fastype_of rhs, tT)
   176           end
   177 
   178         fun get_nths t acc =
   179           case t of
   180             Const(@{const_name "List.nth"},_)$vs$n => insert (fn ((a,_),(b,_)) => a aconv b) (t,(vs,n)) acc
   181           | t1$t2 => get_nths t1 (get_nths t2 acc)
   182           | Abs(_,_,t') => get_nths t'  acc
   183           | _ => acc
   184 
   185         fun
   186            tryeqs [] bds = error "Can not find the atoms equation"
   187          | tryeqs (eq::eqs) bds = ((
   188           let
   189             val rhs = eq |> prop_of |> HOLogic.dest_Trueprop  |> HOLogic.dest_eq |> snd
   190             val nths = get_nths rhs []
   191             val (vss,ns) = fold_rev (fn (_,(vs,n)) => fn (vss,ns) =>
   192                                       (insert (op aconv) vs vss, insert (op aconv) n ns)) nths ([],[])
   193             val (vsns, ctxt') = Variable.variant_fixes (replicate (length vss) "vs") ctxt
   194             val (xns, ctxt'') = Variable.variant_fixes (replicate (length nths) "x") ctxt'
   195             val thy = Proof_Context.theory_of ctxt''
   196             val cert = cterm_of thy
   197             val certT = ctyp_of thy
   198             val vsns_map = vss ~~ vsns
   199             val xns_map = (fst (split_list nths)) ~~ xns
   200             val subst = map (fn (nt, xn) => (nt, Var ((xn,0), fastype_of nt))) xns_map
   201             val rhs_P = subst_free subst rhs
   202             val (tyenv, tmenv) = Pattern.match thy (rhs_P, t) (Vartab.empty, Vartab.empty)
   203             val sbst = Envir.subst_term (tyenv, tmenv)
   204             val sbsT = Envir.subst_type tyenv
   205             val subst_ty = map (fn (n,(s,t)) => (certT (TVar (n, s)), certT t))
   206                                (Vartab.dest tyenv)
   207             val tml = Vartab.dest tmenv
   208             val t's = map (fn xn => snd (the (AList.lookup (op =) tml (xn,0)))) xns (* FIXME : Express with sbst*)
   209             val (subst_ns, bds) = fold_map
   210                 (fn (Const _ $ vs $ n, Var (xn0,T)) => fn bds =>
   211                   let
   212                     val name = snd (the (AList.lookup (op =) tml xn0))
   213                     val (idx, bds) = index_of name bds
   214                   in ((cert n, idx |> (HOLogic.mk_nat #> cert)), bds) end) subst bds
   215             val subst_vs =
   216               let
   217                 fun ty (Const _ $ (vs as Var (vsn,lT)) $ n, Var (xn0,T)) = (certT T, certT (sbsT T))
   218                 fun h (Const _ $ (vs as Var (vsn,lT)) $ n, Var (xn0,T)) =
   219                   let
   220                     val cns = sbst (Const(@{const_name "List.Cons"}, T --> lT --> lT))
   221                     val lT' = sbsT lT
   222                     val (bsT,asT) = the (AList.lookup Type.could_unify bds lT)
   223                     val vsn = the (AList.lookup (op =) vsns_map vs)
   224                     val cvs = cert (fold_rev (fn x => fn xs => cns$x$xs) bsT (Free (vsn, lT')))
   225                   in (cert vs, cvs) end
   226               in map h subst end
   227             val cts = map (fn ((vn,vi),(tT,t)) => (cert(Var ((vn,vi),tT)), cert t))
   228                           (fold (AList.delete (fn (((a: string),_),(b,_)) => a = b))
   229                                 (map (fn n => (n,0)) xns) tml)
   230             val substt =
   231               let val ih = Drule.cterm_rule (Thm.instantiate (subst_ty,[]))
   232               in map (fn (v,t) => (ih v, ih t)) (subst_ns@subst_vs@cts)  end
   233             val th = (Drule.instantiate_normalize (subst_ty, substt)  eq) RS sym
   234           in (hd (Variable.export ctxt'' ctxt [th]), bds) end)
   235           handle Pattern.MATCH => tryeqs eqs bds)
   236       in tryeqs (filter isat eqs) bds end), bds);
   237 
   238   (* Generic reification procedure: *)
   239   (* creates all needed cong rules and then just uses the theorem synthesis *)
   240 
   241     fun mk_congs ctxt raw_eqs =
   242       let
   243         val fs = fold_rev (fn eq =>
   244                            insert (op =) (eq |> prop_of |> HOLogic.dest_Trueprop
   245                            |> HOLogic.dest_eq |> fst |> strip_comb
   246                            |> fst)) raw_eqs []
   247         val tys = fold_rev (fn f => fold (insert (op =)) (f |> fastype_of |> binder_types |> tl)
   248                             ) fs []
   249         val (vs, ctxt') = Variable.variant_fixes (replicate (length tys) "vs") ctxt
   250         val thy = Proof_Context.theory_of ctxt'
   251         val cert = cterm_of thy
   252         val vstys = map (fn (t,v) => (t,SOME (cert (Free(v,t)))))
   253                     (tys ~~ vs)
   254         val is_Var = can dest_Var
   255         fun insteq eq vs =
   256           let
   257             val subst = map (fn (v as Var(n,t)) => (cert v, (the o the) (AList.lookup (op =) vstys t)))
   258                         (filter is_Var vs)
   259           in Thm.instantiate ([],subst) eq
   260           end
   261 
   262         val bds = AList.make (fn _ => ([],[])) tys
   263         val eqs = map (fn eq => eq |> prop_of |> HOLogic.dest_Trueprop
   264                                    |> HOLogic.dest_eq |> fst |> strip_comb |> snd |> tl
   265                                    |> (insteq eq)) raw_eqs
   266         val (ps,congs) = split_list (map (mk_congeq ctxt' fs) eqs)
   267       in (ps ~~ (Variable.export ctxt' ctxt congs), bds)
   268       end
   269 
   270     val (congs, bds) = mk_congs ctxt raw_eqs
   271     val congs = rearrange congs
   272     val (th, bds) = divide_and_conquer' (decomp_genreif (mk_decompatom raw_eqs) congs) (t,ctxt) bds
   273     fun is_listVar (Var (_,t)) = can dest_listT t
   274          | is_listVar _ = false
   275     val vars = th |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd
   276                   |> strip_comb |> snd |> filter is_listVar
   277     val cert = cterm_of (Proof_Context.theory_of ctxt)
   278     val cvs = map (fn (v as Var(n,t)) => (cert v,
   279                   the (AList.lookup Type.could_unify bds t) |> snd |> HOLogic.mk_list (dest_listT t) |> cert)) vars
   280     val th' = Drule.instantiate_normalize ([], cvs) th
   281     val t' = (fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) th'
   282     val th'' = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, t')))
   283                (fn _ => simp_tac (simpset_of ctxt) 1)
   284   in FWD trans [th'',th']
   285   end
   286 
   287 
   288 fun genreflect ctxt conv corr_thms raw_eqs t =
   289   let
   290     val reifth = genreif ctxt raw_eqs t
   291     fun trytrans [] = error "No suitable correctness theorem found"
   292       | trytrans (th::ths) =
   293            (FWD trans [reifth, th RS sym] handle THM _ => trytrans ths)
   294     val th = trytrans corr_thms
   295     val ft = (Thm.dest_arg1 o Thm.dest_arg o Thm.dest_arg o cprop_of) th
   296     val rth = conv ft
   297   in simplify (HOL_basic_ss addsimps raw_eqs addsimps [nth_Cons_0, nth_Cons_Suc])
   298              (simplify (HOL_basic_ss addsimps [rth]) th)
   299   end
   300 
   301 fun genreify_tac ctxt eqs to = SUBGOAL (fn (goal, i) =>
   302   let
   303     val t = (case to of NONE => HOLogic.dest_Trueprop goal | SOME x => x)
   304     val th = genreif ctxt eqs t RS ssubst
   305   in rtac th i end);
   306 
   307     (* Reflection calls reification and uses the correctness *)
   308         (* theorem assumed to be the head of the list *)
   309 fun gen_reflection_tac ctxt conv corr_thms raw_eqs to = SUBGOAL (fn (goal, i) =>
   310   let
   311     val t = (case to of NONE => HOLogic.dest_Trueprop goal | SOME x => x)
   312     val th = genreflect ctxt conv corr_thms raw_eqs t RS ssubst
   313   in rtac th i THEN TRY (rtac TrueI i) end);  (* FIXME THEN_ALL_NEW !? *)
   314 
   315 fun reflection_tac ctxt = gen_reflection_tac ctxt (Codegen.evaluation_conv ctxt);
   316   (*FIXME why Codegen.evaluation_conv?  very specific...*)
   317 
   318 end