author | wenzelm |
Sat, 24 Mar 2018 20:51:42 +0100 | |
changeset 67945 | 984c3dc46cc0 |
parent 67636 | e4eb21f8331c |
child 69597 | ff784d5a5bfb |
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 = []}; |
|
25 |
val extend = I; |
|
33519 | 26 |
fun merge |
27468 | 27 |
({intros = intros1, unfolds = unfolds1, |
28 |
refolds = refolds1, consts = consts1}, |
|
29 |
{intros = intros2, unfolds = unfolds2, |
|
33519 | 30 |
refolds = refolds2, consts = consts2}) : T = |
27468 | 31 |
{intros = Thm.merge_thms (intros1, intros2), |
32 |
unfolds = Thm.merge_thms (unfolds1, unfolds2), |
|
33 |
refolds = Thm.merge_thms (refolds1, refolds2), |
|
33519 | 34 |
consts = Library.merge (op =) (consts1, consts2)}; |
27468 | 35 |
); |
36 |
||
56256 | 37 |
fun unstar_typ (Type (@{type_name star}, [t])) = unstar_typ t |
27468 | 38 |
| unstar_typ (Type (a, Ts)) = Type (a, map unstar_typ Ts) |
39 |
| unstar_typ T = T |
|
40 |
||
41 |
fun unstar_term consts term = |
|
42 |
let |
|
43 |
fun unstar (Const(a,T) $ t) = if member (op =) consts a then (unstar t) |
|
44 |
else (Const(a,unstar_typ T) $ unstar t) |
|
45 |
| unstar (f $ t) = unstar f $ unstar t |
|
46 |
| unstar (Const(a,T)) = Const(a,unstar_typ T) |
|
47 |
| unstar (Abs(a,T,t)) = Abs(a,unstar_typ T,unstar t) |
|
48 |
| unstar t = t |
|
49 |
in |
|
50 |
unstar term |
|
51 |
end |
|
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 |
bf474d719011
Modified transfer principle in HOL/NSA to cause less ho-unficiation
Simon Wimmer <wimmers@in.tum.de>
parents:
62479
diff
changeset
|
57 |
fun thm_of (Const (@{const_name Ifun}, _) $ t $ u) = @{thm transfer_Ifun} OF [thm_of t, thm_of u] |
64435 | 58 |
| thm_of (Const (@{const_name star_of}, _) $ _) = @{thm star_of_def} |
59 |
| thm_of (Const (@{const_name star_n}, _) $ _) = @{thm Pure.reflexive} |
|
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 |
|
bf474d719011
Modified transfer principle in HOL/NSA to cause less ho-unficiation
Simon Wimmer <wimmers@in.tum.de>
parents:
62479
diff
changeset
|
62 |
fun thm_of_goal (Const (@{const_name Pure.eq}, _) $ t $ (Const (@{const_name star_n}, _) $ _)) = |
64435 | 63 |
thm_of t |
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 | 72 |
fun transfer_thm_of ctxt ths t = |
73 |
let |
|
64435 | 74 |
val {intros,unfolds,refolds,consts} = Data.get (Context.Proof ctxt); |
35624 | 75 |
val meta = Local_Defs.meta_rewrite_rule ctxt; |
27468 | 76 |
val ths' = map meta ths; |
77 |
val unfolds' = map meta unfolds and refolds' = map meta refolds; |
|
59643 | 78 |
val (_$_$t') = Thm.concl_of (Raw_Simplifier.rewrite ctxt true unfolds' (Thm.cterm_of ctxt t)) |
27468 | 79 |
val u = unstar_term consts t' |
64435 | 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 | 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 | 86 |
match_tac ctxt [reflexive_thm] 1 |
64435 | 87 |
in Goal.prove ctxt [] [] (Logic.mk_equals (t,u)) (K tac) end; |
27468 | 88 |
|
64435 | 89 |
fun transfer_tac ctxt ths = SUBGOAL (fn (t, _) => (fn th => |
90 |
let |
|
91 |
val tr = transfer_thm_of ctxt ths t |
|
92 |
val (_$ l $ r) = Thm.concl_of tr; |
|
93 |
val trs = if l aconv r then [] else [tr]; |
|
94 |
in rewrite_goals_tac ctxt trs th end)); |
|
27468 | 95 |
|
96 |
local |
|
97 |
||
64435 | 98 |
fun map_intros f = Data.map |
27468 | 99 |
(fn {intros,unfolds,refolds,consts} => |
100 |
{intros=f intros, unfolds=unfolds, refolds=refolds, consts=consts}) |
|
101 |
||
64435 | 102 |
fun map_unfolds f = Data.map |
27468 | 103 |
(fn {intros,unfolds,refolds,consts} => |
104 |
{intros=intros, unfolds=f unfolds, refolds=refolds, consts=consts}) |
|
105 |
||
64435 | 106 |
fun map_refolds f = Data.map |
27468 | 107 |
(fn {intros,unfolds,refolds,consts} => |
108 |
{intros=intros, unfolds=unfolds, refolds=f refolds, consts=consts}) |
|
109 |
in |
|
64435 | 110 |
|
67636 | 111 |
val intro_add = Thm.declaration_attribute (map_intros o Thm.add_thm o Thm.trim_context); |
27468 | 112 |
val intro_del = Thm.declaration_attribute (map_intros o Thm.del_thm); |
113 |
||
67636 | 114 |
val unfold_add = Thm.declaration_attribute (map_unfolds o Thm.add_thm o Thm.trim_context); |
27468 | 115 |
val unfold_del = Thm.declaration_attribute (map_unfolds o Thm.del_thm); |
116 |
||
67636 | 117 |
val refold_add = Thm.declaration_attribute (map_refolds o Thm.add_thm o Thm.trim_context); |
27468 | 118 |
val refold_del = Thm.declaration_attribute (map_refolds o Thm.del_thm); |
64435 | 119 |
|
27468 | 120 |
end |
121 |
||
64435 | 122 |
fun add_const c = Context.theory_map (Data.map |
27468 | 123 |
(fn {intros,unfolds,refolds,consts} => |
124 |
{intros=intros, unfolds=unfolds, refolds=refolds, consts=c::consts})) |
|
125 |
||
58825 | 126 |
val _ = |
127 |
Theory.setup |
|
128 |
(Attrib.setup @{binding transfer_intro} (Attrib.add_del intro_add intro_del) |
|
129 |
"declaration of transfer introduction rule" #> |
|
130 |
Attrib.setup @{binding transfer_unfold} (Attrib.add_del unfold_add unfold_del) |
|
131 |
"declaration of transfer unfolding rule" #> |
|
132 |
Attrib.setup @{binding transfer_refold} (Attrib.add_del refold_add refold_del) |
|
133 |
"declaration of transfer refolding rule") |
|
27468 | 134 |
|
135 |
end; |