src/HOL/NSA/transfer.ML
author haftmann
Sat, 05 Jul 2014 11:01:53 +0200
changeset 57514 bdc2c6b40bf2
parent 56256 1e01c159e7d9
child 58825 2065f49da190
permissions -rw-r--r--
prefer ac_simps collections over separate name bindings for add and mult
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37744
3daaf23b9ab4 tuned titles
haftmann
parents: 35625
diff changeset
     1
(*  Title:      HOL/NSA/transfer.ML
3daaf23b9ab4 tuned titles
haftmann
parents: 35625
diff changeset
     2
    Author:     Brian Huffman
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     3
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     4
Transfer principle tactic for nonstandard analysis.
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     5
*)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     6
47328
9f11a3cd84b1 rename ML structure to avoid shadowing earlier name
huffman
parents: 42361
diff changeset
     7
signature TRANSFER_PRINCIPLE =
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     8
sig
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     9
  val transfer_tac: Proof.context -> thm list -> int -> tactic
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    10
  val add_const: string -> theory -> theory
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    11
  val setup: theory -> theory
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    12
end;
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    13
47328
9f11a3cd84b1 rename ML structure to avoid shadowing earlier name
huffman
parents: 42361
diff changeset
    14
structure Transfer_Principle: TRANSFER_PRINCIPLE =
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    15
struct
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    16
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 30528
diff changeset
    17
structure TransferData = Generic_Data
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    18
(
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    19
  type T = {
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    20
    intros: thm list,
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    21
    unfolds: thm list,
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    22
    refolds: thm list,
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    23
    consts: string list
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    24
  };
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    25
  val empty = {intros = [], unfolds = [], refolds = [], consts = []};
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    26
  val extend = I;
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 30528
diff changeset
    27
  fun merge
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    28
    ({intros = intros1, unfolds = unfolds1,
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    29
      refolds = refolds1, consts = consts1},
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    30
     {intros = intros2, unfolds = unfolds2,
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 30528
diff changeset
    31
      refolds = refolds2, consts = consts2}) : T =
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    32
   {intros = Thm.merge_thms (intros1, intros2),
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    33
    unfolds = Thm.merge_thms (unfolds1, unfolds2),
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    34
    refolds = Thm.merge_thms (refolds1, refolds2),
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 30528
diff changeset
    35
    consts = Library.merge (op =) (consts1, consts2)};
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    36
);
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    37
56256
1e01c159e7d9 more antiquotations;
wenzelm
parents: 54742
diff changeset
    38
fun unstar_typ (Type (@{type_name star}, [t])) = unstar_typ t
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    39
  | unstar_typ (Type (a, Ts)) = Type (a, map unstar_typ Ts)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    40
  | unstar_typ T = T
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    41
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    42
fun unstar_term consts term =
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    43
  let
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    44
    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
    45
          else (Const(a,unstar_typ T) $ unstar t)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    46
      | unstar (f $ t) = unstar f $ unstar t
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    47
      | unstar (Const(a,T)) = Const(a,unstar_typ T)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    48
      | unstar (Abs(a,T,t)) = Abs(a,unstar_typ T,unstar t)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    49
      | unstar t = t
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    50
  in
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    51
    unstar term
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    52
  end
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    53
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    54
fun transfer_thm_of ctxt ths t =
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    55
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41228
diff changeset
    56
    val thy = Proof_Context.theory_of ctxt;
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    57
    val {intros,unfolds,refolds,consts} = TransferData.get (Context.Proof ctxt);
