| author | haftmann | 
| Tue, 05 Jun 2007 15:16:08 +0200 | |
| changeset 23247 | b99dce43d252 | 
| parent 22846 | fb79144af9a3 | 
| child 24039 | 273698405054 | 
| permissions | -rw-r--r-- | 
| 17333 | 1 | (* Title : HOL/Hyperreal/transfer.ML | 
| 2 | ID : $Id$ | |
| 3 | Author : Brian Huffman | |
| 4 | ||
| 18729 | 5 | Transfer principle tactic for nonstandard analysis. | 
| 17333 | 6 | *) | 
| 7 | ||
| 18729 | 8 | signature TRANSFER = | 
| 17329 | 9 | sig | 
| 18729 | 10 | val transfer_tac: Proof.context -> thm list -> int -> tactic | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17333diff
changeset | 11 | val add_const: string -> theory -> theory | 
| 18708 | 12 | val setup: theory -> theory | 
| 17329 | 13 | end; | 
| 14 | ||
| 18729 | 15 | structure Transfer: TRANSFER = | 
| 17329 | 16 | struct | 
| 17 | ||
| 18729 | 18 | structure TransferData = GenericDataFun | 
| 22846 | 19 | ( | 
| 17329 | 20 |   type T = {
 | 
| 21 | intros: thm list, | |
| 22 | unfolds: thm list, | |
| 23 | refolds: thm list, | |
| 24 | consts: string list | |
| 25 | }; | |
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17333diff
changeset | 26 |   val empty = {intros = [], unfolds = [], refolds = [], consts = []};
 | 
| 17329 | 27 | val extend = I; | 
| 28 | fun merge _ | |
| 29 |     ({intros = intros1, unfolds = unfolds1,
 | |
| 30 | refolds = refolds1, consts = consts1}, | |
| 31 |      {intros = intros2, unfolds = unfolds2,
 | |
| 32 | refolds = refolds2, consts = consts2}) = | |
| 33 |    {intros = Drule.merge_rules (intros1, intros2),
 | |
| 34 | unfolds = Drule.merge_rules (unfolds1, unfolds2), | |
| 35 | refolds = Drule.merge_rules (refolds1, refolds2), | |
| 19734 | 36 | consts = Library.merge (op =) (consts1, consts2)} : T; | 
| 22846 | 37 | ); | 
| 17329 | 38 | |
| 39 | val transfer_start = thm "transfer_start" | |
| 40 | ||
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17333diff
changeset | 41 | fun unstar_typ (Type ("StarDef.star",[t])) = unstar_typ t
 | 
| 17329 | 42 | | unstar_typ (Type (a, Ts)) = Type (a, map unstar_typ Ts) | 
| 43 | | unstar_typ T = T | |
| 44 | ||
| 45 | fun unstar_term consts term = | |
| 46 | let | |
| 19734 | 47 | fun unstar (Const(a,T) $ t) = if member (op =) consts a then (unstar t) | 
| 17329 | 48 | else (Const(a,unstar_typ T) $ unstar t) | 
| 49 | | unstar (f $ t) = unstar f $ unstar t | |
| 50 | | unstar (Const(a,T)) = Const(a,unstar_typ T) | |
| 51 | | unstar (Abs(a,T,t)) = Abs(a,unstar_typ T,unstar t) | |
| 52 | | unstar t = t | |
| 53 | in | |
| 54 | unstar term | |
| 55 | end | |
| 56 | ||
| 18729 | 57 | fun transfer_thm_of ctxt ths t = | 
| 17329 | 58 | let | 
| 18729 | 59 | val thy = ProofContext.theory_of ctxt; | 
| 22739 
d12a9d75eee6
transfer_tac accepts also HOL equations as theorems
 haftmann parents: 
21708diff
changeset | 60 |     val {intros,unfolds,refolds,consts} = TransferData.get (Context.Proof ctxt);
 | 
| 
d12a9d75eee6
transfer_tac accepts also HOL equations as theorems
 haftmann parents: 
21708diff
changeset | 61 | val meta = LocalDefs.meta_rewrite_rule ctxt; | 
| 
d12a9d75eee6
transfer_tac accepts also HOL equations as theorems
 haftmann parents: 
