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