35624
c4e29a0bb8c1 modernized structure Local_Defs;
wenzelm
parents: 33519
diff changeset
    58
    val meta = Local_Defs.meta_rewrite_rule ctxt;
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    59
    val ths' = map meta ths;
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    60
    val unfolds' = map meta unfolds and refolds' = map meta refolds;
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 47432
diff changeset
    61
    val (_$_$t') = concl_of (Raw_Simplifier.rewrite ctxt true unfolds' (cterm_of thy t))
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    62
    val u = unstar_term consts t'
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    63
    val tac =
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 47432
diff changeset
    64
      rewrite_goals_tac ctxt (ths' @ refolds' @ unfolds') THEN
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 47432
diff changeset
    65
      ALLGOALS (Object_Logic.full_atomize_tac ctxt) THEN
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    66
      match_tac [transitive_thm] 1 THEN
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    67
      resolve_tac [@{thm transfer_start}] 1 THEN
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    68
      REPEAT_ALL_NEW (resolve_tac intros) 1 THEN
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    69
      match_tac [reflexive_thm] 1
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    70
  in Goal.prove ctxt [] [] (Logic.mk_equals (t,u)) (K tac) end
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    71
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    72
fun transfer_tac ctxt ths =
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    73
    SUBGOAL (fn (t,i) =>
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    74
        (fn th =>
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    75
            let
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    76
              val tr = transfer_thm_of ctxt ths t
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    77
              val (_$l$r) = concl_of tr;
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    78
              val trs = if l aconv r then [] else [tr];
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 47432
diff changeset
    79
            in rewrite_goals_tac ctxt trs th end))
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    80
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    81
local
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    82
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    83
fun map_intros f = TransferData.map
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    84
  (fn {intros,unfolds,refolds,consts} =>
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    85
    {intros=f intros, unfolds=unfolds, refolds=refolds, consts=consts})
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    86
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    87
fun map_unfolds f = TransferData.map
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    88
  (fn {intros,unfolds,refolds,consts} =>
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    89
    {intros=intros, unfolds=f unfolds, refolds=refolds, consts=consts})
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    90
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    91
fun map_refolds f = TransferData.map
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    92
  (fn {intros,unfolds,refolds,consts} =>
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    93
    {intros=intros, unfolds=unfolds, refolds=f refolds, consts=consts})
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    94
in
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    95
val intro_add = Thm.declaration_attribute (map_intros o Thm.add_thm);
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    96
val intro_del = Thm.declaration_attribute (map_intros o Thm.del_thm);
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    97
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    98
val unfold_add = Thm.declaration_attribute (map_unfolds o Thm.add_thm);
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    99
val unfold_del = Thm.declaration_attribute (map_unfolds o Thm.del_thm);
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   100
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   101
val refold_add = Thm.declaration_attribute (map_refolds o Thm.add_thm);
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   102
val refold_del = Thm.declaration_attribute (map_refolds o Thm.del_thm);
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   103
end
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   104
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   105
fun add_const c = Context.theory_map (TransferData.map
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   106
  (fn {intros,unfolds,refolds,consts} =>
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   107
    {intros=intros, unfolds=unfolds, refolds=refolds, consts=c::consts}))
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   108
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   109
val setup =
30528
7173bf123335 simplified attribute setup;
wenzelm
parents: 30510
diff changeset
   110
  Attrib.setup @{binding transfer_intro} (Attrib.add_del intro_add intro_del)
7173bf123335 simplified attribute setup;
wenzelm
parents: 30510
diff changeset
   111
    "declaration of transfer introduction rule" #>
7173bf123335 simplified attribute setup;
wenzelm
parents: 30510
diff changeset
   112
  Attrib.setup @{binding transfer_unfold} (Attrib.add_del unfold_add unfold_del)
7173bf123335 simplified attribute setup;
wenzelm
parents: 30510
diff changeset
   113
    "declaration of transfer unfolding rule" #>
7173bf123335 simplified attribute setup;
wenzelm
parents: 30510
diff changeset
   114
  Attrib.setup @{binding transfer_refold} (Attrib.add_del refold_add refold_del)
47432
e1576d13e933 more standard method setup;
wenzelm
parents: 47328
diff changeset
   115
    "declaration of transfer refolding rule"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   116
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   117
end;