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