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