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