src/HOL/Nonstandard_Analysis/transfer_principle.ML
author wenzelm
Sat, 17 Apr 2021 19:45:12 +0200
changeset 73589 479e9b17090e
parent 69597 ff784d5a5bfb
child 74561 8e6c973003c8
permissions -rw-r--r--
clarified options (again);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
67635
28f926146986 proper file name;
wenzelm
parents: 64435
diff changeset
     1
(*  Title:      HOL/Nonstandard_Analysis/transfer_principle.ML
37744
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
end;
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    12
47328
9f11a3cd84b1 rename ML structure to avoid shadowing earlier name
huffman
parents: 42361
diff changeset
    13
structure Transfer_Principle: TRANSFER_PRINCIPLE =
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    14
struct
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    15
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 64270
diff changeset
    16
structure Data = Generic_Data
27468
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
  type T = {
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    19
    intros: thm list,
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    20
    unfolds: thm list,
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    21
    refolds: thm list,
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    22
    consts: string list
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    23
  };
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    24
  val empty = {intros = [], unfolds = [], refolds = [], consts = []};
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    25
  val extend = I;
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 30528
diff changeset
    26
  fun merge
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    27
    ({intros = intros1, unfolds = unfolds1,
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    28
      refolds = refolds1, consts = consts1},
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    29
     {intros = intros2, unfolds = unfolds2,
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 30528
diff changeset
    30
      refolds = refolds2, consts = consts2}) : T =
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    31
   {intros = Thm.merge_thms (intros1, intros2),
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    32
    unfolds = Thm.merge_thms (unfolds1, unfolds2),
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    33
    refolds = Thm.merge_thms (refolds1, refolds2),
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 30528
diff changeset
    34
    consts = Library.merge (op =) (consts1, consts2)};
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    35
);
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    36
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67636
diff changeset
    37
fun unstar_typ (Type (\<^type_name>\<open>star\<close>, [t])) = unstar_typ t
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    38
  | unstar_typ (Type (a, Ts)) = Type (a, map unstar_typ Ts)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    39
  | unstar_typ T = T
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    40
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    41
fun unstar_term consts term =
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    42
  let
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    43
    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
    44
          else (Const(a,unstar_typ T) $ unstar t)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    45
      | unstar (f $ t) = unstar f $ unstar t
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    46
      | unstar (Const(a,T)) = Const(a,unstar_typ T)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    47
      | unstar (Abs(a,T,t)) = Abs(a,unstar_typ T,unstar t)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    48
      | unstar t = t
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    49
  in
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    50
    unstar term
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    51
  end
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    52
64270
bf474d719011 Modified transfer principle in HOL/NSA to cause less ho-unficiation
Simon Wimmer <wimmers@in.tum.de>
parents: 62479
diff changeset
    53
local exception MATCH
bf474d719011 Modified transfer principle in HOL/NSA to cause less ho-unficiation
Simon Wimmer <wimmers@in.tum.de>
parents: 62479
diff changeset
    54
in
bf474d719011 Modified transfer principle in HOL/NSA to cause less ho-unficiation
Simon Wimmer <wimmers@in.tum.de>
parents: 62479
diff changeset
    55
fun transfer_star_tac ctxt =
bf474d719011 Modified transfer principle in HOL/NSA to cause less ho-unficiation
Simon Wimmer <wimmers@in.tum.de>
parents: 62479
diff changeset
    56
  let
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67636
diff changeset
    57
    fun thm_of (Const (\<^const_name>\<open>Ifun\<close>, _) $ t $ u) = @{thm transfer_Ifun} OF [thm_of t, thm_of u]
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67636
diff changeset
    58
      | thm_of (Const (\<^const_name>\<open>star_of\<close>, _) $ _) = @{thm star_of_def}
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67636
diff changeset
    59
      | thm_of (Const (\<^const_name>\<open>star_n\<close>, _) $ _) = @{thm Pure.reflexive}
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 64270
diff changeset
    60
      | thm_of _ = raise MATCH;
64270
bf474d719011 Modified transfer principle in HOL/NSA to cause less ho-unficiation
Simon Wimmer <wimmers@in.tum.de>
parents: 62479
diff changeset
    61
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67636
diff changeset
    62
    fun thm_of_goal (Const (\<^const_name>\<open>Pure.eq\<close>, _) $ t $ (Const (\<^const_name>\<open>star_n\<close>, _) $ _)) =
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 64270
diff changeset
    63
          thm_of t
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 64270
diff changeset
    64
      | thm_of_goal _ = raise MATCH;
64270
bf474d719011 Modified transfer principle in HOL/NSA to cause less ho-unficiation
Simon Wimmer <wimmers@in.tum.de>
parents: 62479
diff changeset
    65
  in
bf474d719011 Modified transfer principle in HOL/NSA to cause less ho-unficiation
Simon Wimmer <wimmers@in.tum.de>
parents: 62479
diff changeset
    66
    SUBGOAL (fn (t, i) =>
bf474d719011 Modified transfer principle in HOL/NSA to cause less ho-unficiation
Simon Wimmer <wimmers@in.tum.de>
parents: 62479
diff changeset
    67
      resolve_tac ctxt [thm_of_goal (strip_all_body t |> Logic.strip_imp_concl)] i
bf474d719011 Modified transfer principle in HOL/NSA to cause less ho-unficiation
Simon Wimmer <wimmers@in.tum.de>
parents: 62479
diff changeset
    68
      handle MATCH => no_tac)
bf474d719011 Modified transfer principle in HOL/NSA to cause less ho-unficiation
Simon Wimmer <wimmers@in.tum.de>
parents: 62479
diff changeset
    69
  end;
