author | dixon |
Fri, 06 May 2005 18:01:44 +0200 | |
changeset 15936 | 817ac93ee786 |
parent 15929 | 68bd1e16552a |
child 15959 | 366d39e95d3c |
permissions | -rw-r--r-- |
15481 | 1 |
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
2 |
(* Title: Provers/eqsubst.ML |
15481 | 3 |
Author: Lucas Dixon, University of Edinburgh |
4 |
lucas.dixon@ed.ac.uk |
|
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
5 |
Modified: 18 Feb 2005 - Lucas - |
15481 | 6 |
Created: 29 Jan 2005 |
7 |
*) |
|
8 |
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) |
|
9 |
(* DESCRIPTION: |
|
10 |
||
11 |
A Tactic to perform a substiution using an equation. |
|
12 |
||
13 |
*) |
|
14 |
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) |
|
15 |
||
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
16 |
(* Logic specific data stub *) |
15481 | 17 |
signature EQRULE_DATA = |
18 |
sig |
|
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
19 |
|
15481 | 20 |
(* to make a meta equality theorem in the current logic *) |
21 |
val prep_meta_eq : thm -> thm list |
|
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
22 |
|
15481 | 23 |
end; |
24 |
||
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
25 |
|
15481 | 26 |
(* the signature of an instance of the SQSUBST tactic *) |
27 |
signature EQSUBST_TAC = |
|
28 |
sig |
|
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
29 |
|
15550
806214035275
lucas - added more comments and an extra type to clarify the code.
dixon
parents:
15538
diff
changeset
|
30 |
type match = |
15915
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents:
15855
diff
changeset
|
31 |
((Term.indexname * (Term.sort * Term.typ)) list (* type instantiations *) |
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents:
15855
diff
changeset
|
32 |
* (Term.indexname * (Term.typ * Term.term)) list) (* term instantiations *) |
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
33 |
* (string * Term.typ) list (* fake named type abs env *) |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
34 |
* (string * Term.typ) list (* type abs env *) |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
35 |
* Term.term (* outer term *) |
15550
806214035275
lucas - added more comments and an extra type to clarify the code.
dixon
parents:
15538
diff
changeset
|
36 |
|
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
37 |
val prep_subst_in_asm : |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
38 |
(Sign.sg (* sign for matching *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
39 |
-> int (* maxidx *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
40 |
-> 'a (* input object kind *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
41 |
-> BasicIsaFTerm.FcTerm (* focusterm to search under *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
42 |
-> 'b) (* result type *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
43 |
-> int (* subgoal to subst in *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
44 |
-> Thm.thm (* target theorem with subgoals *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
45 |
-> int (* premise to subst in *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
46 |
-> (Thm.cterm list (* certified free var placeholders for vars *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
47 |
* int (* premice no. to subst *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
48 |
* int (* number of assumptions of premice *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
49 |
* Thm.thm) (* premice as a new theorem for forward reasoning *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
50 |
* ('a -> 'b) (* matchf *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
51 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
52 |
val prep_subst_in_asms : |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
53 |
(Sign.sg -> int -> 'a -> BasicIsaFTerm.FcTerm -> 'b) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
54 |
-> int (* subgoal to subst in *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
55 |
-> Thm.thm (* target theorem with subgoals *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
56 |
-> ((Thm.cterm list (* certified free var placeholders for vars *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
57 |
* int (* premice no. to subst *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
58 |
* int (* number of assumptions of premice *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
59 |
* Thm.thm) (* premice as a new theorem for forward reasoning *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
60 |
* ('a -> 'b)) (* matchf *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
61 |
Seq.seq |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
62 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
63 |
val apply_subst_in_asm : |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
64 |
int (* subgoal *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
65 |
-> Thm.thm (* overall theorem *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
66 |
-> (Thm.cterm list (* certified free var placeholders for vars *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
67 |
* int (* assump no being subst *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
68 |
* int (* num of premises of asm *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
69 |
* Thm.thm) (* premthm *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
70 |
-> Thm.thm (* rule *) |
15550
806214035275
lucas - added more comments and an extra type to clarify the code.
dixon
parents:
15538
diff
changeset
|
71 |
-> match |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
72 |
-> Thm.thm Seq.seq |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
73 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
74 |
val prep_concl_subst : |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
75 |
(Sign.sg -> int -> 'a -> BasicIsaFTerm.FcTerm -> 'b) (* searchf *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
76 |
-> int (* subgoal *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
77 |
-> Thm.thm (* overall goal theorem *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
78 |
-> (Thm.cterm list * Thm.thm) * ('a -> 'b) (* (cvfs, conclthm), matchf *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
79 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
80 |
val apply_subst_in_concl : |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
81 |
int (* subgoal *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
82 |
-> Thm.thm (* thm with all goals *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
83 |
-> Thm.cterm list (* certified free var placeholders for vars *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
84 |
* Thm.thm (* trivial thm of goal concl *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
85 |
(* possible matches/unifiers *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
86 |
-> Thm.thm (* rule *) |
15550
806214035275
lucas - added more comments and an extra type to clarify the code.
dixon
parents:
15538
diff
changeset
|
87 |
-> match |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
88 |
-> Thm.thm Seq.seq (* substituted goal *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
89 |
|
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
90 |
val searchf_tlr_unify_all : |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
91 |
(Sign.sg -> int -> |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
92 |
Term.term -> |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
93 |
BasicIsaFTerm.FcTerm -> |
15929
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
94 |
match Seq.seq Seq.seq) |
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
95 |
val searchf_tlr_unify_valid : |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
96 |
(Sign.sg -> int -> |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
97 |
Term.term -> |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
98 |
BasicIsaFTerm.FcTerm -> |
15929
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
99 |
match Seq.seq Seq.seq) |
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
100 |
|
15936
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
101 |
val eqsubst_asm_meth : int list -> Thm.thm list -> Proof.method |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
102 |
val eqsubst_asm_tac : int list -> Thm.thm list -> int -> Thm.thm -> Thm.thm Seq.seq |
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
103 |
val eqsubst_asm_tac' : |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
104 |
(Sign.sg -> int -> |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
105 |
Term.term -> |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
106 |
BasicIsaFTerm.FcTerm -> |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
107 |
match Seq.seq) -> Thm.thm -> int -> Thm.thm -> Thm.thm Seq.seq |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
108 |
|
15936
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
109 |
val eqsubst_meth : int list -> Thm.thm list -> Proof.method |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
110 |
val eqsubst_tac : int list -> Thm.thm list -> int -> Thm.thm -> Thm.thm Seq.seq |
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
111 |
val eqsubst_tac' : |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
112 |
(Sign.sg -> int -> |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
113 |
Term.term -> |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
114 |
BasicIsaFTerm.FcTerm -> |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
115 |
match Seq.seq) -> Thm.thm -> int -> Thm.thm -> Thm.thm Seq.seq |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
116 |
|
15936
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
117 |
val meth : (bool * int list) * Thm.thm list -> Proof.context -> Proof.method |
15481 | 118 |
val setup : (Theory.theory -> Theory.theory) list |
119 |
end; |
|
120 |
||
121 |
functor EQSubstTacFUN (structure EqRuleData : EQRULE_DATA) |
|
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
122 |
: EQSUBST_TAC |
15481 | 123 |
= struct |
124 |
||
15915
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents:
15855
diff
changeset
|
125 |
(* a type abriviation for match information *) |
15550
806214035275
lucas - added more comments and an extra type to clarify the code.
dixon
parents:
15538
diff
changeset
|
126 |
type match = |
15915
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents:
15855
diff
changeset
|
127 |
((Term.indexname * (Term.sort * Term.typ)) list (* type instantiations *) |
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents:
15855
diff
changeset
|
128 |
* (Term.indexname * (Term.typ * Term.term)) list) (* term instantiations *) |
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents:
15855
diff
changeset
|
129 |
* (string * Term.typ) list (* fake named type abs env *) |
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents:
15855
diff
changeset
|
130 |
* (string * Term.typ) list (* type abs env *) |
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents:
15855
diff
changeset
|
131 |
* Term.term (* outer term *) |
15550
806214035275
lucas - added more comments and an extra type to clarify the code.
dixon
parents:
15538
diff
changeset
|
132 |
|
806214035275
lucas - added more comments and an extra type to clarify the code.
dixon
parents:
15538
diff
changeset
|
133 |
|
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
134 |
(* FOR DEBUGGING... |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
135 |
type trace_subst_errT = int (* subgoal *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
136 |
* Thm.thm (* thm with all goals *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
137 |
* (Thm.cterm list (* certified free var placeholders for vars *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
138 |
* Thm.thm) (* trivial thm of goal concl *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
139 |
(* possible matches/unifiers *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
140 |
* Thm.thm (* rule *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
141 |
* (((Term.indexname * Term.typ) list (* type instantiations *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
142 |
* (Term.indexname * Term.term) list ) (* term instantiations *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
143 |
* (string * Term.typ) list (* Type abs env *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
144 |
* Term.term) (* outer term *); |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
145 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
146 |
val trace_subst_err = (ref NONE : trace_subst_errT option ref); |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
147 |
val trace_subst_search = ref false; |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
148 |
exception trace_subst_exp of trace_subst_errT; |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
149 |
*) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
150 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
151 |
(* also defined in /HOL/Tools/inductive_codegen.ML, |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
152 |
maybe move this to seq.ML ? *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
153 |
infix 5 :->; |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
154 |
fun s :-> f = Seq.flat (Seq.map f s); |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
155 |
|
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
156 |
(* search from top, left to right, then down *) |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
157 |
fun search_tlr_all_f f ft = |
15481 | 158 |
let |
159 |
fun maux ft = |
|
160 |
let val t' = (IsaFTerm.focus_of_fcterm ft) |
|
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
161 |
(* val _ = |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
162 |
if !trace_subst_search then |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
163 |
(writeln ("Examining: " ^ (TermLib.string_of_term t')); |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
164 |
TermLib.writeterm t'; ()) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
165 |
else (); *) |
15481 | 166 |
in |
167 |
(case t' of |
|
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
168 |
(_ $ _) => Seq.append(maux (IsaFTerm.focus_left ft), |
15929
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
169 |
Seq.cons(f ft, |
15481 | 170 |
maux (IsaFTerm.focus_right ft))) |
15929
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
171 |
| (Abs _) => Seq.cons(f ft, maux (IsaFTerm.focus_abs ft)) |
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
172 |
| leaf => Seq.single (f ft)) end |
15481 | 173 |
in maux ft end; |
174 |
||
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
175 |
(* search from top, left to right, then down *) |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
176 |
fun search_tlr_valid_f f ft = |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
177 |
let |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
178 |
fun maux ft = |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
179 |
let |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
180 |
val hereseq = if IsaFTerm.valid_match_start ft then f ft else Seq.empty |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
181 |
in |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
182 |
(case (IsaFTerm.focus_of_fcterm ft) of |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
183 |
(_ $ _) => Seq.append(maux (IsaFTerm.focus_left ft), |
15929
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
184 |
Seq.cons(hereseq, |
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
185 |
maux (IsaFTerm.focus_right ft))) |
15929
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
186 |
| (Abs _) => Seq.cons(hereseq, maux (IsaFTerm.focus_abs ft)) |
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
187 |
| leaf => Seq.single (hereseq)) |
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
188 |
end |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
189 |
in maux ft end; |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
190 |
|
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
191 |
(* search all unifications *) |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
192 |
fun searchf_tlr_unify_all sgn maxidx lhs = |
15481 | 193 |
IsaFTerm.find_fcterm_matches |
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
194 |
search_tlr_all_f |
15481 | 195 |
(IsaFTerm.clean_unify_ft sgn maxidx lhs); |
196 |
||
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
197 |
(* search only for 'valid' unifiers (non abs subterms and non vars) *) |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
198 |
fun searchf_tlr_unify_valid sgn maxidx lhs = |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
199 |
IsaFTerm.find_fcterm_matches |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
200 |
search_tlr_valid_f |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
201 |
(IsaFTerm.clean_unify_ft sgn maxidx lhs); |
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
202 |
|
15929
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
203 |
(* special tactic to skip the first "occ" occurances - ie start at the nth match *) |
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
204 |
fun skip_first_occs_search occ searchf sgn i t ft = |
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
205 |
let |
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
206 |
fun skip_occs n sq = |
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
207 |
if n <= 1 then sq |
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
208 |
else |
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
209 |
(case (Seq.pull sq) of NONE => Seq.empty |
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
210 |
| SOME (h,t) => |
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
211 |
(case Seq.pull h of NONE => skip_occs n t |
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
212 |
| SOME _ => skip_occs (n - 1) t)) |
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
213 |
in Seq.flat (skip_occs occ (searchf sgn i t ft)) end; |
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
214 |
|
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
215 |
|
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
216 |
(* apply a substitution in the conclusion of the theorem th *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
217 |
(* cfvs are certified free var placeholders for goal params *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
218 |
(* conclthm is a theorem of for just the conclusion *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
219 |
(* m is instantiation/match information *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
220 |
(* rule is the equation for substitution *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
221 |
fun apply_subst_in_concl i th (cfvs, conclthm) rule m = |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
222 |
(RWInst.rw m rule conclthm) |
15855 | 223 |
|> IsaND.unfix_frees cfvs |
15915
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents:
15855
diff
changeset
|
224 |
|> RWInst.beta_eta_contract |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
225 |
|> (fn r => Tactic.rtac r i th); |
15481 | 226 |
|
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
227 |
(* |
15481 | 228 |
|
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
229 |
|> (fn r => Thm.bicompose false (false, r, Thm.nprems_of r) i th) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
230 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
231 |
*) |
15481 | 232 |
|
233 |
(* substitute within the conclusion of goal i of gth, using a meta |
|
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
234 |
equation rule. Note that we assume rule has var indicies zero'd *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
235 |
fun prep_concl_subst searchf i gth = |
15481 | 236 |
let |
237 |
val th = Thm.incr_indexes 1 gth; |
|
238 |
val tgt_term = Thm.prop_of th; |
|
239 |
||
240 |
val sgn = Thm.sign_of_thm th; |
|
241 |
val ctermify = Thm.cterm_of sgn; |
|
242 |
val trivify = Thm.trivial o ctermify; |
|
243 |
||
244 |
val (fixedbody, fvs) = IsaND.fix_alls_term i tgt_term; |
|
245 |
val cfvs = rev (map ctermify fvs); |
|
246 |
||
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
247 |
val conclterm = Logic.strip_imp_concl fixedbody; |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
248 |
val conclthm = trivify conclterm; |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
249 |
val maxidx = Term.maxidx_of_term conclterm; |
15481 | 250 |
in |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
251 |
((cfvs, conclthm), |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
252 |
(fn lhs => searchf sgn maxidx lhs |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
253 |
((IsaFTerm.focus_right |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
254 |
o IsaFTerm.focus_left |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
255 |
o IsaFTerm.fcterm_of_term |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
256 |
o Thm.prop_of) conclthm))) |
15481 | 257 |
end; |
258 |
||
259 |
(* substitute using an object or meta level equality *) |
|
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
260 |
fun eqsubst_tac' searchf instepthm i th = |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
261 |
let |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
262 |
val (cvfsconclthm, findmatchf) = |
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
263 |
prep_concl_subst searchf i th; |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
264 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
265 |
val stepthms = |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
266 |
Seq.map Drule.zero_var_indexes |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
267 |
(Seq.of_list (EqRuleData.prep_meta_eq instepthm)); |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
268 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
269 |
fun rewrite_with_thm r = |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
270 |
let val (lhs,_) = Logic.dest_equals (Thm.concl_of r); |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
271 |
in (findmatchf lhs) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
272 |
:-> (apply_subst_in_concl i th cvfsconclthm r) end; |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
273 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
274 |
in (stepthms :-> rewrite_with_thm) end; |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
275 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
276 |
|
15936
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
277 |
(* General substiuttion of multiple occurances using one of |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
278 |
the given theorems*) |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
279 |
fun eqsubst_occL tac occL thms i th = |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
280 |
let val nprems = Thm.nprems_of th in |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
281 |
if nprems < i then Seq.empty else |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
282 |
let val thmseq = (Seq.of_list thms) |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
283 |
fun apply_occ occ th = |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
284 |
thmseq :-> |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
285 |
(fn r => tac (skip_first_occs_search |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
286 |
occ searchf_tlr_unify_valid) r |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
287 |
(i + ((Thm.nprems_of th) - nprems)) |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
288 |
th); |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
289 |
in |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
290 |
Seq.EVERY (map apply_occ (Library.sort (Library.rev_order o Library.int_ord) occL)) |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
291 |
th |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
292 |
end |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
293 |
end; |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
294 |
|
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
295 |
(* implicit argus: occL thms i th *) |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
296 |
val eqsubst_tac = eqsubst_occL eqsubst_tac'; |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
297 |
|
15481 | 298 |
|
299 |
(* inthms are the given arguments in Isar, and treated as eqstep with |
|
300 |
the first one, then the second etc *) |
|
15936
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
301 |
fun eqsubst_meth occL inthms = |
15481 | 302 |
Method.METHOD |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
303 |
(fn facts => |
15936
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
304 |
HEADGOAL ( Method.insert_tac facts THEN' eqsubst_tac occL inthms )); |
15481 | 305 |
|
306 |
||
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
307 |
fun apply_subst_in_asm i th (cfvs, j, nprems, pth) rule m = |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
308 |
(RWInst.rw m rule pth) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
309 |
|> Thm.permute_prems 0 ~1 |
15855 | 310 |
|> IsaND.unfix_frees cfvs |
15915
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents:
15855
diff
changeset
|
311 |
|> RWInst.beta_eta_contract |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
312 |
|> (fn r => Tactic.dtac r i th); |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
313 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
314 |
(* |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
315 |
? should I be using bicompose what if we match more than one |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
316 |
assumption, even after instantiation ? (back will work, but it would |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
317 |
be nice to avoid the redudent search) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
318 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
319 |
something like... |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
320 |
|> Thm.lift_rule (th, i) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
321 |
|> (fn r => Thm.bicompose false (false, r, Thm.nprems_of r - nprems) i th) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
322 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
323 |
*) |
15481 | 324 |
|
325 |
||
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
326 |
(* prepare to substitute within the j'th premise of subgoal i of gth, |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
327 |
using a meta-level equation. Note that we assume rule has var indicies |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
328 |
zero'd. Note that we also assume that premt is the j'th premice of |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
329 |
subgoal i of gth. Note the repetition of work done for each |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
330 |
assumption, i.e. this can be made more efficient for search over |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
331 |
multiple assumptions. *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
332 |
fun prep_subst_in_asm searchf i gth j = |
15481 | 333 |
let |
334 |
val th = Thm.incr_indexes 1 gth; |
|
335 |
val tgt_term = Thm.prop_of th; |
|
336 |
||
337 |
val sgn = Thm.sign_of_thm th; |
|
338 |
val ctermify = Thm.cterm_of sgn; |
|
339 |
val trivify = Thm.trivial o ctermify; |
|
340 |
||
341 |
val (fixedbody, fvs) = IsaND.fix_alls_term i tgt_term; |
|
342 |
val cfvs = rev (map ctermify fvs); |
|
343 |
||
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
344 |
val asmt = Library.nth_elem(j - 1,(Logic.strip_imp_prems fixedbody)); |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
345 |
val asm_nprems = length (Logic.strip_imp_prems asmt); |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
346 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
347 |
val pth = trivify asmt; |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
348 |
val maxidx = Term.maxidx_of_term asmt; |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
349 |
|
15481 | 350 |
in |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
351 |
((cfvs, j, asm_nprems, pth), |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
352 |
(fn lhs => (searchf sgn maxidx lhs |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
353 |
((IsaFTerm.focus_right |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
354 |
o IsaFTerm.fcterm_of_term |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
355 |
o Thm.prop_of) pth)))) |
15481 | 356 |
end; |
357 |
||
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
358 |
(* prepare subst in every possible assumption *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
359 |
fun prep_subst_in_asms searchf i gth = |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
360 |
Seq.map |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
361 |
(prep_subst_in_asm searchf i gth) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
362 |
(Seq.of_list (IsaPLib.mk_num_list |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
363 |
(length (Logic.prems_of_goal (Thm.prop_of gth) i)))); |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
364 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
365 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
366 |
(* substitute in an assumption using an object or meta level equality *) |
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
367 |
fun eqsubst_asm_tac' searchf instepthm i th = |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
368 |
let |
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
369 |
val asmpreps = prep_subst_in_asms searchf i th; |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
370 |
val stepthms = |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
371 |
Seq.map Drule.zero_var_indexes |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
372 |
(Seq.of_list (EqRuleData.prep_meta_eq instepthm)) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
373 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
374 |
fun rewrite_with_thm (asminfo, findmatchf) r = |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
375 |
let val (lhs,_) = Logic.dest_equals (Thm.concl_of r); |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
376 |
in (findmatchf lhs) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
377 |
:-> (apply_subst_in_asm i th asminfo r) end; |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
378 |
in |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
379 |
(asmpreps :-> (fn a => stepthms :-> rewrite_with_thm a)) |
15481 | 380 |
end; |
381 |
||
15936
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
382 |
(* substitute using one of the given theorems in an assumption. |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
383 |
Implicit args: occL thms i th *) |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
384 |
val eqsubst_asm_tac = eqsubst_occL eqsubst_asm_tac'; |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
385 |
|
15481 | 386 |
|
387 |
(* inthms are the given arguments in Isar, and treated as eqstep with |
|
388 |
the first one, then the second etc *) |
|
15936
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
389 |
fun eqsubst_asm_meth occL inthms = |
15481 | 390 |
Method.METHOD |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
391 |
(fn facts => |
15936
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
392 |
HEADGOAL (Method.insert_tac facts THEN' eqsubst_asm_tac occL inthms )); |
15481 | 393 |
|
394 |
(* combination method that takes a flag (true indicates that subst |
|
395 |
should be done to an assumption, false = apply to the conclusion of |
|
396 |
the goal) as well as the theorems to use *) |
|
15936
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
397 |
fun meth ((asmflag, occL), inthms) ctxt = |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
398 |
if asmflag then eqsubst_asm_meth occL inthms else eqsubst_meth occL inthms; |
15481 | 399 |
|
400 |
(* syntax for options, given "(asm)" will give back true, without |
|
401 |
gives back false *) |
|
402 |
val options_syntax = |
|
403 |
(Args.parens (Args.$$$ "asm") >> (K true)) || |
|
404 |
(Scan.succeed false); |
|
15936
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
405 |
|
15929
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
406 |
val ith_syntax = |
15936
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
407 |
(Args.parens (Scan.repeat Args.nat)) |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
408 |
|| (Scan.succeed [0]); |
15481 | 409 |
|
410 |
(* method syntax, first take options, then theorems *) |
|
411 |
fun meth_syntax meth src ctxt = |
|
412 |
meth (snd (Method.syntax ((Scan.lift options_syntax) |
|
15929
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
413 |
-- (Scan.lift ith_syntax) |
15481 | 414 |
-- Attrib.local_thms) src ctxt)) |
415 |
ctxt; |
|
416 |
||
417 |
(* setup function for adding method to theory. *) |
|
418 |
val setup = |
|
419 |
[Method.add_method ("subst", meth_syntax meth, "Substiution with an equation. Use \"(asm)\" option to substitute in an assumption.")]; |
|
420 |
||
421 |
end; |