wenzelm@20319: (* wenzelm@20319: ID: $Id$ wenzelm@20319: Author: Amine Chaieb, TU Muenchen wenzelm@20319: wenzelm@20319: A trial for automatical reification. wenzelm@20319: *) wenzelm@20319: wenzelm@20319: signature REFLECTION = sig wenzelm@20319: val genreify_tac: Proof.context -> thm list -> term option -> int -> tactic wenzelm@20319: val reflection_tac: Proof.context -> thm list -> term option -> int -> tactic haftmann@21878: val gen_reflection_tac: Proof.context -> (cterm -> thm) haftmann@21878: -> thm list -> term option -> int -> tactic wenzelm@20319: end; wenzelm@20319: chaieb@23605: structure Reflection : REFLECTION wenzelm@20319: = struct wenzelm@20319: chaieb@20374: val ext2 = thm "ext2"; wenzelm@21669: val nth_Cons_0 = thm "nth_Cons_0"; wenzelm@21669: val nth_Cons_Suc = thm "nth_Cons_Suc"; wenzelm@21669: chaieb@20374: (* Make a congruence rule out of a defining equation for the interpretation *) chaieb@20374: (* th is one defining equation of f, i.e. chaieb@20374: th is "f (Cp ?t1 ... ?tn) = P(f ?t1, .., f ?tn)" *) chaieb@20374: (* Cp is a constructor pattern and P is a pattern *) chaieb@20374: chaieb@20374: (* The result is: chaieb@20374: [|?A1 = f ?t1 ; .. ; ?An= f ?tn |] ==> P (?A1, .., ?An) = f (Cp ?t1 .. ?tn) *) chaieb@20374: (* + the a list of names of the A1 .. An, Those are fresh in the ctxt*) chaieb@20374: chaieb@20374: chaieb@20374: fun mk_congeq ctxt fs th = chaieb@20374: let chaieb@20564: val (f as Const(fN,fT)) = th |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq chaieb@20564: |> fst |> strip_comb |> fst chaieb@20374: val thy = ProofContext.theory_of ctxt chaieb@20374: val cert = Thm.cterm_of thy wenzelm@22568: val (((_,_),[th']), ctxt') = Variable.import_thms true [th] ctxt chaieb@20374: val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of th')) chaieb@20564: fun add_fterms (t as t1 $ t2) = chaieb@22199: if exists (fn f => could_unify (t |> strip_comb |> fst, f)) fs then insert (op aconv) t chaieb@20564: else add_fterms t1 #> add_fterms t2 chaieb@20374: | add_fterms (t as Abs(xn,xT,t')) = chaieb@20564: if (fN mem (term_consts t)) then (fn _ => [t]) else (fn _ => []) chaieb@20374: | add_fterms _ = I chaieb@20374: val fterms = add_fterms rhs [] wenzelm@20797: val (xs, ctxt'') = Variable.variant_fixes (replicate (length fterms) "x") ctxt' chaieb@20374: val tys = map fastype_of fterms chaieb@20374: val vs = map Free (xs ~~ tys) chaieb@20374: val env = fterms ~~ vs chaieb@20564: (* FIXME!!!!*) chaieb@20564: fun replace_fterms (t as t1 $ t2) = chaieb@20374: (case AList.lookup (op aconv) env t of chaieb@20374: SOME v => v chaieb@20564: | NONE => replace_fterms t1 $ replace_fterms t2) chaieb@20374: | replace_fterms t = (case AList.lookup (op aconv) env t of chaieb@20374: SOME v => v chaieb@20374: | NONE => t) chaieb@20374: chaieb@20374: 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))) chaieb@20374: | mk_def (t, v) = HOLogic.mk_Trueprop (HOLogic.mk_eq (v, t)) chaieb@23605: fun tryext x = (x RS ext2 handle THM _ => x) chaieb@20374: val cong = (Goal.prove ctxt'' [] (map mk_def env) chaieb@20374: (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, replace_fterms rhs))) chaieb@20374: (fn x => LocalDefs.unfold_tac (#context x) (map tryext (#prems x)) chaieb@20374: THEN rtac th' 1)) RS sym chaieb@20374: chaieb@20374: val (cong' :: vars') = chaieb@20374: Variable.export ctxt'' ctxt (cong :: map (Drule.mk_term o cert) vs) chaieb@20374: val vs' = map (fst o fst o Term.dest_Var o Thm.term_of o Drule.dest_term) vars' chaieb@20374: chaieb@20374: in (vs', cong') end; wenzelm@20319: (* congs is a list of pairs (P,th) where th is a theorem for *) wenzelm@20319: (* [| f p1 = A1; ...; f pn = An|] ==> f (C p1 .. pn) = P *) wenzelm@20319: val FWD = curry (op OF); wenzelm@20319: wenzelm@20319: (* da is the decomposition for atoms, ie. it returns ([],g) where g wenzelm@20319: returns the right instance f (AtC n) = t , where AtC is the Atoms wenzelm@20319: constructor and n is the number of the atom corresponding to t *) wenzelm@20319: wenzelm@20319: (* Generic decomp for reification : matches the actual term with the wenzelm@20319: rhs of one cong rule. The result of the matching guides the wenzelm@20319: proof synthesis: The matches of the introduced Variables A1 .. An are wenzelm@20319: processed recursively wenzelm@20319: The rest is instantiated in the cong rule,i.e. no reification is needed *) wenzelm@20319: chaieb@20374: exception REIF of string; chaieb@20374: chaieb@20564: val bds = ref ([]: (typ * ((term list) * (term list))) list); chaieb@20564: chaieb@20564: fun index_of t = chaieb@20564: let chaieb@20564: val tt = HOLogic.listT (fastype_of t) chaieb@20564: in chaieb@22199: (case AList.lookup Type.could_unify (!bds) tt of chaieb@20564: NONE => error "index_of : type not found in environements!" chaieb@20564: | SOME (tbs,tats) => chaieb@20564: let chaieb@20564: val i = find_index_eq t tats chaieb@20564: val j = find_index_eq t tbs chaieb@20564: in (if j= ~1 then chaieb@20564: if i= ~1 chaieb@23605: then (bds := AList.update Type.could_unify (tt,(tbs,tats@[t])) (!bds) ; chaieb@20564: length tbs + length tats) chaieb@20564: else i else j) chaieb@20564: end) chaieb@20564: end; chaieb@20564: chaieb@20564: fun dest_listT (Type ("List.list", [T])) = T; chaieb@20374: chaieb@20374: fun decomp_genreif da cgns (t,ctxt) = chaieb@20564: let chaieb@20564: val thy = ProofContext.theory_of ctxt chaieb@20564: val cert = cterm_of thy chaieb@20564: fun tryabsdecomp (s,ctxt) = chaieb@20564: (case s of chaieb@20564: Abs(xn,xT,ta) => chaieb@20564: (let wenzelm@20797: val ([xn],ctxt') = Variable.variant_fixes ["x"] ctxt chaieb@20564: val (xn,ta) = variant_abs (xn,xT,ta) chaieb@20564: val x = Free(xn,xT) chaieb@23605: val _ = (case AList.lookup Type.could_unify (!bds) (HOLogic.listT xT) chaieb@20564: of NONE => error "tryabsdecomp: Type not found in the Environement" chaieb@20564: | SOME (bsT,atsT) => chaieb@23605: (bds := AList.update Type.could_unify (HOLogic.listT xT, ((x::bsT), atsT)) (!bds))) chaieb@20564: in ([(ta, ctxt')] , chaieb@23605: fn [th] => ((let val (bsT,asT) = the(AList.lookup Type.could_unify (!bds) (HOLogic.listT xT)) chaieb@23605: in (bds := AList.update Type.could_unify (HOLogic.listT xT,(tl bsT,asT)) (!bds)) chaieb@20564: end) ; chaieb@20564: hd (Variable.export ctxt' ctxt [(forall_intr (cert x) th) COMP allI]))) chaieb@20374: end) chaieb@20374: | _ => da (s,ctxt)) chaieb@20374: in chaieb@20374: (case cgns of chaieb@20374: [] => tryabsdecomp (t,ctxt) chaieb@20374: | ((vns,cong)::congs) => ((let chaieb@20374: val cert = cterm_of thy chaieb@20564: val certy = ctyp_of thy chaieb@20564: val (tyenv, tmenv) = wenzelm@20319: Pattern.match thy wenzelm@20319: ((fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) (concl_of cong), t) wenzelm@20319: (Envir.type_env (Envir.empty 0),Term.Vartab.empty) wenzelm@20319: val (fnvs,invs) = List.partition (fn ((vn,_),_) => vn mem vns) (Vartab.dest tmenv) chaieb@20564: val (fts,its) = chaieb@20564: (map (snd o snd) fnvs, chaieb@20564: map (fn ((vn,vi),(tT,t)) => (cert(Var ((vn,vi),tT)), cert t)) invs) chaieb@20564: val ctyenv = map (fn ((vn,vi),(s,ty)) => (certy (TVar((vn,vi),s)), certy ty)) (Vartab.dest tyenv) chaieb@20564: in (fts ~~ (replicate (length fts) ctxt), FWD (instantiate (ctyenv, its) cong)) wenzelm@20319: end) chaieb@20564: handle MATCH => decomp_genreif da congs (t,ctxt))) chaieb@20564: end; wenzelm@20319: (* looks for the atoms equation and instantiates it with the right number *) chaieb@20374: chaieb@23548: chaieb@20374: fun mk_decompatom eqs (t,ctxt) = chaieb@23605: let chaieb@23605: val tT = fastype_of t chaieb@23605: fun isat eq = chaieb@23605: let chaieb@23605: val rhs = eq |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd chaieb@20564: in exists_Const chaieb@20564: (fn (n,ty) => n="List.nth" chaieb@20564: andalso chaieb@22199: AList.defined Type.could_unify (!bds) (domain_type ty)) rhs chaieb@22199: andalso Type.could_unify (fastype_of rhs, tT) chaieb@20564: end chaieb@23605: fun get_nths t acc = chaieb@23605: case t of chaieb@23605: Const("List.nth",_)$vs$n => insert (fn ((a,_),(b,_)) => a aconv b) (t,(vs,n)) acc chaieb@23605: | t1$t2 => get_nths t1 (get_nths t2 acc) chaieb@23605: | Abs(_,_,t') => get_nths t' acc chaieb@23605: | _ => acc chaieb@23548: chaieb@23605: fun chaieb@23605: tryeqs [] = error "Can not find the atoms equation" chaieb@23605: | tryeqs (eq::eqs) = (( chaieb@23605: let chaieb@23605: val rhs = eq |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd chaieb@23605: val nths = get_nths rhs [] chaieb@23605: val (vss,ns) = fold_rev (fn (_,(vs,n)) => fn (vss,ns) => chaieb@23605: (insert (op aconv) vs vss, insert (op aconv) n ns)) nths ([],[]) chaieb@23605: val (vsns, ctxt') = Variable.variant_fixes (replicate (length vss) "vs") ctxt chaieb@23605: val (xns, ctxt'') = Variable.variant_fixes (replicate (length nths) "x") ctxt' chaieb@23605: val thy = ProofContext.theory_of ctxt'' chaieb@23605: val cert = cterm_of thy chaieb@23605: val certT = ctyp_of thy chaieb@23605: val vsns_map = vss ~~ vsns chaieb@23605: val xns_map = (fst (split_list nths)) ~~ xns chaieb@23605: val subst = map (fn (nt, xn) => (nt, Var ((xn,0), fastype_of nt))) xns_map chaieb@23605: val rhs_P = subst_free subst rhs chaieb@23605: val (tyenv, tmenv) = Pattern.match chaieb@23605: thy (rhs_P, t) chaieb@23605: (Envir.type_env (Envir.empty 0),Term.Vartab.empty) chaieb@23605: val sbst = Envir.subst_vars (tyenv, tmenv) chaieb@23605: val sbsT = Envir.typ_subst_TVars tyenv chaieb@23605: val subst_ty = map (fn (n,(s,t)) => (certT (TVar (n, s)), certT t)) chaieb@23605: (Vartab.dest tyenv) chaieb@23605: val tml = Vartab.dest tmenv chaieb@23605: val t's = map (fn xn => snd (valOf (AList.lookup (op =) tml (xn,0)))) xns (* FIXME : Express with sbst*) chaieb@23605: val subst_ns = map (fn (Const _ $ vs $ n, Var (xn0,T)) => chaieb@23605: (cert n, snd (valOf (AList.lookup (op =) tml xn0)) chaieb@23605: |> (index_of #> IntInf.fromInt #> HOLogic.mk_nat #> cert))) chaieb@23605: subst chaieb@23605: val subst_vs = chaieb@23605: let chaieb@23605: fun ty (Const _ $ (vs as Var (vsn,lT)) $ n, Var (xn0,T)) = (certT T, certT (sbsT T)) chaieb@23605: fun h (Const _ $ (vs as Var (vsn,lT)) $ n, Var (xn0,T)) = chaieb@20564: let chaieb@23605: val cns = sbst (Const("List.list.Cons", T --> lT --> lT)) chaieb@23605: val lT' = sbsT lT chaieb@23605: val (bsT,asT) = the (AList.lookup Type.could_unify (!bds) lT) chaieb@23605: val vsn = valOf (AList.lookup (op =) vsns_map vs) chaieb@23605: val cvs = cert (fold_rev (fn x => fn xs => cns$x$xs) bsT (Free (vsn, lT'))) chaieb@23605: in (cert vs, cvs) end chaieb@23605: in map h subst end chaieb@23605: val cts = map (fn ((vn,vi),(tT,t)) => (cert(Var ((vn,vi),tT)), cert t)) chaieb@23605: (fold (AList.delete (fn (((a: string),_),(b,_)) => a = b)) chaieb@23605: (map (fn n => (n,0)) xns) tml) chaieb@23605: val substt = chaieb@23605: let val ih = Drule.cterm_rule (Thm.instantiate (subst_ty,[])) chaieb@23605: in map (fn (v,t) => (ih v, ih t)) (subst_ns@subst_vs@cts) end chaieb@23605: val th = (instantiate (subst_ty, substt) eq) RS sym chaieb@23605: in hd (Variable.export ctxt'' ctxt [th]) end) chaieb@23605: handle MATCH => tryeqs eqs) chaieb@23605: in ([], fn _ => tryeqs (filter isat eqs)) chaieb@23605: end; chaieb@20374: wenzelm@20319: (* Generic reification procedure: *) wenzelm@20319: (* creates all needed cong rules and then just uses the theorem synthesis *) chaieb@20564: chaieb@20564: fun mk_congs ctxt raw_eqs = chaieb@20564: let haftmann@21078: val fs = fold_rev (fn eq => haftmann@20853: insert (op =) (eq |> prop_of |> HOLogic.dest_Trueprop chaieb@20564: |> HOLogic.dest_eq |> fst |> strip_comb haftmann@21078: |> fst)) raw_eqs [] haftmann@21078: val tys = fold_rev (fn f => fn ts => (f |> fastype_of |> binder_types |> split_last |> fst) haftmann@21078: union ts) fs [] chaieb@20564: val _ = bds := AList.make (fn _ => ([],[])) tys wenzelm@20797: val (vs, ctxt') = Variable.variant_fixes (replicate (length tys) "vs") ctxt chaieb@20564: val thy = ProofContext.theory_of ctxt' chaieb@20564: val cert = cterm_of thy chaieb@20564: val vstys = map (fn (t,v) => (t,SOME (cert (Free(v,t))))) chaieb@20564: (tys ~~ vs) chaieb@20564: fun insteq eq ts = chaieb@20564: let val itms = map (fn t => t|> (AList.lookup (op =) vstys) |> the) ts chaieb@20564: in instantiate' [] itms eq chaieb@20564: end chaieb@20564: val eqs = map (fn eq => eq |> prop_of |> HOLogic.dest_Trueprop chaieb@20564: |> HOLogic.dest_eq |> fst |> strip_comb |> fst |> fastype_of chaieb@20564: |> binder_types |> split_last |> fst chaieb@20564: |> (insteq eq)) raw_eqs chaieb@20564: val (ps,congs) = split_list (map (mk_congeq ctxt' fs) eqs) chaieb@20564: in ps ~~ (Variable.export ctxt' ctxt congs) chaieb@20564: end; chaieb@20564: wenzelm@20319: fun genreif ctxt raw_eqs t = chaieb@20564: let chaieb@20564: val _ = bds := [] chaieb@20564: val congs = mk_congs ctxt raw_eqs chaieb@20564: val th = divide_and_conquer (decomp_genreif (mk_decompatom raw_eqs) congs) (t,ctxt) chaieb@20564: val tys = th |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd chaieb@20564: |> strip_comb |> fst |> fastype_of |> strip_type |> fst chaieb@20564: |> split_last |> fst chaieb@20564: val cert = cterm_of (ProofContext.theory_of ctxt) chaieb@23605: val cvs = map (fn t => t |> (AList.lookup Type.could_unify (!bds)) |> the |> snd wenzelm@21757: |> HOLogic.mk_list (dest_listT t) |> cert |> SOME) chaieb@20564: tys chaieb@20564: val th' = (instantiate' [] cvs (th RS sym)) RS sym chaieb@20564: val t' = (fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) th' chaieb@20564: val th'' = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, t'))) chaieb@20564: (fn _ => Simp_tac 1) chaieb@20564: val _ = bds := [] chaieb@20564: in FWD trans [th'',th'] chaieb@20564: end; wenzelm@20319: haftmann@21878: fun genreflect ctxt conv corr_thm raw_eqs t = wenzelm@20319: let val th = FWD trans [genreif ctxt raw_eqs t, corr_thm RS sym] wenzelm@20319: val ft = (snd o Thm.dest_comb o snd o Thm.dest_comb o snd o Thm.dest_comb o cprop_of) th haftmann@21878: val rth = conv ft wenzelm@20319: in simplify (HOL_basic_ss addsimps raw_eqs addsimps [nth_Cons_0, nth_Cons_Suc]) wenzelm@20319: (simplify (HOL_basic_ss addsimps [rth]) th) wenzelm@20319: end wenzelm@20319: wenzelm@20319: fun genreify_tac ctxt eqs to i = (fn st => wenzelm@20319: let wenzelm@20319: val P = HOLogic.dest_Trueprop (List.nth (prems_of st, i - 1)) wenzelm@20319: val t = (case to of NONE => P | SOME x => x) wenzelm@20319: val th = (genreif ctxt eqs t) RS ssubst wenzelm@20319: in rtac th i st wenzelm@20319: end); wenzelm@20319: wenzelm@20319: (* Reflection calls reification and uses the correctness *) wenzelm@20319: (* theorem assumed to be the dead of the list *) haftmann@21878: fun gen_reflection_tac ctxt conv (corr_thm :: raw_eqs) to i = (fn st => haftmann@21878: let haftmann@21878: val P = HOLogic.dest_Trueprop (nth (prems_of st) (i - 1)); haftmann@21878: val t = the_default P to; haftmann@21878: val th = genreflect ctxt conv corr_thm raw_eqs t haftmann@21878: RS ssubst; haftmann@21878: in rtac th i st end); haftmann@21878: haftmann@21878: fun reflection_tac ctxt = gen_reflection_tac ctxt NBE.normalization_conv; wenzelm@20797: end