| author | hoelzl |
| Mon, 11 Jan 2010 11:47:38 +0100 | |
| changeset 34872 | 6ca970cfa873 |
| parent 33040 | cffdb7b28498 |
| child 34885 | 6587c24ef6d8 |
| permissions | -rw-r--r-- |
| 22729 | 1 |
(* Title: HOL/Nominal/nominal_fresh_fun.ML |
|
29265
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents:
28397
diff
changeset
|
2 |
Authors: Stefan Berghofer and Julien Narboux, TU Muenchen |
| 22729 | 3 |
|
| 24558 | 4 |
Provides a tactic to generate fresh names and |
5 |
a tactic to analyse instances of the fresh_fun. |
|
| 22729 | 6 |
*) |
7 |
||
| 24558 | 8 |
(* First some functions that should be in the library *) |
| 22729 | 9 |
|
| 24558 | 10 |
(* A tactic which only succeeds when the argument *) |
11 |
(* tactic solves completely the specified subgoal *) |
|
|
23054
d1bbe5afa279
change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents:
22792
diff
changeset
|
12 |
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
|
13 |
|
| 22729 | 14 |
fun gen_res_inst_tac_term instf tyinst tinst elim th i st = |
15 |
let |
|
16 |
val thy = theory_of_thm st; |
|
17 |
val cgoal = nth (cprems_of st) (i - 1); |
|
18 |
val {maxidx, ...} = rep_cterm cgoal;
|
|
19 |
val j = maxidx + 1; |
|
20 |
val tyinst' = map (apfst (Logic.incr_tvar j)) tyinst; |
|
21 |
val ps = Logic.strip_params (term_of cgoal); |
|
22 |
val Ts = map snd ps; |
|
23 |
val tinst' = map (fn (t, u) => |
|
24 |
(head_of (Logic.incr_indexes (Ts, j) t), |
|
25 |
list_abs (ps, u))) tinst; |
|
26 |
val th' = instf |
|
27 |
(map (pairself (ctyp_of thy)) tyinst') |
|
28 |
(map (pairself (cterm_of thy)) tinst') |
|
29 |
(Thm.lift_rule cgoal th) |
|
30 |
in |
|
31 |
compose_tac (elim, th', nprems_of th) i st |
|
32 |
end handle Subscript => Seq.empty; |
|
33 |
||
| 33034 | 34 |
val res_inst_tac_term = |
| 22729 | 35 |
gen_res_inst_tac_term (curry Thm.instantiate); |
36 |
||
| 33034 | 37 |
val res_inst_tac_term' = |
| 22729 | 38 |
gen_res_inst_tac_term (K Drule.cterm_instantiate) []; |
39 |
||
40 |
fun cut_inst_tac_term' tinst th = |
|
| 27239 | 41 |
res_inst_tac_term' tinst false (RuleInsts.make_elim_preserve th); |
| 22729 | 42 |
|
| 26337 | 43 |
fun get_dyn_thm thy name atom_name = |
|
26343
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26337
diff
changeset
|
44 |
PureThy.get_thm thy name handle ERROR _ => |
| 27187 | 45 |
error ("The atom type "^atom_name^" is not defined.");
|
| 22729 | 46 |
|
| 24558 | 47 |
(* End of function waiting to be in the library :o) *) |
| 22729 | 48 |
|
| 24558 | 49 |
(* The theorems needed that are known at compile time. *) |
50 |
val at_exists_fresh' = @{thm "at_exists_fresh'"};
|
|
51 |
val fresh_fun_app' = @{thm "fresh_fun_app'"};
|
|
52 |
val fresh_prod = @{thm "fresh_prod"};
|
|
| 22729 | 53 |
|
| 33034 | 54 |
(* A tactic to generate a name fresh for all the free *) |
| 24558 | 55 |
(* variables and parameters of the goal *) |
| 22729 | 56 |
|
57 |
fun generate_fresh_tac atom_name i thm = |
|
| 33034 | 58 |
let |
| 22729 | 59 |
val thy = theory_of_thm thm; |
60 |
(* the parsing function returns a qualified name, we get back the base name *) |
|
|
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30280
diff
changeset
|
61 |
val atom_basename = Long_Name.base_name atom_name; |
| 22729 | 62 |
val goal = List.nth(prems_of thm, i-1); |
63 |
val ps = Logic.strip_params goal; |
|
64 |
val Ts = rev (map snd ps); |
|
| 33034 | 65 |
fun is_of_fs_name T = Sign.of_sort thy (T, Sign.intern_sort thy ["fs_"^atom_basename]); |
| 22729 | 66 |
(* rebuild de bruijn indices *) |
67 |
val bvs = map_index (Bound o fst) ps; |
|
68 |
(* select variables of the right class *) |
|
69 |
val vs = filter (fn t => is_of_fs_name (fastype_of1 (Ts, t))) |
|
|
29265
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents:
28397
diff
changeset
|
70 |
(OldTerm.term_frees goal @ bvs); |
| 22729 | 71 |
(* build the tuple *) |
|
23368
ad690b9bca1c
generate_fresh works even if there is no free variable in the goal
narboux
parents:
23094
diff
changeset
|
72 |
val s = (Library.foldr1 (fn (v, s) => |
|
28397
389c5e494605
handle _ should be avoided (spurious Interrupt will spoil the game);
wenzelm
parents:
27239
diff
changeset
|
73 |
HOLogic.pair_const (fastype_of1 (Ts, v)) (fastype_of1 (Ts, s)) $ v $ s) vs) handle _ => HOLogic.unit ; (* FIXME avoid handle _ *) |
| 22729 | 74 |
val fs_name_thm = get_dyn_thm thy ("fs_"^atom_basename^"1") atom_basename;
|
75 |
val at_name_inst_thm = get_dyn_thm thy ("at_"^atom_basename^"_inst") atom_basename;
|
|
76 |
val exists_fresh' = at_name_inst_thm RS at_exists_fresh'; |
|
77 |
(* find the variable we want to instantiate *) |
|
|
29265
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents:
28397
diff
changeset
|
78 |
val x = hd (OldTerm.term_vars (prop_of exists_fresh')); |
| 33034 | 79 |
in |
| 22729 | 80 |
(cut_inst_tac_term' [(x,s)] exists_fresh' 1 THEN |
81 |
rtac fs_name_thm 1 THEN |
|
82 |
etac exE 1) thm |
|
83 |
handle Empty => all_tac thm (* if we collected no variables then we do nothing *) |
|
84 |
end; |
|
85 |
||
86 |
fun get_inner_fresh_fun (Bound j) = NONE |
|
| 33034 | 87 |
| get_inner_fresh_fun (v as Free _) = NONE |
| 22729 | 88 |
| get_inner_fresh_fun (v as Var _) = NONE |
89 |
| get_inner_fresh_fun (Const _) = NONE |
|
| 33034 | 90 |
| get_inner_fresh_fun (Abs (_, _, t)) = get_inner_fresh_fun t |
91 |
| get_inner_fresh_fun (Const ("Nominal.fresh_fun",Type("fun",[Type ("fun",[Type (T,_),_]),_])) $ u)
|
|
92 |
= SOME T |
|
93 |
| get_inner_fresh_fun (t $ u) = |
|
| 22729 | 94 |
let val a = get_inner_fresh_fun u in |
| 33034 | 95 |
if a = NONE then get_inner_fresh_fun t else a |
| 22729 | 96 |
end; |
97 |
||
| 33034 | 98 |
(* This tactic generates a fresh name of the atom type *) |
| 24558 | 99 |
(* given by the innermost fresh_fun *) |
| 22729 | 100 |
|
101 |
fun generate_fresh_fun_tac i thm = |
|
102 |
let |
|
103 |
val goal = List.nth(prems_of thm, i-1); |
|
104 |
val atom_name_opt = get_inner_fresh_fun goal; |
|
105 |
in |
|
| 33034 | 106 |
case atom_name_opt of |
| 22729 | 107 |
NONE => all_tac thm |
| 33034 | 108 |
| SOME atom_name => generate_fresh_tac atom_name i thm |
| 22729 | 109 |
end |
110 |
||
| 33034 | 111 |
(* Two substitution tactics which looks for the innermost occurence in |
|
23091
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
112 |
one assumption or in the conclusion *) |
| 22729 | 113 |
|
| 33034 | 114 |
val search_fun = curry (Seq.flat o uncurry EqSubst.searchf_bt_unify_valid); |
| 23065 | 115 |
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
|
116 |
|
| 33034 | 117 |
fun subst_inner_tac ctxt = EqSubst.eqsubst_tac' ctxt search_fun; |
118 |
fun subst_inner_asm_tac_aux i ctxt = EqSubst.eqsubst_asm_tac' ctxt 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
|
119 |
|
| 33034 | 120 |
(* A tactic to substitute in the first assumption |
|
23091
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
121 |
which contains an occurence. *) |
| 22729 | 122 |
|
| 33034 | 123 |
fun subst_inner_asm_tac ctxt th = |
124 |
curry (curry (FIRST' (map uncurry (map uncurry (map subst_inner_asm_tac_aux |
|
125 |
(1 upto Thm.nprems_of th)))))) ctxt th; |
|
|
23091
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
126 |
|
| 33034 | 127 |
fun fresh_fun_tac no_asm i thm = |
| 22729 | 128 |
(* Find the variable we instantiate *) |
129 |
let |
|
130 |
val thy = theory_of_thm thm; |
|
| 33034 | 131 |
val ctxt = ProofContext.init thy; |
|
32149
ef59550a55d3
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
wenzelm
parents:
30549
diff
changeset
|
132 |
val ss = global_simpset_of thy; |
|
26343
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26337
diff
changeset
|
133 |
val abs_fresh = PureThy.get_thms thy "abs_fresh"; |
|
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26337
diff
changeset
|
134 |
val fresh_perm_app = PureThy.get_thms thy "fresh_perm_app"; |
| 22729 | 135 |
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
|
136 |
val ss'' = ss' addsimps fresh_perm_app; |
|
29265
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents:
28397
diff
changeset
|
137 |
val x = hd (tl (OldTerm.term_vars (prop_of exI))); |
| 22787 | 138 |
val goal = nth (prems_of thm) (i-1); |
| 22729 | 139 |
val atom_name_opt = get_inner_fresh_fun goal; |
| 33034 | 140 |
val n = length (Logic.strip_params goal); |
| 24558 | 141 |
(* Here we rely on the fact that the variable introduced by generate_fresh_tac *) |
142 |
(* is the last one in the list, the inner one *) |
|
| 22729 | 143 |
in |
| 33034 | 144 |
case atom_name_opt of |
| 22729 | 145 |
NONE => all_tac thm |
| 33034 | 146 |
| SOME atom_name => |
147 |
let |
|
|
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30280
diff
changeset
|
148 |
val atom_basename = Long_Name.base_name atom_name; |
| 22729 | 149 |
val pt_name_inst = get_dyn_thm thy ("pt_"^atom_basename^"_inst") atom_basename;
|
150 |
val at_name_inst = get_dyn_thm thy ("at_"^atom_basename^"_inst") atom_basename;
|
|
| 22787 | 151 |
fun inst_fresh vars params i st = |
|
29265
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents:
28397
diff
changeset
|
152 |
let val vars' = OldTerm.term_vars (prop_of st); |
| 22787 | 153 |
val thy = theory_of_thm st; |
| 33040 | 154 |
in case subtract (op =) vars vars' of |
| 22787 | 155 |
[x] => Seq.single (Thm.instantiate ([],[(cterm_of thy x,cterm_of thy (list_abs (params,Bound 0)))]) st) |
156 |
| _ => error "fresh_fun_simp: Too many variables, please report." |
|
157 |
end |
|
| 22729 | 158 |
in |
|
23091
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
159 |
((fn st => |
| 33034 | 160 |
let |
|
29265
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents:
28397
diff
changeset
|
161 |
val vars = OldTerm.term_vars (prop_of st); |
| 22787 | 162 |
val params = Logic.strip_params (nth (prems_of st) (i-1)) |
| 33034 | 163 |
(* The tactics which solve the subgoals generated |
|
23091
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
164 |
by the conditionnal rewrite rule. *) |
| 33034 | 165 |
val post_rewrite_tacs = |
|
23054
d1bbe5afa279
change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents:
22792
diff
changeset
|
166 |
[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
|
167 |
rtac at_name_inst, |
| 33034 | 168 |
TRY o SOLVEI (NominalPermeq.finite_guess_tac ss''), |
|
23054
d1bbe5afa279
change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents:
22792
diff
changeset
|
169 |
inst_fresh vars params THEN' |
| 33034 | 170 |
(TRY o SOLVEI (NominalPermeq.fresh_guess_tac ss'')) THEN' |
171 |
(TRY o SOLVEI (asm_full_simp_tac ss''))] |
|
172 |
in |
|
|
23094
f1df8d2da98a
fix a bug : the semantics of no_asm was the opposite
narboux
parents:
23091
diff
changeset
|
173 |
((if no_asm then no_tac else |
| 33034 | 174 |
(subst_inner_asm_tac ctxt fresh_fun_app' i THEN (RANGE post_rewrite_tacs i))) |
|
23091
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
175 |
ORELSE |
| 33034 | 176 |
(subst_inner_tac ctxt fresh_fun_app' i THEN (RANGE post_rewrite_tacs i))) st |
| 22787 | 177 |
end)) thm |
| 33034 | 178 |
|
| 22729 | 179 |
end |
180 |
end |
|
181 |
||
|
23091
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
182 |
(* 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
|
183 |
gives back false *) |
|
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
184 |
val options_syntax = |
|
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
185 |
(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
|
186 |
(Scan.succeed false); |
|
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents:
23065
diff
changeset
|
187 |
|
| 30549 | 188 |
fun setup_generate_fresh x = |
189 |
(Args.goal_spec -- Args.tyname >> |
|
190 |
(fn (quant, s) => K (SIMPLE_METHOD'' quant (generate_fresh_tac s)))) x; |
|
| 22729 | 191 |
|
| 30549 | 192 |
fun setup_fresh_fun_simp x = |
193 |
(Scan.lift options_syntax >> (fn b => K (SIMPLE_METHOD' (fresh_fun_tac b)))) x; |
|
194 |