| author | wenzelm | 
| Wed, 28 Jun 2023 12:25:54 +0200 | |
| changeset 78224 | d85d0d41b2bd | 
| parent 74561 | 8e6c973003c8 | 
| child 82641 | d22294b20573 | 
| permissions | -rw-r--r-- | 
| 67635 | 1  | 
(* Title: HOL/Nonstandard_Analysis/transfer_principle.ML  | 
| 37744 | 2  | 
Author: Brian Huffman  | 
| 27468 | 3  | 
|
4  | 
Transfer principle tactic for nonstandard analysis.  | 
|
5  | 
*)  | 
|
6  | 
||
| 
47328
 
9f11a3cd84b1
rename ML structure to avoid shadowing earlier name
 
huffman 
parents: 
42361 
diff
changeset
 | 
7  | 
signature TRANSFER_PRINCIPLE =  | 
| 27468 | 8  | 
sig  | 
9  | 
val transfer_tac: Proof.context -> thm list -> int -> tactic  | 
|
10  | 
val add_const: string -> theory -> theory  | 
|
11  | 
end;  | 
|
12  | 
||
| 
47328
 
9f11a3cd84b1
rename ML structure to avoid shadowing earlier name
 
huffman 
parents: 
42361 
diff
changeset
 | 
13  | 
structure Transfer_Principle: TRANSFER_PRINCIPLE =  | 
| 27468 | 14  | 
struct  | 
15  | 
||
| 64435 | 16  | 
structure Data = Generic_Data  | 
| 27468 | 17  | 
(  | 
18  | 
  type T = {
 | 
|
19  | 
intros: thm list,  | 
|
20  | 
unfolds: thm list,  | 
|
21  | 
refolds: thm list,  | 
|
22  | 
consts: string list  | 
|
23  | 
};  | 
|
24  | 
  val empty = {intros = [], unfolds = [], refolds = [], consts = []};
 | 
|
| 33519 | 25  | 
fun merge  | 
| 27468 | 26  | 
    ({intros = intros1, unfolds = unfolds1,
 | 
27  | 
refolds = refolds1, consts = consts1},  | 
|
28  | 
     {intros = intros2, unfolds = unfolds2,
 | 
|
| 33519 | 29  | 
refolds = refolds2, consts = consts2}) : T =  | 
| 27468 | 30  | 
   {intros = Thm.merge_thms (intros1, intros2),
 | 
31  | 
unfolds = Thm.merge_thms (unfolds1, unfolds2),  | 
|
32  | 
refolds = Thm.merge_thms (refolds1, refolds2),  | 
|
| 33519 | 33  | 
consts = Library.merge (op =) (consts1, consts2)};  | 
| 27468 | 34  | 
);  | 
35  | 
||
| 69597 | 36  | 
fun unstar_typ (Type (\<^type_name>\<open>star\<close>, [t])) = unstar_typ t  | 
| 27468 | 37  | 
| unstar_typ (Type (a, Ts)) = Type (a, map unstar_typ Ts)  | 
38  | 
| unstar_typ T = T  | 
|
39  | 
||
40  | 
fun unstar_term consts term =  | 
|
41  | 
let  | 
|
42  | 
fun unstar (Const(a,T) $ t) = if member (op =) consts a then (unstar t)  | 
|
43  | 
else (Const(a,unstar_typ T) $ unstar t)  | 
|
44  | 
| unstar (f $ t) = unstar f $ unstar t  | 
|
45  | 
| unstar (Const(a,T)) = Const(a,unstar_typ T)  | 
|
46  | 
| unstar (Abs(a,T,t)) = Abs(a,unstar_typ T,unstar t)  | 
|
47  | 
| unstar t = t  | 
|
48  | 
in  | 
|
49  | 
unstar term  | 
|
50  | 
end  | 
|
51  | 
||
| 
64270
 
bf474d719011
Modified transfer principle in HOL/NSA to cause less ho-unficiation
 
Simon Wimmer <wimmers@in.tum.de> 
parents: 
62479 
diff
changeset
 | 
52  | 
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
 | 
53  | 
in  | 
| 
 
bf474d719011
Modified transfer principle in HOL/NSA to cause less ho-unficiation
 
Simon Wimmer <wimmers@in.tum.de> 
parents: 
62479 
diff
changeset
 | 
54  | 
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
 | 
55  | 
let  | 
| 69597 | 56  | 
    fun thm_of (Const (\<^const_name>\<open>Ifun\<close>, _) $ t $ u) = @{thm transfer_Ifun} OF [thm_of t, thm_of u]
 | 
57  | 
      | thm_of (Const (\<^const_name>\<open>star_of\<close>, _) $ _) = @{thm star_of_def}
 | 
|
58  | 
      | thm_of (Const (\<^const_name>\<open>star_n\<close>, _) $ _) = @{thm Pure.reflexive}
 | 
|
| 64435 | 59  | 
| 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
 | 
60  | 
|
| 69597 | 61  | 
fun thm_of_goal (Const (\<^const_name>\<open>Pure.eq\<close>, _) $ t $ (Const (\<^const_name>\<open>star_n\<close>, _) $ _)) =  | 
| 64435 | 62  | 
thm_of t  | 
63  | 
| 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
 | 
64  | 
in  | 
| 
 
bf474d719011
Modified transfer principle in HOL/NSA to cause less ho-unficiation
 
