| author | nipkow | 
| Tue, 20 Oct 2009 14:44:19 +0200 | |
| changeset 33020 | 0908ed080ccf | 
| parent 32813 | dac196e23093 | 
| child 33037 | b22e44496dc2 | 
| permissions | -rw-r--r-- | 
| 32557 
3cfe4c13aa6e
plain structure name; signature constraint; shorter lines
 haftmann parents: 
32476diff
changeset | 1 | (* Author: Amine Chaieb, University of Cambridge, 2009 | 
| 32813 | 2 | Author: Jeremy Avigad, Carnegie Mellon University | 
| 31706 | 3 | *) | 
| 4 | ||
| 32557 
3cfe4c13aa6e
plain structure name; signature constraint; shorter lines
 haftmann parents: 
32476diff
changeset | 5 | signature TRANSFER = | 
| 31706 | 6 | sig | 
| 7 | type data | |
| 8 | type entry | |
| 9 | val get: Proof.context -> data | |
| 10 | val del: attribute | |
| 11 | val setup: theory -> theory | |
| 12 | end; | |
| 13 | ||
| 32557 
3cfe4c13aa6e
plain structure name; signature constraint; shorter lines
 haftmann parents: 
32476diff
changeset | 14 | structure Transfer : TRANSFER = | 
| 31706 | 15 | struct | 
| 32557 
3cfe4c13aa6e
plain structure name; signature constraint; shorter lines
 haftmann parents: 
32476diff
changeset | 16 | |
| 
3cfe4c13aa6e
plain structure name; signature constraint; shorter lines
 haftmann parents: 
32476diff
changeset | 17 | type entry = {inj : thm list, emb : thm list, ret : thm list, cong : thm list,
 | 
| 32813 | 18 | guess : bool, hints : string list}; | 
| 31706 | 19 | type data = simpset * (thm * entry) list; | 
| 20 | ||
| 21 | structure Data = GenericDataFun | |
| 22 | ( | |
| 23 | type T = data; | |
| 24 | val empty = (HOL_ss, []); | |
| 25 | val extend = I; | |
| 32557 
3cfe4c13aa6e
plain structure name; signature constraint; shorter lines
 haftmann parents: 
32476diff
changeset | 26 | fun merge _ ((ss1, e1), (ss2, e2)) = | 
| 32813 | 27 | (merge_ss (ss1, ss2), AList.merge Thm.eq_thm (K true) (e1, e2)); | 
| 31706 | 28 | ); | 
| 29 | ||
| 30 | val get = Data.get o Context.Proof; | |
| 31 | ||
| 32813 | 32 | fun del_data key = apsnd (remove (eq_fst Thm.eq_thm) (key, [])); | 
| 31706 | 33 | |
| 34 | val del = Thm.declaration_attribute (Data.map o del_data); | |
| 32813 | 35 | val add_ss = Thm.declaration_attribute | 
| 31706 | 36 | (fn th => Data.map (fn (ss,data) => (ss addsimps [th], data))); | 
| 37 | ||
| 32813 | 38 | val del_ss = Thm.declaration_attribute | 
| 31706 | 39 | (fn th => Data.map (fn (ss,data) => (ss delsimps [th], data))); | 
| 40 | ||
| 41 | val transM_pat = (Thm.dest_arg1 o Thm.dest_arg o cprop_of) @{thm TransferMorphism_def};
 | |