bf474d719011 Modified transfer principle in HOL/NSA to cause less ho-unficiation
Simon Wimmer <wimmers@in.tum.de>
parents: 62479
diff changeset
    70
end;
bf474d719011 Modified transfer principle in HOL/NSA to cause less ho-unficiation
Simon Wimmer <wimmers@in.tum.de>
parents: 62479
diff changeset
    71
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    72
fun transfer_thm_of ctxt ths t =
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    73
  let
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 64270
diff changeset
    74
    val {intros,unfolds,refolds,consts} = Data.get (Context.Proof ctxt);
35624
c4e29a0bb8c1 modernized structure Local_Defs;
wenzelm
parents: 33519
diff changeset
    75
    val meta = Local_Defs.meta_rewrite_rule ctxt;
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    76
    val ths' = map meta ths;
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    77
    val unfolds' = map meta unfolds and refolds' = map meta refolds;
59643
f3be9235503d clarified context;
wenzelm
parents: 59621
diff changeset
    78
    val (_$_$t') = Thm.concl_of (Raw_Simplifier.rewrite ctxt true unfolds' (Thm.cterm_of ctxt t))
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    79
    val u = unstar_term consts t'
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 64270
diff changeset
    80
    val tac =
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 47432
diff changeset
    81
      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
    82
      ALLGOALS (Object_Logic.full_atomize_tac ctxt) THEN
58957
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58825
diff changeset
    83
      match_tac ctxt [transitive_thm] 1 THEN
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58957
diff changeset
    84
      resolve_tac ctxt [@{thm transfer_start}] 1 THEN
64270
bf474d719011 Modified transfer principle in HOL/NSA to cause less ho-unficiation
Simon Wimmer <wimmers@in.tum.de>
parents: 62479
diff changeset
    85
      REPEAT_ALL_NEW (resolve_tac ctxt intros ORELSE' transfer_star_tac ctxt) 1 THEN
58957
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58825
diff changeset
    86
      match_tac ctxt [reflexive_thm] 1
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 64270
diff changeset
    87
  in Goal.prove ctxt [] [] (Logic.mk_equals (t,u)) (K tac) end;
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    88
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 64270
diff changeset
    89
fun transfer_tac ctxt ths = SUBGOAL (fn (t, _) => (fn th =>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 64270
diff changeset
    90
  let
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 64270
diff changeset
    91
    val tr = transfer_thm_of ctxt ths t
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 64270
diff changeset
    92
    val (_$ l $ r) = Thm.concl_of tr;
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 64270
diff changeset
    93
    val trs = if l aconv r then [] else [tr];
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 64270
diff changeset
    94
  in rewrite_goals_tac ctxt trs th end));
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    95
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    96
local
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    97
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 64270
diff changeset
    98
fun map_intros f = Data.map
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    99
  (fn {intros,unfolds,refolds,consts} =>
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   100
    {intros=f intros, unfolds=unfolds, refolds=refolds, consts=consts})
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   101
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 64270
diff changeset
   102
fun map_unfolds f = Data.map
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   103
  (fn {intros,unfolds,refolds,consts} =>
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   104
    {intros=intros, unfolds=f unfolds, refolds=refolds, consts=consts})
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   105
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 64270
diff changeset
   106
fun map_refolds f = Data.map
27468
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=f refolds, consts=consts})
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   109
in
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 64270
diff changeset
   110
67636
e4eb21f8331c trim context of persistent data;
wenzelm
parents: 67635
diff changeset
   111
val intro_add = Thm.declaration_attribute (map_intros o Thm.add_thm o Thm.trim_context);
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   112
val intro_del = Thm.declaration_attribute (map_intros o Thm.del_thm);
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   113
67636
e4eb21f8331c trim context of persistent data;
wenzelm
parents: 67635
diff changeset
   114
val unfold_add = Thm.declaration_attribute (map_unfolds o Thm.add_thm o Thm.trim_context);
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   115
val unfold_del = Thm.declaration_attribute (map_unfolds o Thm.del_thm);
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   116
67636
e4eb21f8331c trim context of persistent data;
wenzelm
parents: 67635
diff changeset
   117
val refold_add = Thm.declaration_attribute (map_refolds o Thm.add_thm o Thm.trim_context);
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   118
val refold_del = Thm.declaration_attribute (map_refolds o Thm.del_thm);
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 64270
diff changeset
   119
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   120
end
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   121
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 64270
diff changeset
   122
fun add_const c = Context.theory_map (Data.map
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   123
  (fn {intros,unfolds,refolds,consts} =>
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   124
    {intros=intros, unfolds=unfolds, refolds=refolds, consts=c::consts}))
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   125
58825
2065f49da190 modernized setup;
wenzelm
parents: 56256
diff changeset
   126
val _ =
2065f49da190 modernized setup;
wenzelm
parents: 56256
diff changeset
   127
  Theory.setup
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67636
diff changeset
   128
   (Attrib.setup \<^binding>\<open>transfer_intro\<close> (Attrib.add_del intro_add intro_del)
58825
2065f49da190 modernized setup;
wenzelm
parents: 56256
diff changeset
   129
      "declaration of transfer introduction rule" #>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67636
diff changeset
   130
    Attrib.setup \<^binding>\<open>transfer_unfold\<close> (Attrib.add_del unfold_add unfold_del)
58825
2065f49da190 modernized setup;
wenzelm
parents: 56256
diff changeset
   131
      "declaration of transfer unfolding rule" #>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67636
diff changeset
   132
    Attrib.setup \<^binding>\<open>transfer_refold\<close> (Attrib.add_del refold_add refold_del)
58825
2065f49da190 modernized setup;
wenzelm
parents: 56256
diff changeset
   133
      "declaration of transfer refolding rule")
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   134
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   135
end;