src/HOL/Library/reflection.ML
changeset 29650 cc3958d31b1d
parent 29273 285c00993bc2
child 29805 a5da150bd0ab
equal deleted inserted replaced
29648:ead544f3d6a1 29650:cc3958d31b1d
       
     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 
       
    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 = ProofContext.theory_of ctxt
       
    37    val cert = Thm.cterm_of thy
       
    38    val (((_,_),[th']), ctxt') = Variable.import_thms 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 => LocalDefs.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  (* da is the decomposition for atoms, ie. it returns ([],g) where g
       
    78  returns the right instance f (AtC n) = t , where AtC is the Atoms
       
    79  constructor and n is the number of the atom corresponding to t *)
       
    80 
       
    81 (* Generic decomp for reification : matches the actual term with the
       
    82 rhs of one cong rule. The result of the matching guides the
       
    83 proof synthesis: The matches of the introduced Variables A1 .. An are
       
    84 processed recursively
       
    85  The rest is instantiated in the cong rule,i.e. no reification is needed *)
       
    86 
       
    87 exception REIF of string;
       
    88 
       
    89 val bds = ref ([]: (typ * ((term list) * (term list))) list);
       
    90 
       
    91 fun index_of t = 
       
    92  let 
       
    93   val tt = HOLogic.listT (fastype_of t)
       
    94  in 
       
    95   (case AList.lookup Type.could_unify (!bds) tt of
       
    96     NONE => error "index_of : type not found in environements!"
       
    97   | SOME (tbs,tats) =>
       
    98     let
       
    99      val i = find_index_eq t tats
       
   100      val j = find_index_eq t tbs 
       
   101     in (if j= ~1 then 
       
   102 	    if i= ~1 
       
   103 	    then (bds := AList.update Type.could_unify (tt,(tbs,tats@[t])) (!bds) ; 
       
   104 		  length tbs + length tats) 
       
   105 	    else i else j)
       
   106     end)
       
   107  end;
       
   108 
       
   109 fun dest_listT (Type ("List.list", [T])) = T;
       
   110 
       
   111 fun decomp_genreif da cgns (t,ctxt) =
       
   112  let 
       
   113   val thy = ProofContext.theory_of ctxt 
       
   114   val cert = cterm_of thy
       
   115   fun tryabsdecomp (s,ctxt) = 
       
   116    (case s of 
       
   117      Abs(xn,xT,ta) => 
       
   118      (let
       
   119        val ([xn],ctxt') = Variable.variant_fixes ["x"] ctxt
       
   120        val (xn,ta) = variant_abs (xn,xT,ta)
       
   121        val x = Free(xn,xT)
       
   122        val _ = (case AList.lookup Type.could_unify (!bds) (HOLogic.listT xT)
       
   123 		 of NONE => error "tryabsdecomp: Type not found in the Environement"
       
   124 		  | SOME (bsT,atsT) => 
       
   125 		    (bds := AList.update Type.could_unify (HOLogic.listT xT, ((x::bsT), atsT)) (!bds)))
       
   126       in ([(ta, ctxt')] , 
       
   127 	  fn [th] => ((let val (bsT,asT) = the(AList.lookup Type.could_unify (!bds) (HOLogic.listT xT))
       
   128 		       in (bds := AList.update Type.could_unify (HOLogic.listT xT,(tl bsT,asT)) (!bds))
       
   129 		       end) ; 
       
   130 		      hd (Variable.export ctxt' ctxt [(forall_intr (cert x) th) COMP allI])))
       
   131 	end)
       
   132     | _ => da (s,ctxt))
       
   133   in 
       
   134   (case cgns of 
       
   135     [] => tryabsdecomp (t,ctxt)
       
   136   | ((vns,cong)::congs) => ((let
       
   137         val cert = cterm_of thy
       
   138 	val certy = ctyp_of thy
       
   139         val (tyenv, tmenv) =
       
   140         Pattern.match thy
       
   141         ((fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) (concl_of cong), t)
       
   142         (Envir.type_env (Envir.empty 0), Vartab.empty)
       
   143         val (fnvs,invs) = List.partition (fn ((vn,_),_) => vn mem vns) (Vartab.dest tmenv)
       
   144         val (fts,its) = 
       
   145 	    (map (snd o snd) fnvs,
       
   146              map (fn ((vn,vi),(tT,t)) => (cert(Var ((vn,vi),tT)), cert t)) invs)
       
   147 	val ctyenv = map (fn ((vn,vi),(s,ty)) => (certy (TVar((vn,vi),s)), certy ty)) (Vartab.dest tyenv)
       
   148     in (fts ~~ (replicate (length fts) ctxt), FWD (instantiate (ctyenv, its) cong))
       
   149     end)
       
   150       handle MATCH => decomp_genreif da congs (t,ctxt)))
       
   151   end;
       
   152 
       
   153  (* looks for the atoms equation and instantiates it with the right number *)
       
   154 
       
   155 
       
   156 fun mk_decompatom eqs (t,ctxt) =
       
   157 let 
       
   158  val tT = fastype_of t
       
   159  fun isat eq = 
       
   160   let 
       
   161    val rhs = eq |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd
       
   162    in exists_Const 
       
   163 	  (fn (n,ty) => n="List.nth" 
       
   164 			andalso 
       
   165 			AList.defined Type.could_unify (!bds) (domain_type ty)) rhs 
       
   166 	  andalso Type.could_unify (fastype_of rhs, tT)
       
   167    end
       
   168  fun get_nths t acc = 
       
   169   case t of
       
   170     Const("List.nth",_)$vs$n => insert (fn ((a,_),(b,_)) => a aconv b) (t,(vs,n)) acc
       
   171   | t1$t2 => get_nths t1 (get_nths t2 acc)
       
   172   | Abs(_,_,t') => get_nths t'  acc
       
   173   | _ => acc
       
   174 
       
   175  fun 
       
   176    tryeqs [] = error "Can not find the atoms equation"
       
   177  | tryeqs (eq::eqs) = ((
       
   178   let 
       
   179    val rhs = eq |> prop_of |> HOLogic.dest_Trueprop  |> HOLogic.dest_eq |> snd
       
   180    val nths = get_nths rhs []
       
   181    val (vss,ns) = fold_rev (fn (_,(vs,n)) => fn (vss,ns) => 
       
   182                              (insert (op aconv) vs vss, insert (op aconv) n ns)) nths ([],[]) 
       
   183    val (vsns, ctxt') = Variable.variant_fixes (replicate (length vss) "vs") ctxt
       
   184    val (xns, ctxt'') = Variable.variant_fixes (replicate (length nths) "x") ctxt' 
       
   185    val thy = ProofContext.theory_of ctxt''
       
   186    val cert = cterm_of thy
       
   187    val certT = ctyp_of thy
       
   188    val vsns_map = vss ~~ vsns
       
   189    val xns_map = (fst (split_list nths)) ~~ xns
       
   190    val subst = map (fn (nt, xn) => (nt, Var ((xn,0), fastype_of nt))) xns_map
       
   191    val rhs_P = subst_free subst rhs
       
   192    val (tyenv, tmenv) = Pattern.match 
       
   193 	                    thy (rhs_P, t)
       
   194 	                    (Envir.type_env (Envir.empty 0), Vartab.empty)
       
   195    val sbst = Envir.subst_vars (tyenv, tmenv)
       
   196    val sbsT = Envir.typ_subst_TVars tyenv
       
   197    val subst_ty = map (fn (n,(s,t)) => (certT (TVar (n, s)), certT t)) 
       
   198                       (Vartab.dest tyenv)
       
   199    val tml = Vartab.dest tmenv
       
   200    val t's = map (fn xn => snd (valOf (AList.lookup (op =) tml (xn,0)))) xns (* FIXME : Express with sbst*)
       
   201    val subst_ns = map (fn (Const _ $ vs $ n, Var (xn0,T)) => 
       
   202                           (cert n, snd (valOf (AList.lookup (op =) tml xn0)) 
       
   203                              |> (index_of #> HOLogic.mk_nat #> cert))) 
       
   204                       subst
       
   205    val subst_vs = 
       
   206     let 
       
   207      fun ty (Const _ $ (vs as Var (vsn,lT)) $ n, Var (xn0,T)) = (certT T, certT (sbsT T))
       
   208      fun h (Const _ $ (vs as Var (vsn,lT)) $ n, Var (xn0,T)) = 
       
   209       let 
       
   210        val cns = sbst (Const("List.list.Cons", T --> lT --> lT))
       
   211        val lT' = sbsT lT
       
   212        val (bsT,asT) = the (AList.lookup Type.could_unify (!bds) lT)
       
   213        val vsn = valOf (AList.lookup (op =) vsns_map vs)
       
   214        val cvs = cert (fold_rev (fn x => fn xs => cns$x$xs) bsT (Free (vsn, lT')))
       
   215       in (cert vs, cvs) end
       
   216     in map h subst end
       
   217    val cts = map (fn ((vn,vi),(tT,t)) => (cert(Var ((vn,vi),tT)), cert t)) 
       
   218                  (fold (AList.delete (fn (((a: string),_),(b,_)) => a = b)) 
       
   219                        (map (fn n => (n,0)) xns) tml)
       
   220    val substt = 
       
   221     let val ih = Drule.cterm_rule (Thm.instantiate (subst_ty,[]))
       
   222     in map (fn (v,t) => (ih v, ih t)) (subst_ns@subst_vs@cts)  end
       
   223    val th = (instantiate (subst_ty, substt)  eq) RS sym
       
   224   in  hd (Variable.export ctxt'' ctxt [th]) end)
       
   225  handle MATCH => tryeqs eqs)
       
   226 in ([], fn _ => tryeqs (filter isat eqs))
       
   227 end;
       
   228 
       
   229   (* Generic reification procedure: *)
       
   230   (* creates all needed cong rules and then just uses the theorem synthesis *)
       
   231 
       
   232   fun mk_congs ctxt raw_eqs = 
       
   233  let
       
   234   val fs = fold_rev (fn eq =>
       
   235 		     insert (op =) (eq |> prop_of |> HOLogic.dest_Trueprop 
       
   236 			 |> HOLogic.dest_eq |> fst |> strip_comb 
       
   237 			 |> fst)) raw_eqs []
       
   238   val tys = fold_rev (fn f => fn ts => (f |> fastype_of |> binder_types |> tl) 
       
   239 				    union ts) fs []
       
   240   val _ = bds := AList.make (fn _ => ([],[])) tys
       
   241   val (vs, ctxt') = Variable.variant_fixes (replicate (length tys) "vs") ctxt
       
   242   val thy = ProofContext.theory_of ctxt'
       
   243   val cert = cterm_of thy
       
   244   val vstys = map (fn (t,v) => (t,SOME (cert (Free(v,t))))) 
       
   245 		  (tys ~~ vs)
       
   246   val is_Var = can dest_Var
       
   247   fun insteq eq vs = 
       
   248    let
       
   249      val subst = map (fn (v as Var(n,t)) => (cert v, (valOf o valOf) (AList.lookup (op =) vstys t)))  
       
   250   (filter is_Var vs)
       
   251    in Thm.instantiate ([],subst) eq
       
   252    end
       
   253   val eqs = map (fn eq => eq |> prop_of |> HOLogic.dest_Trueprop 
       
   254 			     |> HOLogic.dest_eq |> fst |> strip_comb |> snd |> tl
       
   255 			     |> (insteq eq)) raw_eqs
       
   256   val (ps,congs) = split_list (map (mk_congeq ctxt' fs) eqs)
       
   257 in ps ~~ (Variable.export ctxt' ctxt congs)
       
   258 end
       
   259 
       
   260 fun partition P [] = ([],[])
       
   261   | partition P (x::xs) = 
       
   262      let val (yes,no) = partition P xs
       
   263      in if P x then (x::yes,no) else (yes, x::no) end
       
   264 
       
   265 fun rearrange congs = 
       
   266 let 
       
   267  fun P (_, th) = 
       
   268   let val @{term "Trueprop"}$(Const ("op =",_) $l$_) = concl_of th
       
   269   in can dest_Var l end
       
   270  val (yes,no) = partition P congs 
       
   271  in no @ yes end
       
   272 
       
   273 
       
   274 
       
   275 fun genreif ctxt raw_eqs t =
       
   276  let 
       
   277   val congs = rearrange (mk_congs ctxt raw_eqs)
       
   278   val th = divide_and_conquer (decomp_genreif (mk_decompatom raw_eqs) congs) (t,ctxt)
       
   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, the (AList.lookup Type.could_unify (!bds) t) |> snd |> HOLogic.mk_list (dest_listT t) |> cert)) vars
       
   285   val th' = instantiate ([], cvs) th
       
   286   val t' = (fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) th'
       
   287   val th'' = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, t')))
       
   288 			(fn _ => simp_tac (local_simpset_of ctxt) 1)
       
   289   val _ = bds := []
       
   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     val P = HOLogic.dest_Trueprop (List.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 
       
   327 end