--- a/src/HOL/Tools/transfer.ML Thu Oct 01 16:03:43 2009 +0200
+++ b/src/HOL/Tools/transfer.ML Thu Oct 01 16:09:47 2009 +0200
@@ -1,5 +1,5 @@
(* Author: Amine Chaieb, University of Cambridge, 2009
- Jeremy Avigad, Carnegie Mellon University
+ Author: Jeremy Avigad, Carnegie Mellon University
*)
signature TRANSFER =
@@ -14,10 +14,8 @@
structure Transfer : TRANSFER =
struct
-val eq_thm = Thm.eq_thm;
-
type entry = {inj : thm list, emb : thm list, ret : thm list, cong : thm list,
- guess : bool, hints : string list};
+ guess : bool, hints : string list};
type data = simpset * (thm * entry) list;
structure Data = GenericDataFun
@@ -26,36 +24,34 @@
val empty = (HOL_ss, []);
val extend = I;
fun merge _ ((ss1, e1), (ss2, e2)) =
- (merge_ss (ss1, ss2), AList.merge eq_thm (K true) (e1, e2));
+ (merge_ss (ss1, ss2), AList.merge Thm.eq_thm (K true) (e1, e2));
);
val get = Data.get o Context.Proof;
-fun del_data key = apsnd (remove (eq_fst eq_thm) (key, []));
+fun del_data key = apsnd (remove (eq_fst Thm.eq_thm) (key, []));
val del = Thm.declaration_attribute (Data.map o del_data);
-val add_ss = Thm.declaration_attribute
+val add_ss = Thm.declaration_attribute
(fn th => Data.map (fn (ss,data) => (ss addsimps [th], data)));
-val del_ss = Thm.declaration_attribute
+val del_ss = Thm.declaration_attribute
(fn th => Data.map (fn (ss,data) => (ss delsimps [th], data)));
val transM_pat = (Thm.dest_arg1 o Thm.dest_arg o cprop_of) @{thm TransferMorphism_def};
fun merge_update eq m (k,v) [] = [(k,v)]
- | merge_update eq m (k,v) ((k',v')::al) =
+ | merge_update eq m (k,v) ((k',v')::al) =
if eq (k,k') then (k',m (v,v')):: al else (k',v') :: merge_update eq m (k,v) al
-fun C f x y = f y x
-
-fun simpset_of_entry injonly {inj = inj, emb = emb, ret = ret, cong = cg, guess = g, hints = hints} =
+fun simpset_of_entry injonly {inj = inj, emb = emb, ret = ret, cong = cg, guess = g, hints = hints} =
HOL_ss addsimps inj addsimps (if injonly then [] else emb@ret) addcongs cg;
-fun basic_transfer_rule injonly a0 D0 e leave ctxt0 th =
- let
+fun basic_transfer_rule injonly a0 D0 e leave ctxt0 th =
+ let
val ([a,D], ctxt) = apfst (map Drule.dest_term o snd) (Variable.import true (map Drule.mk_term [a0, D0]) ctxt0)
- val (aT,bT) =
- let val T = typ_of (ctyp_of_term a)
+ val (aT,bT) =
+ let val T = typ_of (ctyp_of_term a)
in (Term.range_type T, Term.domain_type T)
end
val ctxt' = (Variable.declare_term (term_of a) o Variable.declare_term (term_of D) o Variable.declare_thm th) ctxt
@@ -65,60 +61,64 @@
val cfis = map ((cterm_of o ProofContext.theory_of) ctxt'' o (fn n => Free (n, bT))) ins
val cis = map (Thm.capply a) cfis
val (hs,ctxt''') = Assumption.add_assumes (map (fn ct => Thm.capply @{cterm "Trueprop"} (Thm.capply D ct)) cfis) ctxt''
- val th1 = Drule.cterm_instantiate (cns~~ cis) th
- val th2 = fold (C implies_elim) hs (fold_rev implies_intr (map cprop_of hs) th1)
- val th3 = Simplifier.asm_full_simplify (Simplifier.context ctxt''' (simpset_of_entry injonly e))
+ val th1 = Drule.cterm_instantiate (cns ~~ cis) th
+ val th2 = fold Thm.elim_implies hs (fold_rev implies_intr (map cprop_of hs) th1)
+ val th3 = Simplifier.asm_full_simplify (Simplifier.context ctxt''' (simpset_of_entry injonly e))
(fold_rev implies_intr (map cprop_of hs) th2)
in hd (Variable.export ctxt''' ctxt0 [th3]) end;
local
-fun transfer_ruleh a D leave ctxt th =
+fun transfer_ruleh a D leave ctxt th =
let val (ss,al) = get ctxt
val a0 = cterm_of (ProofContext.theory_of ctxt) a
val D0 = cterm_of (ProofContext.theory_of ctxt) D
- fun h (th', e) = let val (a',D') = (Thm.dest_binop o Thm.dest_arg o cprop_of) th'
+ fun h (th', e) = let val (a',D') = (Thm.dest_binop o Thm.dest_arg o cprop_of) th'
in if a0 aconvc a' andalso D0 aconvc D' then SOME e else NONE
end
in case get_first h al of
SOME e => basic_transfer_rule false a0 D0 e leave ctxt th
| NONE => error "Transfer: corresponding instance not found in context-data"
end
-in fun transfer_rule (a,D) leave (gctxt,th) =
+in fun transfer_rule (a,D) leave (gctxt,th) =
(gctxt, transfer_ruleh a D leave (Context.proof_of gctxt) th)
end;
fun splits P [] = []
- | splits P (xxs as (x::xs)) =
+ | splits P (xxs as (x::xs)) =
let val pss = filter (P x) xxs
val qss = filter_out (P x) xxs
in if null pss then [qss] else if null qss then [pss] else pss:: splits P qss
end
-fun all_transfers leave (gctxt,th) =
- let
+fun all_transfers leave (gctxt,th) =
+ let
val ctxt = Context.proof_of gctxt
val tys = map snd (Term.add_vars (prop_of th) [])
val _ = if null tys then error "transfer: Unable to guess instance" else ()
- val tyss = splits (curry Type.could_unify) tys
+ val tyss = splits (curry Type.could_unify) tys
val get_ty = typ_of o ctyp_of_term o fst o Thm.dest_binop o Thm.dest_arg o cprop_of
val get_aD = Thm.dest_binop o Thm.dest_arg o cprop_of
- val insts =
- map_filter (fn tys =>
- get_first (fn (k,ss) => if Type.could_unify (hd tys, range_type (get_ty k))
- then SOME (get_aD k, ss)
- else NONE) (snd (get ctxt))) tyss
- val _ = if null insts then error "Transfer guesser: there were no possible instances, use direction: in order to provide a direction" else ()
+ val insts =
+ map_filter (fn tys =>
+ get_first (fn (k,ss) =>
+ if Type.could_unify (hd tys, range_type (get_ty k))
+ then SOME (get_aD k, ss)
+ else NONE) (snd (get ctxt))) tyss
+ val _ =
+ if null insts then
+ error "Transfer guesser: there were no possible instances, use direction: in order to provide a direction"
+ else ()
val ths = map (fn ((a,D),e) => basic_transfer_rule false a D e leave ctxt th) insts
val cth = Conjunction.intr_balanced ths
in (gctxt, cth)
end;
-fun transfer_rule_by_hint ls leave (gctxt,th) =
- let
+fun transfer_rule_by_hint ls leave (gctxt,th) =
+ let
val ctxt = Context.proof_of gctxt
val get_aD = Thm.dest_binop o Thm.dest_arg o cprop_of
- val insts =
- map_filter (fn (k,e) => if exists (fn l => l mem_string (#hints e)) ls
+ val insts =
+ map_filter (fn (k,e) => if exists (member (op =) (#hints e)) ls
then SOME (get_aD k, e) else NONE)
(snd (get ctxt))
val _ = if null insts then error "Transfer: No labels provided are stored in the context" else ()
@@ -128,53 +128,58 @@
end;
-fun transferred_attribute ls NONE leave =
+fun transferred_attribute ls NONE leave =
if null ls then all_transfers leave else transfer_rule_by_hint ls leave
| transferred_attribute _ (SOME (a,D)) leave = transfer_rule (a,D) leave
- (* Add data to the context *)
+
+(* Add data to the context *)
+
fun gen_merge_entries {inj = inj0, emb = emb0, ret = ret0, cong = cg0, guess = g0, hints = hints0}
- ({inj = inj1, emb = emb1, ret = ret1, cong = cg1, guess = g1, hints = hints1},
+ ({inj = inj1, emb = emb1, ret = ret1, cong = cg1, guess = g1, hints = hints1},
{inj = inj2, emb = emb2, ret = ret2, cong = cg2, guess = g2, hints = hints2})
- =
+ =
let fun h xs0 xs ys = subtract Thm.eq_thm xs0 (merge Thm.eq_thm (xs,ys)) in
- {inj = h inj0 inj1 inj2, emb = h emb0 emb1 emb2,
+ {inj = h inj0 inj1 inj2, emb = h emb0 emb1 emb2,
ret = h ret0 ret1 ret2, cong = h cg0 cg1 cg2, guess = g1 andalso g2,
- hints = subtract (op = : string*string -> bool) hints0
+ hints = subtract (op = : string*string -> bool) hints0
(hints1 union_string hints2)}
end;
local
val h = curry (merge Thm.eq_thm)
in
-fun merge_entries ({inj = inj1, emb = emb1, ret = ret1, cong = cg1, guess = g1, hints = hints1},
- {inj = inj2, emb = emb2, ret = ret2, cong = cg2, guess = g2, hints = hints2}) =
+fun merge_entries ({inj = inj1, emb = emb1, ret = ret1, cong = cg1, guess = g1, hints = hints1},
+ {inj = inj2, emb = emb2, ret = ret2, cong = cg2, guess = g2, hints = hints2}) =
{inj = h inj1 inj2, emb = h emb1 emb2, ret = h ret1 ret2, cong = h cg1 cg2, guess = g1 andalso g2, hints = hints1 union_string hints2}
-end;
+end;
fun add ((inja,injd), (emba,embd), (reta,retd), (cga,cgd), g, (hintsa, hintsd)) =
Thm.declaration_attribute (fn key => fn context => context |> Data.map
- (fn (ss, al) =>
+ (fn (ss, al) =>
let
- val _ = ((let val _ = Thm.match (transM_pat, (Thm.dest_arg o cprop_of) key)
- in 0 end)
- handle MATCH => error "Attribute expected Theorem of the form : TransferMorphism A a B b")
+ val _ = Thm.match (transM_pat, Thm.dest_arg (Thm.cprop_of key))
+ handle Pattern.MATCH =>
+ error "Attribute expected Theorem of the form : TransferMorphism A a B b"
val e0 = {inj = inja, emb = emba, ret = reta, cong = cga, guess = g, hints = hintsa}
val ed = {inj = injd, emb = embd, ret = retd, cong = cgd, guess = g, hints = hintsd}
- val entry =
- if g then
+ val entry =
+ if g then
let val (a0,D0) = (Thm.dest_binop o Thm.dest_arg o cprop_of) key
val ctxt0 = ProofContext.init (Thm.theory_of_thm key)
- val inj' = if null inja then #inj (case AList.lookup eq_thm al key of SOME e => e | NONE => error "Transfer: can not generate return rules on the fly, either add injectivity axiom or force manual mode with mode: manual")
- else inja
+ val inj' =
+ if null inja then
+ #inj
+ (case AList.lookup Thm.eq_thm al key of SOME e => e
+ | NONE => error "Transfer: can not generate return rules on the fly, either add injectivity axiom or force manual mode with mode: manual")
+ else inja
val ret' = merge Thm.eq_thm (reta, map (fn th => basic_transfer_rule true a0 D0 {inj = inj', emb = [], ret = [], cong = cga, guess = g, hints = hintsa} [] ctxt0 th RS sym) emba)
- in {inj = inja, emb = emba, ret = ret', cong = cga, guess = g, hints = hintsa} end
+ in {inj = inja, emb = emba, ret = ret', cong = cga, guess = g, hints = hintsa} end
else e0
- in (ss, merge_update eq_thm (gen_merge_entries ed) (key, entry) al)
+ in (ss, merge_update Thm.eq_thm (gen_merge_entries ed) (key, entry) al)
end));
-
(* concrete syntax *)
local
@@ -199,7 +204,7 @@
val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat
val terms = thms >> map Drule.dest_term
-val types = thms >> (Logic.dest_type o HOLogic.dest_Trueprop o prop_of o hd)
+val types = thms >> (Logic.dest_type o HOLogic.dest_Trueprop o prop_of o hd)
val name = Scan.lift Args.name
val names = Scan.repeat (Scan.unless any_keyword name)
fun optional scan = Scan.optional scan []