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