]> matita.cs.unibo.it Git - helm.git/blob - components/tactics/primitiveTactics.ml
07723ea9fe9483fd18afda332169373a7a8f6666
[helm.git] / components / tactics / primitiveTactics.ml
1 (* Copyright (C) 2002, HELM Team.
2  * 
3  * This file is part of HELM, an Hypertextual, Electronic
4  * Library of Mathematics, developed at the Computer Science
5  * Department, University of Bologna, Italy.
6  * 
7  * HELM is free software; you can redistribute it and/or
8  * modify it under the terms of the GNU General Public License
9  * as published by the Free Software Foundation; either version 2
10  * of the License, or (at your option) any later version.
11  * 
12  * HELM is distributed in the hope that it will be useful,
13  * but WITHOUT ANY WARRANTY; without even the implied warranty of
14  * MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE.  See the
15  * GNU General Public License for more details.
16  *
17  * You should have received a copy of the GNU General Public License
18  * along with HELM; if not, write to the Free Software
19  * Foundation, Inc., 59 Temple Place - Suite 330, Boston,
20  * MA  02111-1307, USA.
21  * 
22  * For details, see the HELM World-Wide-Web page,
23  * http://cs.unibo.it/helm/.
24  *)
25
26 (* $Id$ *)
27
28 open ProofEngineTypes
29
30 exception TheTypeOfTheCurrentGoalIsAMetaICannotChooseTheRightElimiantionPrinciple
31 exception NotAnInductiveTypeToEliminate
32 exception WrongUriToVariable of string
33 exception NotAnEliminator
34
35 (* lambda_abstract newmeta ty *)
36 (* returns a triple [bo],[context],[ty'] where              *)
37 (* [ty] = Pi/LetIn [context].[ty'] ([context] is a vector!) *)
38 (* and [bo] = Lambda/LetIn [context].(Meta [newmeta])       *)
39 (* So, lambda_abstract is the core of the implementation of *)
40 (* the Intros tactic.                                       *)
41 (* howmany = -1 means Intros, howmany > 0 means Intros n    *)
42 let lambda_abstract ?(howmany=(-1)) metasenv context newmeta ty mk_fresh_name =
43  let module C = Cic in
44   let rec collect_context context howmany do_whd ty =
45    match howmany with
46    | 0 ->  
47         let irl =
48           CicMkImplicit.identity_relocation_list_for_metavariable context
49         in
50          context, ty, (C.Meta (newmeta,irl))
51    | _ -> 
52       match ty with 
53         C.Cast (te,_)   -> collect_context context howmany do_whd te 
54       | C.Prod (n,s,t)  ->
55          let n' = mk_fresh_name metasenv context n ~typ:s in
56           let (context',ty,bo) =
57            let ctx = (Some (n',(C.Decl s)))::context in
58            collect_context ctx (howmany - 1) do_whd t 
59           in
60            (context',ty,C.Lambda(n',s,bo))
61       | C.LetIn (n,s,t) ->
62          let (context',ty,bo) =
63           collect_context ((Some (n,(C.Def (s,None))))::context) (howmany - 1) do_whd t
64          in
65           (context',ty,C.LetIn(n,s,bo))
66       | _ as t ->
67         if howmany <= 0 then
68          let irl =
69           CicMkImplicit.identity_relocation_list_for_metavariable context
70          in
71           context, t, (C.Meta (newmeta,irl))
72         else if do_whd then
73           let t = CicReduction.whd ~delta:true context t in
74           collect_context context howmany false t
75         else
76          raise (Fail (lazy "intro(s): not enough products or let-ins"))
77   in
78    collect_context context howmany true ty 
79
80 let eta_expand metasenv context t arg =
81  let module T = CicTypeChecker in
82  let module S = CicSubstitution in
83  let module C = Cic in
84   let rec aux n =
85    function
86       t' when t' = S.lift n arg -> C.Rel (1 + n)
87     | C.Rel m  -> if m <= n then C.Rel m else C.Rel (m+1)
88     | C.Var (uri,exp_named_subst) ->
89        let exp_named_subst' = aux_exp_named_subst n exp_named_subst in
90         C.Var (uri,exp_named_subst')
91     | C.Meta (i,l) ->
92        let l' =
93         List.map (function None -> None | Some t -> Some (aux n t)) l
94        in
95         C.Meta (i, l')
96     | C.Sort _
97     | C.Implicit _ as t -> t
98     | C.Cast (te,ty) -> C.Cast (aux n te, aux n ty)
99     | C.Prod (nn,s,t) -> C.Prod (nn, aux n s, aux (n+1) t)
100     | C.Lambda (nn,s,t) -> C.Lambda (nn, aux n s, aux (n+1) t)
101     | C.LetIn (nn,s,t) -> C.LetIn (nn, aux n s, aux (n+1) t)
102     | C.Appl l -> C.Appl (List.map (aux n) l)
103     | C.Const (uri,exp_named_subst) ->
104        let exp_named_subst' = aux_exp_named_subst n exp_named_subst in
105         C.Const (uri,exp_named_subst')
106     | C.MutInd (uri,i,exp_named_subst) ->
107        let exp_named_subst' = aux_exp_named_subst n exp_named_subst in
108         C.MutInd (uri,i,exp_named_subst')
109     | C.MutConstruct (uri,i,j,exp_named_subst) ->
110        let exp_named_subst' = aux_exp_named_subst n exp_named_subst in
111         C.MutConstruct (uri,i,j,exp_named_subst')
112     | C.MutCase (sp,i,outt,t,pl) ->
113        C.MutCase (sp,i,aux n outt, aux n t,
114         List.map (aux n) pl)
115     | C.Fix (i,fl) ->
116        let tylen = List.length fl in
117         let substitutedfl =
118          List.map
119           (fun (name,i,ty,bo) -> (name, i, aux n ty, aux (n+tylen) bo))
120            fl
121         in
122          C.Fix (i, substitutedfl)
123     | C.CoFix (i,fl) ->
124        let tylen = List.length fl in
125         let substitutedfl =
126          List.map
127           (fun (name,ty,bo) -> (name, aux n ty, aux (n+tylen) bo))
128            fl
129         in
130          C.CoFix (i, substitutedfl)
131   and aux_exp_named_subst n =
132    List.map (function uri,t -> uri,aux n t)
133   in
134    let argty,_ = 
135     T.type_of_aux' metasenv context arg CicUniv.empty_ugraph (* TASSI: FIXME *)
136    in
137     let fresh_name =
138      FreshNamesGenerator.mk_fresh_name ~subst:[]
139       metasenv context (Cic.Name "Heta") ~typ:argty
140     in
141      (C.Appl [C.Lambda (fresh_name,argty,aux 0 t) ; arg])
142
143 (*CSC: ma serve solamente la prima delle new_uninst e l'unione delle due!!! *)
144 let classify_metas newmeta in_subst_domain subst_in metasenv =
145  List.fold_right
146   (fun (i,canonical_context,ty) (old_uninst,new_uninst) ->
147     if in_subst_domain i then
148      old_uninst,new_uninst
149     else
150      let ty' = subst_in canonical_context ty in
151       let canonical_context' =
152        List.fold_right
153         (fun entry canonical_context' ->
154           let entry' =
155            match entry with
156               Some (n,Cic.Decl s) ->
157                Some (n,Cic.Decl (subst_in canonical_context' s))
158             | Some (n,Cic.Def (s,None)) ->
159                Some (n,Cic.Def ((subst_in canonical_context' s),None))
160             | None -> None
161             | Some (n,Cic.Def (bo,Some ty)) ->
162                Some
163                 (n,
164                   Cic.Def
165                    (subst_in canonical_context' bo,
166                     Some (subst_in canonical_context' ty)))
167           in
168            entry'::canonical_context'
169         ) canonical_context []
170      in
171       if i < newmeta then
172        ((i,canonical_context',ty')::old_uninst),new_uninst
173       else
174        old_uninst,((i,canonical_context',ty')::new_uninst)
175   ) metasenv ([],[])
176
177 (* Useful only inside apply_tac *)
178 let
179  generalize_exp_named_subst_with_fresh_metas context newmeta uri exp_named_subst
180 =
181  let module C = Cic in
182   let params =
183     let o,_ = CicEnvironment.get_obj CicUniv.empty_ugraph uri in
184     CicUtil.params_of_obj o
185   in
186    let exp_named_subst_diff,new_fresh_meta,newmetasenvfragment,exp_named_subst'=
187     let next_fresh_meta = ref newmeta in
188     let newmetasenvfragment = ref [] in
189     let exp_named_subst_diff = ref [] in
190      let rec aux =
191       function
192          [],[] -> []
193        | uri::tl,[] ->
194           let ty =
195             let o,_ = CicEnvironment.get_obj CicUniv.empty_ugraph uri in
196               match o with
197                   C.Variable (_,_,ty,_,_) ->
198                     CicSubstitution.subst_vars !exp_named_subst_diff ty
199                 | _ -> raise (WrongUriToVariable (UriManager.string_of_uri uri))
200           in
201 (* CSC: patch to generate ?1 : ?2 : Type in place of ?1 : Type to simulate ?1 :< Type
202            (match ty with
203                C.Sort (C.Type _) as s -> (* TASSI: ?? *)
204                  let fresh_meta = !next_fresh_meta in
205                  let fresh_meta' = fresh_meta + 1 in
206                   next_fresh_meta := !next_fresh_meta + 2 ;
207                   let subst_item = uri,C.Meta (fresh_meta',[]) in
208                    newmetasenvfragment :=
209                     (fresh_meta,[],C.Sort (C.Type (CicUniv.fresh()))) ::
210                      (* TASSI: ?? *)
211                      (fresh_meta',[],C.Meta (fresh_meta,[])) :: !newmetasenvfragment ;
212                    exp_named_subst_diff := !exp_named_subst_diff @ [subst_item] ;
213                    subst_item::(aux (tl,[]))
214              | _ ->
215 *)
216               let irl =
217                 CicMkImplicit.identity_relocation_list_for_metavariable context
218               in
219               let subst_item = uri,C.Meta (!next_fresh_meta,irl) in
220                newmetasenvfragment :=
221                 (!next_fresh_meta,context,ty)::!newmetasenvfragment ;
222                exp_named_subst_diff := !exp_named_subst_diff @ [subst_item] ;
223                incr next_fresh_meta ;
224                subst_item::(aux (tl,[]))(*)*)
225        | uri::tl1,((uri',_) as s)::tl2 ->
226           assert (UriManager.eq uri uri') ;
227           s::(aux (tl1,tl2))
228        | [],_ -> assert false
229      in
230       let exp_named_subst' = aux (params,exp_named_subst) in
231        !exp_named_subst_diff,!next_fresh_meta,
232         List.rev !newmetasenvfragment, exp_named_subst'
233    in
234     new_fresh_meta,newmetasenvfragment,exp_named_subst',exp_named_subst_diff
235 ;;
236
237 let new_metasenv_and_unify_and_t newmeta' metasenv' context term' ty termty goal_arity =
238   let (consthead,newmetasenv,arguments,_) =
239    TermUtil.saturate_term newmeta' metasenv' context termty
240     goal_arity in
241   let subst,newmetasenv',_ = 
242    CicUnification.fo_unif newmetasenv context consthead ty CicUniv.empty_ugraph
243   in
244   let t = 
245     if List.length arguments = 0 then term' else Cic.Appl (term'::arguments)
246   in
247   subst,newmetasenv',t
248
249 let rec count_prods context ty =
250  match CicReduction.whd context ty with
251     Cic.Prod (n,s,t) -> 1 + count_prods (Some (n,Cic.Decl s)::context) t
252   | _ -> 0
253
254 let apply_with_subst ~term ~subst ~maxmeta (proof, goal) =
255   (* Assumption: The term "term" must be closed in the current context *)
256  let module T = CicTypeChecker in
257  let module R = CicReduction in
258  let module C = Cic in
259   let (_,metasenv,_,_, _) = proof in
260   let metano,context,ty = CicUtil.lookup_meta goal metasenv in
261   let newmeta = max (CicMkImplicit.new_meta metasenv subst) maxmeta in
262    let exp_named_subst_diff,newmeta',newmetasenvfragment,term' =
263     match term with
264        C.Var (uri,exp_named_subst) ->
265         let newmeta',newmetasenvfragment,exp_named_subst',exp_named_subst_diff =
266          generalize_exp_named_subst_with_fresh_metas context newmeta uri
267           exp_named_subst
268         in
269          exp_named_subst_diff,newmeta',newmetasenvfragment,
270           C.Var (uri,exp_named_subst')
271      | C.Const (uri,exp_named_subst) ->
272         let newmeta',newmetasenvfragment,exp_named_subst',exp_named_subst_diff =
273          generalize_exp_named_subst_with_fresh_metas context newmeta uri
274           exp_named_subst
275         in
276          exp_named_subst_diff,newmeta',newmetasenvfragment,
277           C.Const (uri,exp_named_subst')
278      | C.MutInd (uri,tyno,exp_named_subst) ->
279         let newmeta',newmetasenvfragment,exp_named_subst',exp_named_subst_diff =
280          generalize_exp_named_subst_with_fresh_metas context newmeta uri
281           exp_named_subst
282         in
283          exp_named_subst_diff,newmeta',newmetasenvfragment,
284           C.MutInd (uri,tyno,exp_named_subst')
285      | C.MutConstruct (uri,tyno,consno,exp_named_subst) ->
286         let newmeta',newmetasenvfragment,exp_named_subst',exp_named_subst_diff =
287          generalize_exp_named_subst_with_fresh_metas context newmeta uri
288           exp_named_subst
289         in
290          exp_named_subst_diff,newmeta',newmetasenvfragment,
291           C.MutConstruct (uri,tyno,consno,exp_named_subst')
292      | _ -> [],newmeta,[],term
293    in
294    let metasenv' = metasenv@newmetasenvfragment in
295    let termty,_ = 
296      CicTypeChecker.type_of_aux' metasenv' context term' CicUniv.empty_ugraph
297    in
298    let termty =
299      CicSubstitution.subst_vars exp_named_subst_diff termty in
300    let goal_arity = count_prods context ty in
301    let subst,newmetasenv',t = 
302     let rec add_one_argument n =
303      try
304       new_metasenv_and_unify_and_t newmeta' metasenv' context term' ty
305         termty n
306      with CicUnification.UnificationFailure _ when n > 0 ->
307       add_one_argument (n - 1)
308     in
309      add_one_argument goal_arity
310    in
311    let in_subst_domain i = List.exists (function (j,_) -> i=j) subst in
312    let apply_subst = CicMetaSubst.apply_subst subst in
313    let old_uninstantiatedmetas,new_uninstantiatedmetas =
314      (* subst_in doesn't need the context. Hence the underscore. *)
315      let subst_in _ = CicMetaSubst.apply_subst subst in
316      classify_metas newmeta in_subst_domain subst_in newmetasenv'
317    in
318    let bo' = apply_subst t in
319    let newmetasenv'' = new_uninstantiatedmetas@old_uninstantiatedmetas in
320    let subst_in =
321      (* if we just apply the subtitution, the type is irrelevant:
322               we may use Implicit, since it will be dropped *)
323      CicMetaSubst.apply_subst ((metano,(context,bo',Cic.Implicit None))::subst)
324    in
325    let (newproof, newmetasenv''') = 
326     ProofEngineHelpers.subst_meta_and_metasenv_in_proof proof metano subst_in
327      newmetasenv''
328    in
329    let subst = ((metano,(context,bo',Cic.Implicit None))::subst) in
330    subst,
331    (newproof, List.map (function (i,_,_) -> i) new_uninstantiatedmetas),
332    max maxmeta (CicMkImplicit.new_meta newmetasenv''' subst)
333
334
335 (* ALB *)
336 let apply_with_subst ~term ?(subst=[]) ?(maxmeta=0) status =
337   try
338 (*     apply_tac_verbose ~term status *)
339     apply_with_subst ~term ~subst ~maxmeta status
340       (* TODO cacciare anche altre eccezioni? *)
341   with 
342   | CicUnification.UnificationFailure msg
343   | CicTypeChecker.TypeCheckerFailure msg -> raise (Fail msg)
344
345 (* ALB *)
346 let apply_tac_verbose ~term status =
347   let subst, status, _ = apply_with_subst ~term status in
348   (CicMetaSubst.apply_subst subst), status
349
350 let apply_tac ~term status = snd (apply_tac_verbose ~term status)
351
352   (* TODO per implementare i tatticali e' necessario che tutte le tattiche
353   sollevino _solamente_ Fail *)
354 let apply_tac ~term =
355  let apply_tac ~term status =
356   try
357     apply_tac ~term status
358       (* TODO cacciare anche altre eccezioni? *)
359   with 
360   | CicUnification.UnificationFailure msg
361   | CicTypeChecker.TypeCheckerFailure msg ->
362       raise (Fail msg)
363  in
364   mk_tactic (apply_tac ~term)
365
366 let intros_tac ?howmany ?(mk_fresh_name_callback = FreshNamesGenerator.mk_fresh_name ~subst:[]) ()=
367  let intros_tac
368   ?(mk_fresh_name_callback = (FreshNamesGenerator.mk_fresh_name ~subst:[])) ()
369   (proof, goal)
370  =
371   let module C = Cic in
372   let module R = CicReduction in
373    let (_,metasenv,_,_, _) = proof in
374    let metano,context,ty = CicUtil.lookup_meta goal metasenv in
375     let newmeta = ProofEngineHelpers.new_meta_of_proof ~proof in
376      let (context',ty',bo') =
377       lambda_abstract ?howmany metasenv context newmeta ty mk_fresh_name_callback
378      in
379       let (newproof, _) =
380        ProofEngineHelpers.subst_meta_in_proof proof metano bo'
381         [newmeta,context',ty']
382       in
383        (newproof, [newmeta])
384  in
385   mk_tactic (intros_tac ~mk_fresh_name_callback ())
386   
387 let cut_tac ?(mk_fresh_name_callback = FreshNamesGenerator.mk_fresh_name ~subst:[]) term =
388  let cut_tac
389   ?(mk_fresh_name_callback = FreshNamesGenerator.mk_fresh_name ~subst:[])
390   term (proof, goal)
391  =
392   let module C = Cic in
393    let curi,metasenv,pbo,pty, attrs = proof in
394    let metano,context,ty = CicUtil.lookup_meta goal metasenv in
395     let newmeta1 = ProofEngineHelpers.new_meta_of_proof ~proof in
396     let newmeta2 = newmeta1 + 1 in
397     let fresh_name =
398      mk_fresh_name_callback metasenv context (Cic.Name "Hcut") ~typ:term in
399     let context_for_newmeta1 =
400      (Some (fresh_name,C.Decl term))::context in
401     let irl1 =
402      CicMkImplicit.identity_relocation_list_for_metavariable
403       context_for_newmeta1
404     in
405     let irl2 =
406       CicMkImplicit.identity_relocation_list_for_metavariable context
407     in
408      let newmeta1ty = CicSubstitution.lift 1 ty in
409      let bo' =
410       C.Appl
411        [C.Lambda (fresh_name,term,C.Meta (newmeta1,irl1)) ;
412         C.Meta (newmeta2,irl2)]
413      in
414       let (newproof, _) =
415        ProofEngineHelpers.subst_meta_in_proof proof metano bo'
416         [newmeta2,context,term; newmeta1,context_for_newmeta1,newmeta1ty];
417       in
418        (newproof, [newmeta1 ; newmeta2])
419  in
420   mk_tactic (cut_tac ~mk_fresh_name_callback term)
421
422 let letin_tac ?(mk_fresh_name_callback=FreshNamesGenerator.mk_fresh_name ~subst:[]) term =
423  let letin_tac
424   ?(mk_fresh_name_callback = FreshNamesGenerator.mk_fresh_name ~subst:[])
425   term (proof, goal)
426  =
427   let module C = Cic in
428    let curi,metasenv,pbo,pty, attrs = proof in
429    (* occur check *)
430    let occur i t =
431      let m = CicUtil.metas_of_term t in 
432      List.exists (fun (j,_) -> i=j) m
433    in
434    let metano,context,ty = CicUtil.lookup_meta goal metasenv in
435    if occur metano term then
436      raise 
437        (ProofEngineTypes.Fail (lazy
438          "You can't letin a term containing the current goal"));
439     let _,_ =
440       CicTypeChecker.type_of_aux' metasenv context term CicUniv.empty_ugraph in
441      let newmeta = ProofEngineHelpers.new_meta_of_proof ~proof in
442      let fresh_name =
443       mk_fresh_name_callback metasenv context (Cic.Name "Hletin") ~typ:term in
444      let context_for_newmeta =
445       (Some (fresh_name,C.Def (term,None)))::context in
446      let irl =
447       CicMkImplicit.identity_relocation_list_for_metavariable
448        context_for_newmeta
449      in
450       let newmetaty = CicSubstitution.lift 1 ty in
451       let bo' = C.LetIn (fresh_name,term,C.Meta (newmeta,irl)) in
452        let (newproof, _) =
453          ProofEngineHelpers.subst_meta_in_proof
454            proof metano bo'[newmeta,context_for_newmeta,newmetaty]
455        in
456         (newproof, [newmeta])
457  in
458   mk_tactic (letin_tac ~mk_fresh_name_callback term)
459
460   (** functional part of the "exact" tactic *)
461 let exact_tac ~term =
462  let exact_tac ~term (proof, goal) =
463   (* Assumption: the term bo must be closed in the current context *)
464   let (_,metasenv,_,_, _) = proof in
465   let metano,context,ty = CicUtil.lookup_meta goal metasenv in
466   let module T = CicTypeChecker in
467   let module R = CicReduction in
468   let ty_term,u = T.type_of_aux' metasenv context term CicUniv.empty_ugraph in
469   let b,_ = R.are_convertible context ty_term ty u in (* TASSI: FIXME *)
470   if b then
471    begin
472     let (newproof, metasenv') =
473       ProofEngineHelpers.subst_meta_in_proof proof metano term [] in
474     (newproof, [])
475    end
476   else
477    raise (Fail (lazy "The type of the provided term is not the one expected."))
478  in
479   mk_tactic (exact_tac ~term)
480
481 (* not really "primitive" tactics .... *)
482   
483 module TC  = CicTypeChecker
484 module U   = UriManager
485 module R   = CicReduction
486 module C   = Cic
487 module PET = ProofEngineTypes
488 module PEH = ProofEngineHelpers
489 module PER = ProofEngineReduction
490 module MS  = CicMetaSubst 
491 module S   = CicSubstitution 
492 module T   = Tacticals
493 module RT  = ReductionTactics
494
495 let elim_tac ?using ?(pattern = PET.conclusion_pattern None) term = 
496  let elim_tac (proof, goal) =
497    let ugraph = CicUniv.empty_ugraph in
498    let curi, metasenv, proofbo, proofty, attrs = proof in
499    let conjecture = CicUtil.lookup_meta goal metasenv in
500    let metano, context, ty = conjecture in
501 (*    let (term, metasenv, _ugraph), cpatt = match pattern with 
502        | Some f, [], Some cpatt -> f context metasenv ugraph, cpatt
503        | _                      -> assert false
504     in
505 *)    
506     let termty,_ugraph = TC.type_of_aux' metasenv context term ugraph in
507     let termty = CicReduction.whd context termty in
508     let (termty,metasenv',arguments,_fresh_meta) =
509      TermUtil.saturate_term
510       (ProofEngineHelpers.new_meta_of_proof proof) metasenv context termty 0 in
511     let term = if arguments = [] then term else Cic.Appl (term::arguments) in
512     let uri,exp_named_subst,typeno,args =
513      match termty with
514         C.MutInd (uri,typeno,exp_named_subst) -> (uri,exp_named_subst,typeno,[])
515       | C.Appl ((C.MutInd (uri,typeno,exp_named_subst))::args) ->
516           (uri,exp_named_subst,typeno,args)
517       | _ -> raise NotAnInductiveTypeToEliminate
518     in
519      let eliminator_uri =
520       let buri = U.buri_of_uri uri in
521       let name = 
522         let o,_ugraph = CicEnvironment.get_obj ugraph uri in
523        match o with
524           C.InductiveDefinition (tys,_,_,_) ->
525            let (name,_,_,_) = List.nth tys typeno in
526             name
527         | _ -> assert false
528       in
529       let ty_ty,_ugraph = TC.type_of_aux' metasenv' context ty ugraph in
530       let ext =
531        match ty_ty with
532           C.Sort C.Prop -> "_ind"
533         | C.Sort C.Set  -> "_rec"
534         | C.Sort C.CProp -> "_rec"
535         | C.Sort (C.Type _)-> "_rect" 
536         | C.Meta (_,_) -> raise TheTypeOfTheCurrentGoalIsAMetaICannotChooseTheRightElimiantionPrinciple
537         | _ -> assert false
538       in
539        U.uri_of_string (buri ^ "/" ^ name ^ ext ^ ".con")
540      in
541       let eliminator_ref = match using with
542          | None   -> C.Const (eliminator_uri,exp_named_subst)
543          | Some t -> t 
544        in
545        let ety,_ugraph = 
546          TC.type_of_aux' metasenv' context eliminator_ref ugraph in
547 (* FG: ADDED PART ***********************************************************)
548 (* FG: we can not assume eliminator is the default eliminator ***************)
549 (*
550    let add_lambdas n t =
551       let rec aux n t =
552          if n <= 0 then t 
553          else C.Lambda (C.Anonymous, C.Implicit None, aux (pred n) t)
554       in
555       aux n (S.lift n t)
556    in
557 *)
558    let rec args_init n f =
559       if n <= 0 then [] else f n :: args_init (pred n) f
560    in
561    let splits, args_no = PEH.split_with_whd (context, ety) in
562    let pred_pos = match List.hd splits with
563       | _, C.Rel i when i > 1 && i <= args_no -> i
564       | _, C.Appl (C.Rel i :: _) when i > 1 && i <= args_no -> i
565       | _ -> raise NotAnEliminator
566    in
567 (*
568    let _, lambdas = PEH.split_with_whd (List.nth splits pred_pos) in
569    let termty_ty =
570       let termty_ty,_ugraph = TC.type_of_aux' metasenv' context termty ugraph in
571       CicReduction.whd context termty_ty
572    in
573 *)
574 (*   
575    let metasenv', term, pred, upto = match cpatt, termty_ty with
576       | C.Implicit (Some `Hole), _ 
577       | _, C.Sort C.Prop when lambdas = 0 -> metasenv', term, C.Implicit None, 0
578       | _                                 ->
579 (* FG: we find the predicate for the eliminator as in the rewrite tactic ****)
580          let fresh_name = 
581              FreshNamesGenerator.mk_fresh_name 
582              ~subst:[] metasenv' context C.Anonymous ~typ:termty
583          in
584          let lazy_term c m u =  
585             let distance  = List.length c - List.length context in
586             S.lift distance term, m, u
587          in
588          let pattern = Some lazy_term, [], Some cpatt in
589          let subst, metasenv', _ugraph, _conjecture, selected_terms =
590             ProofEngineHelpers.select
591             ~metasenv:metasenv' ~ugraph ~conjecture ~pattern
592          in
593          let metasenv' = MS.apply_subst_metasenv subst metasenv' in  
594          let map (_context_of_t, t) l = t :: l in
595          let what = List.fold_right map selected_terms [] in
596          let ty = MS.apply_subst subst ty in
597          let term = MS.apply_subst subst term in
598          let termty = MS.apply_subst subst termty in
599          let abstr_ty = PER.replace_with_rel_1_from ~equality:(==) ~what 1 ty in
600          let abstr_ty = MS.apply_subst subst abstr_ty in
601          let pred_body = C.Lambda (fresh_name, termty, abstr_ty) in
602          metasenv', term, add_lambdas (pred lambdas) pred_body, lambdas 
603    in
604 (* FG: END OF ADDED PART ****************************************************)
605 *)
606          let pred, upto = C.Implicit None, 0 in
607          
608          let term_to_refine =
609           let f n =
610            if n = pred_pos then pred else
611            if n = 1 then term else C.Implicit None
612           in
613            C.Appl (eliminator_ref :: args_init args_no f)
614          in
615           let refined_term,_refined_termty,metasenv'',_ugraph = 
616            CicRefine.type_of_aux' metasenv' context term_to_refine
617              ugraph
618           in
619            let new_goals =
620             ProofEngineHelpers.compare_metasenvs
621              ~oldmetasenv:metasenv ~newmetasenv:metasenv''
622            in
623            let proof' = curi,metasenv'',proofbo,proofty, attrs in
624             let proof'', new_goals' =
625              apply_tactic (apply_tac ~term:refined_term) (proof',goal)
626             in
627              (* The apply_tactic can have closed some of the new_goals *)
628              let patched_new_goals =
629               let (_,metasenv''',_,_, _) = proof'' in
630                List.filter
631                 (function i -> List.exists (function (j,_,_) -> j=i) metasenv'''
632                 ) new_goals @ new_goals'
633              in
634               let res = proof'', patched_new_goals in
635                if upto = 0 then res else 
636                 let pattern = PET.conclusion_pattern None in
637                 let continuation =
638                  RT.simpl_tac ~pattern
639                  (* RT.head_beta_reduce_tac ~delta:false ~upto ~pattern *)
640                 in
641                  let dummy_status = proof,goal in
642                   PET.apply_tactic
643                    (T.then_ ~start:(PET.mk_tactic (fun _ -> res)) ~continuation)
644                    dummy_status
645  in
646   mk_tactic elim_tac
647 ;;
648
649 let cases_intros_tac ?(mk_fresh_name_callback = FreshNamesGenerator.mk_fresh_name ~subst:[]) term =
650  let cases_tac ~term (proof, goal) =
651   let module TC = CicTypeChecker in
652   let module U = UriManager in
653   let module R = CicReduction in
654   let module C = Cic in
655   let (curi,metasenv,proofbo,proofty, attrs) = proof in
656   let metano,context,ty = CicUtil.lookup_meta goal metasenv in
657   let termty,_ = TC.type_of_aux' metasenv context term CicUniv.empty_ugraph in
658   let termty = CicReduction.whd context termty in
659   let (termty,metasenv',arguments,fresh_meta) =
660    TermUtil.saturate_term
661     (ProofEngineHelpers.new_meta_of_proof proof) metasenv context termty 0 in
662   let term = if arguments = [] then term else Cic.Appl (term::arguments) in
663   let uri,exp_named_subst,typeno,args =
664     match termty with
665     | C.MutInd (uri,typeno,exp_named_subst) -> (uri,exp_named_subst,typeno,[])
666     | C.Appl ((C.MutInd (uri,typeno,exp_named_subst))::args) ->
667         (uri,exp_named_subst,typeno,args)
668     | _ -> raise NotAnInductiveTypeToEliminate
669   in
670   let paramsno,itty,patterns,right_args =
671     match CicEnvironment.get_obj CicUniv.empty_ugraph uri with
672     | C.InductiveDefinition (tys,_,paramsno,_),_ ->
673        let _,left_parameters,right_args = 
674          List.fold_right 
675            (fun x (n,acc1,acc2) -> 
676              if n > 0 then (n-1,acc1,x::acc2) else (n,x::acc1,acc2)) 
677            args (List.length args - paramsno, [],[])
678        in
679        let _,_,itty,cl = List.nth tys typeno in
680        let rec aux left_parameters context t =
681          match left_parameters,CicReduction.whd context t with
682          | [],C.Prod (name,source,target) ->
683             let fresh_name =
684               mk_fresh_name_callback metasenv' context name ~typ:source
685             in
686              C.Lambda (fresh_name,C.Implicit None,
687              aux [] (Some (fresh_name,C.Decl source)::context) target)
688          | hd::tl,C.Prod (name,source,target) ->
689              (* left parameters instantiation *)
690              aux tl context (CicSubstitution.subst hd target)
691          | [],_ -> C.Implicit None
692          | _ -> assert false
693        in
694         paramsno,itty,
695         List.map (function (_,cty) -> aux left_parameters context cty) cl,
696         right_args
697     | _ -> assert false
698   in
699   let outtype =
700     let n_right_args = List.length right_args in
701     let n_lambdas = n_right_args + 1 in
702     let lifted_ty = CicSubstitution.lift n_lambdas ty in
703     let replace = ProofEngineReduction.replace_lifting
704        ~equality:(ProofEngineReduction.alpha_equivalence)
705     in
706     let captured_ty = 
707       let what = 
708         List.map (CicSubstitution.lift n_lambdas) (right_args@[term])
709       in
710       let with_what = 
711         let rec mkargs = function 
712           | 0 -> []
713           | 1 -> [Cic.Rel 1]
714           | n -> (Cic.Implicit None)::(mkargs (n-1)) 
715         in
716         mkargs n_lambdas 
717       in
718       replace ~what ~with_what ~where:lifted_ty
719     in
720     let captured_term_ty = 
721       let term_ty = CicSubstitution.lift n_right_args termty in
722       let rec mkrels = function 0 -> []|n -> (Cic.Rel n)::(mkrels (n-1)) in
723       let rec fstn acc l n = 
724         if n = 0 then acc else fstn (acc@[List.hd l]) (List.tl l) (n-1) 
725       in
726       match term_ty with
727       | C.MutInd _ -> term_ty
728       | C.Appl ((C.MutInd (a,b,c))::args) -> 
729            C.Appl ((C.MutInd (a,b,c))::
730                fstn [] args paramsno @ mkrels n_right_args)
731       | _ -> raise NotAnInductiveTypeToEliminate
732     in
733     let rec add_lambdas = function
734       | 0 -> captured_ty
735       | 1 -> 
736           C.Lambda (C.Name "matched", captured_term_ty, (add_lambdas 0))
737       | n -> 
738            C.Lambda (C.Name ("right_"^(string_of_int (n-1))),
739                      C.Implicit None, (add_lambdas (n-1)))
740     in
741     add_lambdas n_lambdas
742   in
743   let term_to_refine = C.MutCase (uri,typeno,outtype,term,patterns) in
744   let refined_term,_,metasenv'',_ = 
745     CicRefine.type_of_aux' metasenv' context term_to_refine
746       CicUniv.empty_ugraph
747   in
748   let new_goals =
749     ProofEngineHelpers.compare_metasenvs
750       ~oldmetasenv:metasenv ~newmetasenv:metasenv''
751   in
752   let proof' = curi,metasenv'',proofbo,proofty, attrs in
753   let proof'', new_goals' =
754     apply_tactic (apply_tac ~term:refined_term) (proof',goal)
755   in
756   (* The apply_tactic can have closed some of the new_goals *)
757   let patched_new_goals =
758     let (_,metasenv''',_,_,_) = proof'' in
759       List.filter
760         (function i -> List.exists (function (j,_,_) -> j=i) metasenv''')
761         new_goals @ new_goals'
762     in
763     proof'', patched_new_goals
764   in
765   mk_tactic (cases_tac ~term)
766 ;;
767
768
769 let elim_intros_tac ?(mk_fresh_name_callback = FreshNamesGenerator.mk_fresh_name ~subst:[]) 
770                     ?depth ?using ?pattern what =
771  Tacticals.then_ ~start:(elim_tac ?using ?pattern what)
772   ~continuation:(intros_tac ~mk_fresh_name_callback ?howmany:depth ())
773 ;;
774
775 (* The simplification is performed only on the conclusion *)
776 let elim_intros_simpl_tac ?(mk_fresh_name_callback = FreshNamesGenerator.mk_fresh_name ~subst:[])
777                           ?depth ?using ?pattern what =
778  Tacticals.then_ ~start:(elim_tac ?using ?pattern what)
779   ~continuation:
780    (Tacticals.thens
781      ~start:(intros_tac ~mk_fresh_name_callback ?howmany:depth ())
782      ~continuations:
783        [ReductionTactics.simpl_tac
784          ~pattern:(ProofEngineTypes.conclusion_pattern None)])
785 ;;
786
787 (* FG: insetrts a "hole" in the context (derived from letin_tac) *)
788
789 let letout_tac =
790    let mk_fresh_name_callback = FreshNamesGenerator.mk_fresh_name ~subst:[] in
791    let term = C.Sort C.Set in
792    let letout_tac (proof, goal) =
793       let curi, metasenv, pbo, pty, attrs = proof in
794       let metano, context, ty = CicUtil.lookup_meta goal metasenv in
795       let newmeta = ProofEngineHelpers.new_meta_of_proof ~proof in
796       let fresh_name = mk_fresh_name_callback metasenv context (Cic.Name "hole") ~typ:term in
797       let context_for_newmeta = None :: context in
798       let irl = CicMkImplicit.identity_relocation_list_for_metavariable context_for_newmeta in
799       let newmetaty = CicSubstitution.lift 1 ty in
800       let bo' = C.LetIn (fresh_name, term, C.Meta (newmeta,irl)) in
801       let newproof, _ = ProofEngineHelpers.subst_meta_in_proof proof metano bo'[newmeta,context_for_newmeta,newmetaty] in
802       newproof, [newmeta]
803    in
804    mk_tactic letout_tac