21708diff
changeset | 62 | val ths' = map meta ths; | 
| 19734 | 63 | val unfolds' = map meta unfolds and refolds' = map meta refolds; | 
| 21708 | 64 | val (_$_$t') = concl_of (MetaSimplifier.rewrite true unfolds' (cterm_of thy t)) | 
| 17329 | 65 | val u = unstar_term consts t' | 
| 18729 | 66 | val tac = | 
| 22739 
d12a9d75eee6
transfer_tac accepts also HOL equations as theorems
 haftmann parents: 
21708diff
changeset | 67 | rewrite_goals_tac (ths' @ refolds' @ unfolds') THEN | 
| 19734 | 68 | ALLGOALS ObjectLogic.full_atomize_tac THEN | 
| 18729 | 69 | match_tac [transitive_thm] 1 THEN | 
| 70 | resolve_tac [transfer_start] 1 THEN | |
| 71 | REPEAT_ALL_NEW (resolve_tac intros) 1 THEN | |
| 72 | match_tac [reflexive_thm] 1 | |
| 20049 | 73 | in Goal.prove ctxt [] [] (Logic.mk_equals (t,u)) (K tac) end | 
| 17329 | 74 | |
| 18729 | 75 | fun transfer_tac ctxt ths = | 
| 17329 | 76 | SUBGOAL (fn (t,i) => | 
| 77 | (fn th => | |
| 18729 | 78 | let val tr = transfer_thm_of ctxt ths t | 
| 79 | in rewrite_goals_tac [tr] th end)) | |
| 17329 | 80 | |
| 81 | local | |
| 19734 | 82 | |
| 17329 | 83 | fun map_intros f = TransferData.map | 
| 84 |   (fn {intros,unfolds,refolds,consts} =>
 | |
| 85 |     {intros=f intros, unfolds=unfolds, refolds=refolds, consts=consts})
 | |
| 86 | ||
| 87 | fun map_unfolds f = TransferData.map | |
| 88 |   (fn {intros,unfolds,refolds,consts} =>
 | |
| 89 |     {intros=intros, unfolds=f unfolds, refolds=refolds, consts=consts})
 | |
| 90 | ||
| 91 | fun map_refolds f = TransferData.map | |
| 92 |   (fn {intros,unfolds,refolds,consts} =>
 | |
| 93 |     {intros=intros, unfolds=unfolds, refolds=f refolds, consts=consts})
 | |
| 94 | in | |
| 18729 | 95 | val intro_add = Thm.declaration_attribute (map_intros o Drule.add_rule); | 
| 96 | val intro_del = Thm.declaration_attribute (map_intros o Drule.del_rule); | |
| 17329 | 97 | |
| 18729 | 98 | val unfold_add = Thm.declaration_attribute (map_unfolds o Drule.add_rule); | 
| 99 | val unfold_del = Thm.declaration_attribute (map_unfolds o Drule.del_rule); | |
| 17329 | 100 | |
| 18729 | 101 | val refold_add = Thm.declaration_attribute (map_refolds o Drule.add_rule); | 
| 102 | val refold_del = Thm.declaration_attribute (map_refolds o Drule.del_rule); | |
| 17329 | 103 | end | 
| 104 | ||
| 18729 | 105 | fun add_const c = Context.theory_map (TransferData.map | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17333diff
changeset | 106 |   (fn {intros,unfolds,refolds,consts} =>
 | 
| 18729 | 107 |     {intros=intros, unfolds=unfolds, refolds=refolds, consts=c::consts}))
 | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17333diff
changeset | 108 | |
| 17329 | 109 | val setup = | 
| 18708 | 110 | Attrib.add_attributes | 
| 18729 | 111 |     [("transfer_intro", Attrib.add_del_args intro_add intro_del,
 | 
| 18708 | 112 | "declaration of transfer introduction rule"), | 
| 18729 | 113 |      ("transfer_unfold", Attrib.add_del_args unfold_add unfold_del,
 | 
| 18708 | 114 | "declaration of transfer unfolding rule"), | 
| 18729 | 115 |      ("transfer_refold", Attrib.add_del_args refold_add refold_del,
 | 
| 18708 | 116 | "declaration of transfer refolding rule")] #> | 
| 21588 | 117 |   Method.add_method ("transfer", Method.thms_ctxt_args (fn ths => fn ctxt =>
 | 
| 118 | Method.SIMPLE_METHOD' (transfer_tac ctxt ths)), "transfer principle"); | |
| 17329 | 119 | |
| 120 | end; |