author | wenzelm |
Sat, 18 Aug 2007 13:32:28 +0200 | |
changeset 24325 | 5c29e8822f50 |
parent 24121 | a93b0f4df838 |
permissions | -rw-r--r-- |
17333 | 1 |
(* Title : HOL/Hyperreal/transfer.ML |
2 |
ID : $Id$ |
|
3 |
Author : Brian Huffman |
|
4 |
||
18729 | 5 |
Transfer principle tactic for nonstandard analysis. |
17333 | 6 |
*) |
7 |
||
18729 | 8 |
signature TRANSFER = |
17329 | 9 |
sig |
18729 | 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 | 12 |
val setup: theory -> theory |
17329 | 13 |
end; |
14 |
||
18729 | 15 |
structure Transfer: TRANSFER = |
17329 | 16 |
struct |
17 |
||
18729 | 18 |
structure TransferData = GenericDataFun |
22846 | 19 |
( |
17329 | 20 |
type T = { |
21 |
intros: thm list, |
|
22 |
unfolds: thm list, |
|
23 |
refolds: thm list, |
|
24 |
consts: string list |
|
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 | 27 |
val extend = I; |
28 |
fun merge _ |
|
29 |
({intros = intros1, unfolds = unfolds1, |
|
30 |
refolds = refolds1, consts = consts1}, |
|
31 |
{intros = intros2, unfolds = unfolds2, |
|
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 | 36 |
consts = Library.merge (op =) (consts1, consts2)} : T; |
22846 | 37 |
); |
17329 | 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 | 40 |
| unstar_typ (Type (a, Ts)) = Type (a, map unstar_typ Ts) |
41 |
| unstar_typ T = T |
|
42 |
||
43 |
fun unstar_term consts term = |
|
44 |
let |
|
19734 | 45 |
fun unstar (Const(a,T) $ t) = if member (op =) consts a then (unstar t) |
17329 | 46 |
else (Const(a,unstar_typ T) $ unstar t) |
47 |
| unstar (f $ t) = unstar f $ unstar t |
|
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 | 50 |
| unstar t = t |
51 |
in |
|
52 |
unstar term |
|
53 |
end |
|
54 |
||
18729 | 55 |
fun transfer_thm_of ctxt ths t = |
17329 | 56 |
let |
18729 | 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 | 61 |
val unfolds' = map meta unfolds and refolds' = map meta refolds; |
21708 | 62 |
val (_$_$t') = concl_of (MetaSimplifier.rewrite true unfolds' (cterm_of thy t)) |
17329 | 63 |
val u = unstar_term consts t' |
18729 | 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 | 66 |
ALLGOALS ObjectLogic.full_atomize_tac THEN |
18729 | 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 | 69 |
REPEAT_ALL_NEW (resolve_tac intros) 1 THEN |
70 |
match_tac [reflexive_thm] 1 |
|
20049 | 71 |
in Goal.prove ctxt [] [] (Logic.mk_equals (t,u)) (K tac) end |
17329 | 72 |
|
18729 | 73 |
fun transfer_tac ctxt ths = |
17329 | 74 |
SUBGOAL (fn (t,i) => |
75 |
(fn th => |
|
24121 | 76 |
let |
77 |
val tr = transfer_thm_of ctxt ths t |
|
78 |
val (_$l$r) = concl_of tr; |
|
79 |
val trs = if l aconv r then [] else [tr]; |
|
80 |
in rewrite_goals_tac trs th end)) |
|
17329 | 81 |
|
82 |
local |
|
19734 | 83 |
|
17329 | 84 |
fun map_intros f = TransferData.map |
85 |
(fn {intros,unfolds,refolds,consts} => |
|
86 |
{intros=f intros, unfolds=unfolds, refolds=refolds, consts=consts}) |
|
87 |
||
88 |
fun map_unfolds f = TransferData.map |
|
89 |
(fn {intros,unfolds,refolds,consts} => |
|
90 |
{intros=intros, unfolds=f unfolds, refolds=refolds, consts=consts}) |
|
91 |
||
92 |
fun map_refolds f = TransferData.map |
|
93 |
(fn {intros,unfolds,refolds,consts} => |
|
94 |
{intros=intros, unfolds=unfolds, refolds=f refolds, consts=consts}) |
|
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 | 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 | 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 | 104 |
end |
105 |
||
18729 | 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 | 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 | 110 |
val setup = |
18708 | 111 |
Attrib.add_attributes |
18729 | 112 |
[("transfer_intro", Attrib.add_del_args intro_add intro_del, |
18708 | 113 |
"declaration of transfer introduction rule"), |
18729 | 114 |
("transfer_unfold", Attrib.add_del_args unfold_add unfold_del, |
18708 | 115 |
"declaration of transfer unfolding rule"), |
18729 | 116 |
("transfer_refold", Attrib.add_del_args refold_add refold_del, |
18708 | 117 |
"declaration of transfer refolding rule")] #> |
21588 | 118 |
Method.add_method ("transfer", Method.thms_ctxt_args (fn ths => fn ctxt => |
119 |
Method.SIMPLE_METHOD' (transfer_tac ctxt ths)), "transfer principle"); |
|
17329 | 120 |
|
121 |
end; |