author | paulson |
Thu, 06 Sep 2007 17:03:32 +0200 | |
changeset 24546 | c90cee3163b7 |
parent 24271 | 499608101177 |
child 24558 | 419f7cde7f59 |
permissions | -rw-r--r-- |
22729 | 1 |
(* Title: HOL/Nominal/nominal_fresh_fun.ML |
2 |
ID: $Id$ |
|
3 |
Authors: Stefan Berghofer, Julien Narboux, TU Muenchen |
|
4 |
||
5 |
A tactic to generate fresh names. |
|
6 |
A tactic to get rid of the fresh_fun. |
|
7 |
*) |
|
8 |
||
23091
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
9 |
(* First some functions that could be in the library *) |
22729 | 10 |
|
22774
8c64803fae48
adds op in front of an infix to fix SML compilation
narboux
parents:
22753
diff
changeset
|
11 |
(* A tactical which applies a list of int -> tactic to the corresponding subgoals present after the application of another tactic. |
8c64803fae48
adds op in front of an infix to fix SML compilation
narboux
parents:
22753
diff
changeset
|
12 |
T THENL [A,B,C] is equivalent to T THEN (C 3 THEN B 2 THEN A 1) |
8c64803fae48
adds op in front of an infix to fix SML compilation
narboux
parents:
22753
diff
changeset
|
13 |
*) |
22753 | 14 |
|
15 |
infix 1 THENL |
|
16 |
||
22792 | 17 |
fun tac THENL tacs = |
22753 | 18 |
tac THEN |
19 |
(EVERY (map (fn (tac,i) => tac i) (rev tacs ~~ (length tacs downto 1)))) |
|
20 |
||
23091
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
21 |
(* A tactical to test if a tactic completly solve a subgoal *) |
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
22 |
|
23054
d1bbe5afa279
change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents:
22792
diff
changeset
|
23 |
fun SOLVEI t = t THEN_ALL_NEW (fn i => no_tac); |
23091
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
24 |
|
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
25 |
(* A version of TRY for int -> tactic *) |
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
26 |
|
23054
d1bbe5afa279
change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents:
22792
diff
changeset
|
27 |
fun TRY' tac i = TRY (tac i); |
d1bbe5afa279
change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents:
22792
diff
changeset
|
28 |
|
22729 | 29 |
fun gen_res_inst_tac_term instf tyinst tinst elim th i st = |
30 |
let |
|
31 |
val thy = theory_of_thm st; |
|
32 |
val cgoal = nth (cprems_of st) (i - 1); |
|
33 |
val {maxidx, ...} = rep_cterm cgoal; |
|
34 |
val j = maxidx + 1; |
|
35 |
val tyinst' = map (apfst (Logic.incr_tvar j)) tyinst; |
|
36 |
val ps = Logic.strip_params (term_of cgoal); |
|
37 |
val Ts = map snd ps; |
|
38 |
val tinst' = map (fn (t, u) => |
|
39 |
(head_of (Logic.incr_indexes (Ts, j) t), |
|
40 |
list_abs (ps, u))) tinst; |
|
41 |
val th' = instf |
|
42 |
(map (pairself (ctyp_of thy)) tyinst') |
|
43 |
(map (pairself (cterm_of thy)) tinst') |
|
44 |
(Thm.lift_rule cgoal th) |
|
45 |
in |
|
46 |
compose_tac (elim, th', nprems_of th) i st |
|
47 |
end handle Subscript => Seq.empty; |
|
48 |
||
49 |
val res_inst_tac_term = |
|
50 |
gen_res_inst_tac_term (curry Thm.instantiate); |
|
51 |
||
52 |
val res_inst_tac_term' = |
|
53 |
gen_res_inst_tac_term (K Drule.cterm_instantiate) []; |
|
54 |
||
55 |
fun cut_inst_tac_term' tinst th = |
|
56 |
res_inst_tac_term' tinst false (Tactic.make_elim_preserve th); |
|
57 |
||
58 |
fun get_dyn_thm thy name atom_name = (PureThy.get_thm thy (Name name)) handle _ => raise ERROR ("The atom type "^atom_name^" is not defined."); |
|
59 |
||
60 |
(* End of function waiting to be |
|
61 |
in the library *) |
|
62 |
||
63 |
(* The theorems needed that are known at |
|
64 |
compile time. *) |
|
65 |
||
66 |
val at_exists_fresh' = thm "at_exists_fresh'"; |
|
22787 | 67 |
val fresh_fun_app' = thm "fresh_fun_app'"; |
22729 | 68 |
val fresh_prod = thm "fresh_prod"; |
69 |
||
70 |
(* A tactic to generate a name fresh for |
|
71 |
all the free variables and parameters of the goal *) |
|
72 |
||
73 |
fun generate_fresh_tac atom_name i thm = |
|
74 |
let |
|
75 |
val thy = theory_of_thm thm; |
|
76 |
(* the parsing function returns a qualified name, we get back the base name *) |
|
77 |
val atom_basename = Sign.base_name atom_name; |
|
78 |
val goal = List.nth(prems_of thm, i-1); |
|
79 |
val ps = Logic.strip_params goal; |
|
80 |
val Ts = rev (map snd ps); |
|
24271 | 81 |
fun is_of_fs_name T = Sign.of_sort thy (T, Sign.intern_sort thy ["fs_"^atom_basename]); |
22729 | 82 |
(* rebuild de bruijn indices *) |
83 |
val bvs = map_index (Bound o fst) ps; |
|
84 |
(* select variables of the right class *) |
|
85 |
val vs = filter (fn t => is_of_fs_name (fastype_of1 (Ts, t))) |
|
86 |
(term_frees goal @ bvs); |
|
87 |
(* build the tuple *) |
|
23368
ad690b9bca1c
generate_fresh works even if there is no free variable in the goal
narboux
parents:
23094
diff
changeset
|
88 |
val s = (Library.foldr1 (fn (v, s) => |
ad690b9bca1c
generate_fresh works even if there is no free variable in the goal
narboux
parents:
23094
diff
changeset
|
89 |
HOLogic.pair_const (fastype_of1 (Ts, v)) (fastype_of1 (Ts, s)) $ v $ s) vs) handle _ => HOLogic.unit ; |
22729 | 90 |
val fs_name_thm = get_dyn_thm thy ("fs_"^atom_basename^"1") atom_basename; |
91 |
val at_name_inst_thm = get_dyn_thm thy ("at_"^atom_basename^"_inst") atom_basename; |
|
92 |
val exists_fresh' = at_name_inst_thm RS at_exists_fresh'; |
|
93 |
(* find the variable we want to instantiate *) |
|
94 |
val x = hd (term_vars (prop_of exists_fresh')); |
|
95 |
in |
|
96 |
(cut_inst_tac_term' [(x,s)] exists_fresh' 1 THEN |
|
97 |
rtac fs_name_thm 1 THEN |
|
98 |
etac exE 1) thm |
|
99 |
handle Empty => all_tac thm (* if we collected no variables then we do nothing *) |
|
100 |
end; |
|
101 |
||
102 |
fun get_inner_fresh_fun (Bound j) = NONE |
|
103 |
| get_inner_fresh_fun (v as Free _) = NONE |
|
104 |
| get_inner_fresh_fun (v as Var _) = NONE |
|
105 |
| get_inner_fresh_fun (Const _) = NONE |
|
106 |
| get_inner_fresh_fun (Abs (_, _, t)) = get_inner_fresh_fun t |
|
107 |
| get_inner_fresh_fun (Const ("Nominal.fresh_fun",Type("fun",[Type ("fun",[Type (T,_),_]),_])) $ u) = SOME T |
|
108 |
| get_inner_fresh_fun (t $ u) = |
|
109 |
let val a = get_inner_fresh_fun u in |
|
110 |
if a = NONE then get_inner_fresh_fun t else a |
|
111 |
end; |
|
112 |
||
113 |
(* This tactic generates a fresh name |
|
114 |
of the atom type given by the inner most fresh_fun *) |
|
115 |
||
116 |
fun generate_fresh_fun_tac i thm = |
|
117 |
let |
|
118 |
val goal = List.nth(prems_of thm, i-1); |
|
119 |
val atom_name_opt = get_inner_fresh_fun goal; |
|
120 |
in |
|
121 |
case atom_name_opt of |
|
122 |
NONE => all_tac thm |
|
123 |
| SOME atom_name => generate_fresh_tac atom_name i thm |
|
124 |
end |
|
125 |
||
23091
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
126 |
(* Two substitution tactics which looks for the inner most occurence in |
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
127 |
one assumption or in the conclusion *) |
22729 | 128 |
|
23065 | 129 |
val search_fun = curry (Seq.flat o (uncurry EqSubst.searchf_bt_unify_valid)); |
130 |
val search_fun_asm = EqSubst.skip_first_asm_occs_search EqSubst.searchf_bt_unify_valid; |
|
23054
d1bbe5afa279
change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents:
22792
diff
changeset
|
131 |
|
23091
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
132 |
fun subst_inner_tac ctx = EqSubst.eqsubst_tac' ctx search_fun; |
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
133 |
fun subst_inner_asm_tac_aux i ctx = EqSubst.eqsubst_asm_tac' ctx search_fun_asm i; |
23054
d1bbe5afa279
change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents:
22792
diff
changeset
|
134 |
|
23091
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
135 |
(* A tactic to substitute in the first assumption |
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
136 |
which contains an occurence. *) |
22729 | 137 |
|
23091
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
138 |
fun subst_inner_asm_tac ctx th = curry (curry (FIRST' (map uncurry (map uncurry (map subst_inner_asm_tac_aux (1 upto Thm.nprems_of th)))))) ctx th; |
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
139 |
|
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
140 |
fun fresh_fun_tac no_asm i thm = |
22729 | 141 |
(* Find the variable we instantiate *) |
142 |
let |
|
143 |
val thy = theory_of_thm thm; |
|
144 |
val ctx = Context.init_proof thy; |
|
145 |
val ss = simpset_of thy; |
|
146 |
val abs_fresh = PureThy.get_thms thy (Name "abs_fresh"); |
|
23054
d1bbe5afa279
change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents:
22792
diff
changeset
|
147 |
val fresh_perm_app = PureThy.get_thms thy (Name "fresh_perm_app"); |
22729 | 148 |
val ss' = ss addsimps fresh_prod::abs_fresh; |
23054
d1bbe5afa279
change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents:
22792
diff
changeset
|
149 |
val ss'' = ss' addsimps fresh_perm_app; |
22729 | 150 |
val x = hd (tl (term_vars (prop_of exI))); |
22787 | 151 |
val goal = nth (prems_of thm) (i-1); |
22729 | 152 |
val atom_name_opt = get_inner_fresh_fun goal; |
153 |
val n = List.length (Logic.strip_params goal); |
|
154 |
(* Here we rely on the fact that the variable introduced by generate_fresh_tac is the last one in the list, the inner one *) |
|
155 |
in |
|
156 |
case atom_name_opt of |
|
157 |
NONE => all_tac thm |
|
158 |
| SOME atom_name => |
|
159 |
let |
|
160 |
val atom_basename = Sign.base_name atom_name; |
|
161 |
val pt_name_inst = get_dyn_thm thy ("pt_"^atom_basename^"_inst") atom_basename; |
|
162 |
val at_name_inst = get_dyn_thm thy ("at_"^atom_basename^"_inst") atom_basename; |
|
22787 | 163 |
fun inst_fresh vars params i st = |
164 |
let val vars' = term_vars (prop_of st); |
|
165 |
val thy = theory_of_thm st; |
|
166 |
in case vars' \\ vars of |
|
167 |
[x] => Seq.single (Thm.instantiate ([],[(cterm_of thy x,cterm_of thy (list_abs (params,Bound 0)))]) st) |
|
168 |
| _ => error "fresh_fun_simp: Too many variables, please report." |
|
169 |
end |
|
22729 | 170 |
in |
23091
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
171 |
((fn st => |
22787 | 172 |
let |
173 |
val vars = term_vars (prop_of st); |
|
174 |
val params = Logic.strip_params (nth (prems_of st) (i-1)) |
|
23091
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
175 |
(* The tactics which solve the subgoals generated |
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
176 |
by the conditionnal rewrite rule. *) |
23054
d1bbe5afa279
change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents:
22792
diff
changeset
|
177 |
val post_rewrite_tacs = |
d1bbe5afa279
change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents:
22792
diff
changeset
|
178 |
[rtac pt_name_inst, |
d1bbe5afa279
change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents:
22792
diff
changeset
|
179 |
rtac at_name_inst, |
d1bbe5afa279
change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents:
22792
diff
changeset
|
180 |
TRY' (SOLVEI (NominalPermeq.finite_guess_tac ss'')), |
d1bbe5afa279
change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents:
22792
diff
changeset
|
181 |
inst_fresh vars params THEN' |
d1bbe5afa279
change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents:
22792
diff
changeset
|
182 |
(TRY' (SOLVEI (NominalPermeq.fresh_guess_tac ss''))) THEN' |
d1bbe5afa279
change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents:
22792
diff
changeset
|
183 |
(TRY' (SOLVEI (asm_full_simp_tac ss'')))] |
d1bbe5afa279
change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents:
22792
diff
changeset
|
184 |
in |
23094
f1df8d2da98a
fix a bug : the semantics of no_asm was the opposite
narboux
parents:
23091
diff
changeset
|
185 |
((if no_asm then no_tac else |
f1df8d2da98a
fix a bug : the semantics of no_asm was the opposite
narboux
parents:
23091
diff
changeset
|
186 |
(subst_inner_asm_tac ctx fresh_fun_app' i THENL post_rewrite_tacs)) |
23091
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
187 |
ORELSE |
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
188 |
(subst_inner_tac ctx fresh_fun_app' i THENL post_rewrite_tacs)) st |
22787 | 189 |
end)) thm |
190 |
||
22729 | 191 |
end |
192 |
end |
|
193 |
||
23091
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
194 |
(* syntax for options, given "(no_asm)" will give back true, without |
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
195 |
gives back false *) |
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
196 |
val options_syntax = |
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
197 |
(Args.parens (Args.$$$ "no_asm") >> (K true)) || |
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
198 |
(Scan.succeed false); |
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
199 |
|
22729 | 200 |
val setup_generate_fresh = |
201 |
Method.goal_args_ctxt' Args.tyname (fn ctxt => generate_fresh_tac) |
|
202 |
||
203 |
val setup_fresh_fun_simp = |
|
23091
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
204 |
Method.simple_args options_syntax |
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
205 |
(fn b => fn _ => Method.SIMPLE_METHOD (fresh_fun_tac b 1)) |