Simon Wimmer <wimmers@in.tum.de> 
parents: 
62479 
diff
changeset
 | 
65  | 
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
 | 
66  | 
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
 | 
67  | 
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
 | 
68  | 
end;  | 
| 
 
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  | 
|
| 27468 | 71  | 
fun transfer_thm_of ctxt ths t =  | 
72  | 
let  | 
|
| 64435 | 73  | 
    val {intros,unfolds,refolds,consts} = Data.get (Context.Proof ctxt);
 | 
| 35624 | 74  | 
val meta = Local_Defs.meta_rewrite_rule ctxt;  | 
| 27468 | 75  | 
val ths' = map meta ths;  | 
76  | 
val unfolds' = map meta unfolds and refolds' = map meta refolds;  | 
|
| 59643 | 77  | 
val (_$_$t') = Thm.concl_of (Raw_Simplifier.rewrite ctxt true unfolds' (Thm.cterm_of ctxt t))  | 
| 27468 | 78  | 
val u = unstar_term consts t'  | 
| 64435 | 79  | 
val tac =  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
47432 
diff
changeset
 | 
80  | 
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
 | 
81  | 
ALLGOALS (Object_Logic.full_atomize_tac ctxt) THEN  | 
| 58957 | 82  | 
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
 | 
83  | 
      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
 | 
84  | 
REPEAT_ALL_NEW (resolve_tac ctxt intros ORELSE' transfer_star_tac ctxt) 1 THEN  | 
| 58957 | 85  | 
match_tac ctxt [reflexive_thm] 1  | 
| 64435 | 86  | 
in Goal.prove ctxt [] [] (Logic.mk_equals (t,u)) (K tac) end;  | 
| 27468 | 87  | 
|
| 64435 | 88  | 
fun transfer_tac ctxt ths = SUBGOAL (fn (t, _) => (fn th =>  | 
89  | 
let  | 
|
90  | 
val tr = transfer_thm_of ctxt ths t  | 
|
91  | 
val (_$ l $ r) = Thm.concl_of tr;  | 
|
92  | 
val trs = if l aconv r then [] else [tr];  | 
|
93  | 
in rewrite_goals_tac ctxt trs th end));  | 
|
| 27468 | 94  | 
|
95  | 
local  | 
|
96  | 
||
| 64435 | 97  | 
fun map_intros f = Data.map  | 
| 27468 | 98  | 
  (fn {intros,unfolds,refolds,consts} =>
 | 
99  | 
    {intros=f intros, unfolds=unfolds, refolds=refolds, consts=consts})
 | 
|
100  | 
||
| 64435 | 101  | 
fun map_unfolds f = Data.map  | 
| 27468 | 102  | 
  (fn {intros,unfolds,refolds,consts} =>
 | 
103  | 
    {intros=intros, unfolds=f unfolds, refolds=refolds, consts=consts})
 | 
|
104  | 
||
| 64435 | 105  | 
fun map_refolds f = Data.map  | 
| 27468 | 106  | 
  (fn {intros,unfolds,refolds,consts} =>
 | 
107  | 
    {intros=intros, unfolds=unfolds, refolds=f refolds, consts=consts})
 | 
|
108  | 
in  | 
|
| 64435 | 109  | 
|
| 67636 | 110  | 
val intro_add = Thm.declaration_attribute (map_intros o Thm.add_thm o Thm.trim_context);  | 
| 27468 | 111  | 
val intro_del = Thm.declaration_attribute (map_intros o Thm.del_thm);  | 
112  | 
||
| 67636 | 113  | 
val unfold_add = Thm.declaration_attribute (map_unfolds o Thm.add_thm o Thm.trim_context);  | 
| 27468 | 114  | 
val unfold_del = Thm.declaration_attribute (map_unfolds o Thm.del_thm);  | 
115  | 
||
| 67636 | 116  | 
val refold_add = Thm.declaration_attribute (map_refolds o Thm.add_thm o Thm.trim_context);  | 
| 27468 | 117  | 
val refold_del = Thm.declaration_attribute (map_refolds o Thm.del_thm);  | 
| 64435 | 118  | 
|
| 27468 | 119  | 
end  | 
120  | 
||
| 64435 | 121  | 
fun add_const c = Context.theory_map (Data.map  | 
| 27468 | 122  | 
  (fn {intros,unfolds,refolds,consts} =>
 | 
123  | 
    {intros=intros, unfolds=unfolds, refolds=refolds, consts=c::consts}))
 | 
|
124  | 
||
| 58825 | 125  | 
val _ =  | 
126  | 
Theory.setup  | 
|
| 69597 | 127  | 
(Attrib.setup \<^binding>\<open>transfer_intro\<close> (Attrib.add_del intro_add intro_del)  | 
| 58825 | 128  | 
"declaration of transfer introduction rule" #>  | 
| 69597 | 129  | 
Attrib.setup \<^binding>\<open>transfer_unfold\<close> (Attrib.add_del unfold_add unfold_del)  | 
| 58825 | 130  | 
"declaration of transfer unfolding rule" #>  | 
| 69597 | 131  | 
Attrib.setup \<^binding>\<open>transfer_refold\<close> (Attrib.add_del refold_add refold_del)  | 
| 58825 | 132  | 
"declaration of transfer refolding rule")  | 
| 27468 | 133  | 
|
134  | 
end;  |