| 42 | ||
| 43 | fun merge_update eq m (k,v) [] = [(k,v)] | |
| 32813 | 44 | | merge_update eq m (k,v) ((k',v')::al) = | 
| 31706 | 45 | if eq (k,k') then (k',m (v,v')):: al else (k',v') :: merge_update eq m (k,v) al | 
| 46 | ||
| 32813 | 47 | fun simpset_of_entry injonly {inj = inj, emb = emb, ret = ret, cong = cg, guess = g, hints = hints} =
 | 
| 31706 | 48 | HOL_ss addsimps inj addsimps (if injonly then [] else emb@ret) addcongs cg; | 
| 49 | ||
| 32813 | 50 | fun basic_transfer_rule injonly a0 D0 e leave ctxt0 th = | 
| 51 | let | |
| 31794 
71af1fd6a5e4
renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
 wenzelm parents: 
31706diff
changeset | 52 | val ([a,D], ctxt) = apfst (map Drule.dest_term o snd) (Variable.import true (map Drule.mk_term [a0, D0]) ctxt0) | 
| 32813 | 53 | val (aT,bT) = | 
| 54 | let val T = typ_of (ctyp_of_term a) | |
| 31706 | 55 | in (Term.range_type T, Term.domain_type T) | 
| 56 | end | |
| 57 | val ctxt' = (Variable.declare_term (term_of a) o Variable.declare_term (term_of D) o Variable.declare_thm th) ctxt | |
| 58 | val ns = filter (fn i => Type.could_unify (snd i, aT) andalso not (fst (fst i) mem_string leave)) (Term.add_vars (prop_of th) []) | |
| 59 | val (ins, ctxt'') = Variable.variant_fixes (map (fst o fst) ns) ctxt' | |
| 60 | val cns = map ((cterm_of o ProofContext.theory_of) ctxt'' o Var) ns | |
| 61 | val cfis = map ((cterm_of o ProofContext.theory_of) ctxt'' o (fn n => Free (n, bT))) ins | |
| 62 | val cis = map (Thm.capply a) cfis | |
| 63 |   val (hs,ctxt''') = Assumption.add_assumes (map (fn ct => Thm.capply @{cterm "Trueprop"} (Thm.capply D ct)) cfis) ctxt''
 | |
| 32813 | 64 | val th1 = Drule.cterm_instantiate (cns ~~ cis) th | 
| 65 | val th2 = fold Thm.elim_implies hs (fold_rev implies_intr (map cprop_of hs) th1) | |
| 66 | val th3 = Simplifier.asm_full_simplify (Simplifier.context ctxt''' (simpset_of_entry injonly e)) | |
| 31706 | 67 | (fold_rev implies_intr (map cprop_of hs) th2) | 
| 68 | in hd (Variable.export ctxt''' ctxt0 [th3]) end; | |
| 69 | ||
| 70 | local | |
| 32813 | 71 | fun transfer_ruleh a D leave ctxt th = | 
| 31706 | 72 | let val (ss,al) = get ctxt | 
| 73 | val a0 = cterm_of (ProofContext.theory_of ctxt) a | |
| 74 | val D0 = cterm_of (ProofContext.theory_of ctxt) D | |
| 32813 | 75 | fun h (th', e) = let val (a',D') = (Thm.dest_binop o Thm.dest_arg o cprop_of) th' | 
| 31706 | 76 | in if a0 aconvc a' andalso D0 aconvc D' then SOME e else NONE | 
| 77 | end | |
| 78 | in case get_first h al of | |
| 79 | SOME e => basic_transfer_rule false a0 D0 e leave ctxt th | |
| 80 | | NONE => error "Transfer: corresponding instance not found in context-data" | |
| 81 | end | |
| 32813 | 82 | in fun transfer_rule (a,D) leave (gctxt,th) = | 
| 31706 | 83 | (gctxt, transfer_ruleh a D leave (Context.proof_of gctxt) th) | 
| 84 | end; | |
| 85 | ||
| 86 | fun splits P [] = [] | |
| 32813 | 87 | | splits P (xxs as (x::xs)) = | 
| 31706 | 88 | let val pss = filter (P x) xxs | 
| 89 | val qss = filter_out (P x) xxs | |
| 90 | in if null pss then [qss] else if null qss then [pss] else pss:: splits P qss | |
| 91 | end | |
| 92 | ||
| 32813 | 93 | fun all_transfers leave (gctxt,th) = | 
| 94 | let | |
| 31706 | 95 | val ctxt = Context.proof_of gctxt | 
| 96 | val tys = map snd (Term.add_vars (prop_of th) []) | |
| 97 | val _ = if null tys then error "transfer: Unable to guess instance" else () | |
| 32813 | 98 | val tyss = splits (curry Type.could_unify) tys | 
| 31706 | 99 | val get_ty = typ_of o ctyp_of_term o fst o Thm.dest_binop o Thm.dest_arg o cprop_of | 
| 100 | val get_aD = Thm.dest_binop o Thm.dest_arg o cprop_of | |
| 32813 | 101 | val insts = | 
| 102 | map_filter (fn tys => | |
| 103 | get_first (fn (k,ss) => | |
| 104 | if Type.could_unify (hd tys, range_type (get_ty k)) | |
| 105 | then SOME (get_aD k, ss) | |
| 106 | else NONE) (snd (get ctxt))) tyss | |
| 107 | val _ = | |
| 108 | if null insts then | |
| 109 | error "Transfer guesser: there were no possible instances, use direction: in order to provide a direction" | |
| 110 | else () | |
| 31706 | 111 | val ths = map (fn ((a,D),e) => basic_transfer_rule false a D e leave ctxt th) insts | 
| 112 | val cth = Conjunction.intr_balanced ths | |
| 113 | in (gctxt, cth) | |
| 114 | end; | |
| 115 | ||
| 32813 | 116 | fun transfer_rule_by_hint ls leave (gctxt,th) = | 
| 117 | let | |
| 31706 | 118 | val ctxt = Context.proof_of gctxt | 
| 119 | val get_aD = Thm.dest_binop o Thm.dest_arg o cprop_of | |
| 32813 | 120 | val insts = | 
| 121 | map_filter (fn (k,e) => if exists (member (op =) (#hints e)) ls | |
| 32557 
3cfe4c13aa6e
plain structure name; signature constraint; shorter lines
 haftmann parents: 
32476diff
changeset | 122 | then SOME (get_aD k, e) else NONE) | 
| 31706 | 123 | (snd (get ctxt)) | 
| 124 | val _ = if null insts then error "Transfer: No labels provided are stored in the context" else () | |
| 125 | val ths = map (fn ((a,D),e) => basic_transfer_rule false a D e leave ctxt th) insts | |
| 126 | val cth = Conjunction.intr_balanced ths | |
| 127 | in (gctxt, cth) | |
| 128 | end; | |
| 129 | ||
| 130 | ||
| 32813 | 131 | fun transferred_attribute ls NONE leave = | 
| 31706 | 132 | if null ls then all_transfers leave else transfer_rule_by_hint ls leave | 
| 133 | | transferred_attribute _ (SOME (a,D)) leave = transfer_rule (a,D) leave | |
| 134 | ||
| 32813 | 135 | |
| 136 | (* Add data to the context *) | |
| 137 | ||
| 31706 | 138 | fun gen_merge_entries {inj = inj0, emb = emb0, ret = ret0, cong = cg0, guess = g0, hints = hints0}
 | 
| 32813 | 139 |                       ({inj = inj1, emb = emb1, ret = ret1, cong = cg1, guess = g1, hints = hints1},
 | 
| 31706 | 140 |                        {inj = inj2, emb = emb2, ret = ret2, cong = cg2, guess = g2, hints = hints2})
 | 
| 32813 | 141 | = | 
| 31706 | 142 | let fun h xs0 xs ys = subtract Thm.eq_thm xs0 (merge Thm.eq_thm (xs,ys)) in | 
| 32813 | 143 |  {inj = h inj0 inj1 inj2, emb = h emb0 emb1 emb2,
 | 
| 31706 | 144 | ret = h ret0 ret1 ret2, cong = h cg0 cg1 cg2, guess = g1 andalso g2, | 
| 32813 | 145 | hints = subtract (op = : string*string -> bool) hints0 | 
| 31706 | 146 | (hints1 union_string hints2)} | 
| 147 | end; | |
| 148 | ||
| 149 | local | |
| 150 | val h = curry (merge Thm.eq_thm) | |
| 151 | in | |
| 32813 | 152 | fun merge_entries ({inj = inj1, emb = emb1, ret = ret1, cong = cg1, guess = g1, hints = hints1},
 | 
| 153 |                    {inj = inj2, emb = emb2, ret = ret2, cong = cg2, guess = g2, hints = hints2}) =
 | |
| 31706 | 154 |     {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}
 | 
| 32813 | 155 | end; | 
| 31706 | 156 | |
| 157 | fun add ((inja,injd), (emba,embd), (reta,retd), (cga,cgd), g, (hintsa, hintsd)) = | |
| 158 | Thm.declaration_attribute (fn key => fn context => context |> Data.map | |
| 32813 | 159 | (fn (ss, al) => | 
| 31706 | 160 | let | 
| 32813 | 161 | val _ = Thm.match (transM_pat, Thm.dest_arg (Thm.cprop_of key)) | 
| 162 | handle Pattern.MATCH => | |
| 163 | error "Attribute expected Theorem of the form : TransferMorphism A a B b" | |
| 31706 | 164 |       val e0 = {inj = inja, emb = emba, ret = reta, cong = cga, guess = g, hints = hintsa}
 | 
| 165 |       val ed = {inj = injd, emb = embd, ret = retd, cong = cgd, guess = g, hints = hintsd}
 | |
| 32813 | 166 | val entry = | 
| 167 | if g then | |
| 31706 | 168 | let val (a0,D0) = (Thm.dest_binop o Thm.dest_arg o cprop_of) key | 
| 169 | val ctxt0 = ProofContext.init (Thm.theory_of_thm key) | |
| 32813 | 170 | val inj' = | 
| 171 | if null inja then | |
| 172 | #inj | |
| 173 | (case AList.lookup Thm.eq_thm al key of SOME e => e | |
| 174 | | NONE => error "Transfer: can not generate return rules on the fly, either add injectivity axiom or force manual mode with mode: manual") | |
| 175 | else inja | |
| 31706 | 176 |              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)
 | 
| 32813 | 177 |          in {inj = inja, emb = emba, ret = ret', cong = cga, guess = g, hints = hintsa} end
 | 
| 31706 | 178 | else e0 | 
| 32813 | 179 | in (ss, merge_update Thm.eq_thm (gen_merge_entries ed) (key, entry) al) | 
| 31706 | 180 | end)); | 
| 181 | ||
| 182 | ||
| 183 | (* concrete syntax *) | |
| 184 | ||
| 185 | local | |
| 186 | ||
| 187 | fun keyword k = Scan.lift (Args.$$$ k) >> K () | |
| 188 | fun keywordC k = Scan.lift (Args.$$$ k -- Args.colon) >> K () | |
| 189 | ||
| 190 | val congN = "cong" | |
| 191 | val injN = "inj" | |
| 192 | val embedN = "embed" | |
| 193 | val returnN = "return" | |
| 194 | val addN = "add" | |
| 195 | val delN = "del" | |
| 196 | val modeN = "mode" | |
| 197 | val automaticN = "automatic" | |
| 198 | val manualN = "manual" | |
| 199 | val directionN = "direction" | |
| 200 | val labelsN = "labels" | |
| 201 | val leavingN = "leaving" | |
| 202 | ||
| 203 | val any_keyword = keywordC congN || keywordC injN || keywordC embedN || keywordC returnN || keywordC directionN || keywordC modeN || keywordC delN || keywordC labelsN || keywordC leavingN | |
| 204 | ||
| 205 | val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat | |
| 206 | val terms = thms >> map Drule.dest_term | |
| 32813 | 207 | val types = thms >> (Logic.dest_type o HOLogic.dest_Trueprop o prop_of o hd) | 
| 31706 | 208 | val name = Scan.lift Args.name | 
| 209 | val names = Scan.repeat (Scan.unless any_keyword name) | |
| 210 | fun optional scan = Scan.optional scan [] | |
| 211 | fun optional2 scan = Scan.optional scan ([],[]) | |
| 212 | ||
| 213 | val mode = keywordC modeN |-- ((Scan.lift (Args.$$$ manualN) >> K false) || (Scan.lift (Args.$$$ automaticN) >> K true)) | |
| 214 | val inj = (keywordC injN |-- thms) -- optional (keywordC delN |-- thms) | |
| 215 | val embed = (keywordC embedN |-- thms) -- optional (keywordC delN |-- thms) | |
| 216 | val return = (keywordC returnN |-- thms) -- optional (keywordC delN |-- thms) | |
| 217 | val cong = (keywordC congN |-- thms) -- optional (keywordC delN |-- thms) | |
| 218 | val addscan = Scan.unless any_keyword (keyword addN) | |
| 219 | val labels = (keywordC labelsN |-- names) -- optional (keywordC delN |-- names) | |
| 220 | val entry = Scan.optional mode true -- optional2 inj -- optional2 embed -- optional2 return -- optional2 cong -- optional2 labels | |
| 221 | ||
| 222 | val transf_add = addscan |-- entry | |
| 223 | in | |
| 224 | ||
| 225 | val install_att_syntax = | |
| 226 | (Scan.lift (Args.$$$ delN >> K del) || | |
| 227 | transf_add | |
| 228 | >> (fn (((((g, inj), embed), ret), cg), hints) => add (inj, embed, ret, cg, g, hints))) | |
| 229 | ||
| 32476 | 230 | val transferred_att_syntax = (optional names -- Scan.option (keywordC directionN |-- (Args.term -- Args.term)) | 
| 231 | -- optional (keywordC leavingN |-- names) >> (fn ((hints, aD),leave) => transferred_attribute hints aD leave)); | |
| 31706 | 232 | |
| 233 | end; | |
| 234 | ||
| 235 | ||
| 236 | (* theory setup *) | |
| 237 | ||
| 238 | val setup = | |
| 239 |   Attrib.setup @{binding transfer} install_att_syntax
 | |
| 240 | "Installs transfer data" #> | |
| 241 |   Attrib.setup @{binding transfer_simps} (Attrib.add_del add_ss del_ss)
 | |
| 242 | "simp rules for transfer" #> | |
| 243 |   Attrib.setup @{binding transferred} transferred_att_syntax
 | |
| 244 | "Transfers theorems"; | |
| 245 | ||
| 246 | end; |