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