]> matita.cs.unibo.it Git - helm.git/blob - components/tactics/primitiveTactics.ml
elim with a pattern now works correctly (hopefully)
[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 exception TheTypeOfTheCurrentGoalIsAMetaICannotChooseTheRightElimiantionPrinciple
29 exception NotAnInductiveTypeToEliminate
30 exception WrongUriToVariable of string
31 exception NotAnEliminator
32
33 module PET = ProofEngineTypes
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 (PET.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' subst 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_subst 
243      subst context newmetasenv consthead ty CicUniv.empty_ugraph
244   in
245   let t = 
246     if List.length arguments = 0 then term' else Cic.Appl (term'::arguments)
247   in
248   subst,newmetasenv',t
249
250 let rec count_prods context ty =
251  match CicReduction.whd context ty with
252     Cic.Prod (n,s,t) -> 1 + count_prods (Some (n,Cic.Decl s)::context) t
253   | _ -> 0
254
255 let apply_with_subst ~term ~subst ~maxmeta (proof, goal) =
256   (* Assumption: The term "term" must be closed in the current context *)
257  let module T = CicTypeChecker in
258  let module R = CicReduction in
259  let module C = Cic in
260   let (_,metasenv,_,_, _) = proof in
261   let metano,context,ty = CicUtil.lookup_meta goal metasenv in
262   let newmeta = max (CicMkImplicit.new_meta metasenv subst) maxmeta in
263    let exp_named_subst_diff,newmeta',newmetasenvfragment,term' =
264     match term with
265        C.Var (uri,exp_named_subst) ->
266         let newmeta',newmetasenvfragment,exp_named_subst',exp_named_subst_diff =
267          generalize_exp_named_subst_with_fresh_metas context newmeta uri
268           exp_named_subst
269         in
270          exp_named_subst_diff,newmeta',newmetasenvfragment,
271           C.Var (uri,exp_named_subst')
272      | C.Const (uri,exp_named_subst) ->
273         let newmeta',newmetasenvfragment,exp_named_subst',exp_named_subst_diff =
274          generalize_exp_named_subst_with_fresh_metas context newmeta uri
275           exp_named_subst
276         in
277          exp_named_subst_diff,newmeta',newmetasenvfragment,
278           C.Const (uri,exp_named_subst')
279      | C.MutInd (uri,tyno,exp_named_subst) ->
280         let newmeta',newmetasenvfragment,exp_named_subst',exp_named_subst_diff =
281          generalize_exp_named_subst_with_fresh_metas context newmeta uri
282           exp_named_subst
283         in
284          exp_named_subst_diff,newmeta',newmetasenvfragment,
285           C.MutInd (uri,tyno,exp_named_subst')
286      | C.MutConstruct (uri,tyno,consno,exp_named_subst) ->
287         let newmeta',newmetasenvfragment,exp_named_subst',exp_named_subst_diff =
288          generalize_exp_named_subst_with_fresh_metas context newmeta uri
289           exp_named_subst
290         in
291          exp_named_subst_diff,newmeta',newmetasenvfragment,
292           C.MutConstruct (uri,tyno,consno,exp_named_subst')
293      | _ -> [],newmeta,[],term
294    in
295    let metasenv' = metasenv@newmetasenvfragment in
296    let termty,_ = 
297      CicTypeChecker.type_of_aux' metasenv' context term' CicUniv.empty_ugraph
298    in
299    let termty =
300      CicSubstitution.subst_vars exp_named_subst_diff termty in
301    let goal_arity = count_prods context ty in
302    let subst,newmetasenv',t = 
303     let rec add_one_argument n =
304      try
305       new_metasenv_and_unify_and_t newmeta' metasenv' subst context term' ty
306         termty n
307      with CicUnification.UnificationFailure _ when n > 0 ->
308       add_one_argument (n - 1)
309     in
310      add_one_argument goal_arity
311    in
312    let in_subst_domain i = List.exists (function (j,_) -> i=j) subst in
313    let apply_subst = CicMetaSubst.apply_subst subst in
314    let old_uninstantiatedmetas,new_uninstantiatedmetas =
315      (* subst_in doesn't need the context. Hence the underscore. *)
316      let subst_in _ = CicMetaSubst.apply_subst subst in
317      classify_metas newmeta in_subst_domain subst_in newmetasenv'
318    in
319    let bo' = apply_subst t in
320    let newmetasenv'' = new_uninstantiatedmetas@old_uninstantiatedmetas in
321    let subst_in =
322      (* if we just apply the subtitution, the type is irrelevant:
323               we may use Implicit, since it will be dropped *)
324      CicMetaSubst.apply_subst ((metano,(context,bo',Cic.Implicit None))::subst)
325    in
326    let (newproof, newmetasenv''') = 
327     ProofEngineHelpers.subst_meta_and_metasenv_in_proof proof metano subst_in
328      newmetasenv''
329    in
330    let subst = ((metano,(context,bo',ty))::subst) in
331    subst,
332    (newproof, List.map (function (i,_,_) -> i) new_uninstantiatedmetas),
333    max maxmeta (CicMkImplicit.new_meta newmetasenv''' subst)
334
335
336 (* ALB *)
337 let apply_with_subst ~term ?(subst=[]) ?(maxmeta=0) status =
338   try
339 (*     apply_tac_verbose ~term status *)
340     apply_with_subst ~term ~subst ~maxmeta status
341       (* TODO cacciare anche altre eccezioni? *)
342   with 
343   | CicUnification.UnificationFailure msg
344   | CicTypeChecker.TypeCheckerFailure msg -> raise (PET.Fail msg)
345
346 (* ALB *)
347 let apply_tac_verbose ~term status =
348   let subst, status, _ = apply_with_subst ~term status in
349   (CicMetaSubst.apply_subst subst), status
350
351 let apply_tac ~term status = snd (apply_tac_verbose ~term status)
352
353   (* TODO per implementare i tatticali e' necessario che tutte le tattiche
354   sollevino _solamente_ Fail *)
355 let apply_tac ~term =
356  let apply_tac ~term status =
357   try
358     apply_tac ~term status
359       (* TODO cacciare anche altre eccezioni? *)
360   with 
361   | CicUnification.UnificationFailure msg
362   | CicTypeChecker.TypeCheckerFailure msg ->
363       raise (PET.Fail msg)
364  in
365   PET.mk_tactic (apply_tac ~term)
366
367 let intros_tac ?howmany ?(mk_fresh_name_callback = FreshNamesGenerator.mk_fresh_name ~subst:[]) ()=
368  let intros_tac
369   ?(mk_fresh_name_callback = (FreshNamesGenerator.mk_fresh_name ~subst:[])) ()
370   (proof, goal)
371  =
372   let module C = Cic in
373   let module R = CicReduction in
374    let (_,metasenv,_,_, _) = proof in
375    let metano,context,ty = CicUtil.lookup_meta goal metasenv in
376     let newmeta = ProofEngineHelpers.new_meta_of_proof ~proof in
377      let (context',ty',bo') =
378       lambda_abstract ?howmany metasenv context newmeta ty mk_fresh_name_callback
379      in
380       let (newproof, _) =
381        ProofEngineHelpers.subst_meta_in_proof proof metano bo'
382         [newmeta,context',ty']
383       in
384        (newproof, [newmeta])
385  in
386   PET.mk_tactic (intros_tac ~mk_fresh_name_callback ())
387   
388 let cut_tac ?(mk_fresh_name_callback = FreshNamesGenerator.mk_fresh_name ~subst:[]) term =
389  let cut_tac
390   ?(mk_fresh_name_callback = FreshNamesGenerator.mk_fresh_name ~subst:[])
391   term (proof, goal)
392  =
393   let module C = Cic in
394    let curi,metasenv,pbo,pty, attrs = proof in
395    let metano,context,ty = CicUtil.lookup_meta goal metasenv in
396     let newmeta1 = ProofEngineHelpers.new_meta_of_proof ~proof in
397     let newmeta2 = newmeta1 + 1 in
398     let fresh_name =
399      mk_fresh_name_callback metasenv context (Cic.Name "Hcut") ~typ:term in
400     let context_for_newmeta1 =
401      (Some (fresh_name,C.Decl term))::context in
402     let irl1 =
403      CicMkImplicit.identity_relocation_list_for_metavariable
404       context_for_newmeta1
405     in
406     let irl2 =
407       CicMkImplicit.identity_relocation_list_for_metavariable context
408     in
409      let newmeta1ty = CicSubstitution.lift 1 ty in
410      let bo' =
411       C.Appl
412        [C.Lambda (fresh_name,term,C.Meta (newmeta1,irl1)) ;
413         C.Meta (newmeta2,irl2)]
414      in
415       let (newproof, _) =
416        ProofEngineHelpers.subst_meta_in_proof proof metano bo'
417         [newmeta2,context,term; newmeta1,context_for_newmeta1,newmeta1ty];
418       in
419        (newproof, [newmeta1 ; newmeta2])
420  in
421   PET.mk_tactic (cut_tac ~mk_fresh_name_callback term)
422
423 let letin_tac ?(mk_fresh_name_callback=FreshNamesGenerator.mk_fresh_name ~subst:[]) term =
424  let letin_tac
425   ?(mk_fresh_name_callback = FreshNamesGenerator.mk_fresh_name ~subst:[])
426   term (proof, goal)
427  =
428   let module C = Cic in
429    let curi,metasenv,pbo,pty, attrs = proof in
430    (* occur check *)
431    let occur i t =
432      let m = CicUtil.metas_of_term t in 
433      List.exists (fun (j,_) -> i=j) m
434    in
435    let metano,context,ty = CicUtil.lookup_meta goal metasenv in
436    if occur metano term then
437      raise 
438        (ProofEngineTypes.Fail (lazy
439          "You can't letin a term containing the current goal"));
440     let _,_ =
441       CicTypeChecker.type_of_aux' metasenv context term CicUniv.empty_ugraph in
442      let newmeta = ProofEngineHelpers.new_meta_of_proof ~proof in
443      let fresh_name =
444       mk_fresh_name_callback metasenv context (Cic.Name "Hletin") ~typ:term in
445      let context_for_newmeta =
446       (Some (fresh_name,C.Def (term,None)))::context in
447      let irl =
448       CicMkImplicit.identity_relocation_list_for_metavariable
449        context_for_newmeta
450      in
451       let newmetaty = CicSubstitution.lift 1 ty in
452       let bo' = C.LetIn (fresh_name,term,C.Meta (newmeta,irl)) in
453        let (newproof, _) =
454          ProofEngineHelpers.subst_meta_in_proof
455            proof metano bo'[newmeta,context_for_newmeta,newmetaty]
456        in
457         (newproof, [newmeta])
458  in
459   PET.mk_tactic (letin_tac ~mk_fresh_name_callback term)
460
461   (** functional part of the "exact" tactic *)
462 let exact_tac ~term =
463  let exact_tac ~term (proof, goal) =
464   (* Assumption: the term bo must be closed in the current context *)
465   let (_,metasenv,_,_, _) = proof in
466   let metano,context,ty = CicUtil.lookup_meta goal metasenv in
467   let module T = CicTypeChecker in
468   let module R = CicReduction in
469   let ty_term,u = T.type_of_aux' metasenv context term CicUniv.empty_ugraph in
470   let b,_ = R.are_convertible context ty_term ty u in (* TASSI: FIXME *)
471   if b then
472    begin
473     let (newproof, metasenv') =
474       ProofEngineHelpers.subst_meta_in_proof proof metano term [] in
475     (newproof, [])
476    end
477   else
478    raise (PET.Fail (lazy "The type of the provided term is not the one expected."))
479  in
480   PET.mk_tactic (exact_tac ~term)
481
482 (* not really "primitive" tactics .... *)
483   
484 module TC  = CicTypeChecker
485 module UM  = UriManager
486 module R   = CicReduction
487 module C   = Cic
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 beta_after_elim_tac upto predicate =
496    let beta_after_elim_tac status =
497       let proof, goal = status in
498       let _, metasenv, _, _, _ = proof in
499       let _, _, ty = CicUtil.lookup_meta goal metasenv in
500       let mk_pattern ~equality ~upto ~predicate ty =
501          (* code adapted from ProceduralConversion.generalize *)
502          let meta = C.Implicit None in
503          let hole = C.Implicit (Some `Hole) in
504          let anon = C.Anonymous in
505          let is_meta =
506             let map b = function
507                | C.Implicit None when b -> b
508                | _                      -> false
509             in
510             List.fold_left map true
511          in
512          let rec gen_fix len k (name, i, ty, bo) =
513             name, i, gen_term k ty, gen_term (k + len) bo
514          and gen_cofix len k (name, ty, bo) =
515             name, gen_term k ty, gen_term (k + len) bo
516          and gen_term k = function
517             | C.Sort _ 
518             | C.Implicit _
519             | C.Const (_, _)
520             | C.Var (_, _)
521             | C.MutInd (_, _, _)
522             | C.MutConstruct (_, _, _, _)
523             | C.Meta (_, _) 
524             | C.Rel _ -> meta
525             | C.Appl (hd :: tl) when equality hd (S.lift k predicate) ->
526                assert (List.length tl = upto);
527                hole
528             | C.Appl ts -> 
529                let ts = List.map (gen_term k) ts in
530                if is_meta ts then meta else C.Appl ts
531             | C.Cast (te, ty) -> 
532                let te, ty = gen_term k te, gen_term k ty in
533                if is_meta [te; ty] then meta else C.Cast (te, ty)
534             | C.MutCase (sp, i, outty, t, pl) ->         
535                let outty, t, pl = gen_term k outty, gen_term k t, List.map (gen_term k) pl in
536                if is_meta (outty :: t :: pl) then meta else hole (* C.MutCase (sp, i, outty, t, pl) *)
537             | C.Prod (_, s, t) -> 
538                let s, t = gen_term k s, gen_term (succ k) t in
539                if is_meta [s; t] then meta else C.Prod (anon, s, t)
540             | C.Lambda (_, s, t) ->
541                let s, t = gen_term k s, gen_term (succ k) t in
542                if is_meta [s; t] then meta else C.Lambda (anon, s, t)
543             | C.LetIn (_, s, t) -> 
544                let s, t = gen_term k s, gen_term (succ k) t in
545                if is_meta [s; t] then meta else C.LetIn (anon, s, t)
546             | C.Fix (i, fl) -> C.Fix (i, List.map (gen_fix (List.length fl) k) fl)
547             | C.CoFix (i, fl) -> C.CoFix (i, List.map (gen_cofix (List.length fl) k) fl)
548          in
549          None, [], Some (gen_term 0 ty)
550       in
551       let equality = CicUtil.alpha_equivalence in
552       let pattern = mk_pattern ~equality ~upto ~predicate ty in
553       let tactic = RT.head_beta_reduce_tac ~delta:false ~upto ~pattern in
554       PET.apply_tactic tactic status
555    in
556    PET.mk_tactic beta_after_elim_tac
557    
558 let elim_tac ?using ?(pattern = PET.conclusion_pattern None) term = 
559  let elim_tac (proof, goal) =
560    let cpatt = match pattern with 
561       | None, [], Some cpatt -> cpatt
562       | _                    -> raise (PET.Fail (lazy "not implemented"))
563    in    
564    let ugraph = CicUniv.empty_ugraph in
565    let curi, metasenv, proofbo, proofty, attrs = proof in
566    let conjecture = CicUtil.lookup_meta goal metasenv in
567    let metano, context, ty = conjecture in 
568     let termty,_ugraph = TC.type_of_aux' metasenv context term ugraph in
569     let termty = CicReduction.whd context termty in
570     let termty, metasenv', arguments, _fresh_meta =
571      TermUtil.saturate_term
572       (ProofEngineHelpers.new_meta_of_proof proof) metasenv context termty 0 in
573     let term = if arguments = [] then term else Cic.Appl (term::arguments) in
574     let uri, exp_named_subst, typeno, _args =
575      match termty with
576         C.MutInd (uri,typeno,exp_named_subst) -> (uri,exp_named_subst,typeno,[])
577       | C.Appl ((C.MutInd (uri,typeno,exp_named_subst))::args) ->
578           (uri,exp_named_subst,typeno,args)
579       | _ -> raise NotAnInductiveTypeToEliminate
580     in
581      let eliminator_uri =
582       let buri = UM.buri_of_uri uri in
583       let name = 
584         let o,_ugraph = CicEnvironment.get_obj ugraph uri in
585        match o with
586           C.InductiveDefinition (tys,_,_,_) ->
587            let (name,_,_,_) = List.nth tys typeno in
588             name
589         | _ -> assert false
590       in
591       let ty_ty,_ugraph = TC.type_of_aux' metasenv' context ty ugraph in
592       let ext =
593        match ty_ty with
594           C.Sort C.Prop -> "_ind"
595         | C.Sort C.Set  -> "_rec"
596         | C.Sort C.CProp -> "_rec"
597         | C.Sort (C.Type _)-> "_rect" 
598         | C.Meta (_,_) -> raise TheTypeOfTheCurrentGoalIsAMetaICannotChooseTheRightElimiantionPrinciple
599         | _ -> assert false
600       in
601        UM.uri_of_string (buri ^ "/" ^ name ^ ext ^ ".con")
602      in
603       let eliminator_ref = match using with
604          | None   -> C.Const (eliminator_uri, exp_named_subst)
605          | Some t -> t 
606        in
607        let ety, _ugraph = 
608          TC.type_of_aux' metasenv' context eliminator_ref ugraph in
609 (* FG: ADDED PART ***********************************************************)
610 (* FG: we can not assume eliminator is the default eliminator ***************)
611    let rec args_init n f =
612       if n <= 0 then [] else f n :: args_init (pred n) f
613    in
614    let splits, args_no = PEH.split_with_whd (context, ety) in
615    let pred_pos = match List.hd splits with
616       | _, C.Rel i when i > 1 && i <= args_no -> i
617       | _, C.Appl (C.Rel i :: _) when i > 1 && i <= args_no -> i
618       | _ -> raise NotAnEliminator
619    in
620    let upto, metasenv', pred, term = match pattern with 
621       | None, [], Some (C.Implicit (Some `Hole)) ->
622          0, metasenv', C.Implicit None, term
623       | _                                        ->
624          let instantiated_eliminator =
625             let f n = if n = 1 then term else C.Implicit None in
626             C.Appl (eliminator_ref :: args_init args_no f)
627          in
628          let _actual_term, iety, _metasenv'', _ugraph = 
629             CicRefine.type_of_aux' metasenv' context instantiated_eliminator ugraph
630          in
631          let _actual_meta, actual_args = match iety with
632             | C.Meta (i, _)                  -> i, []
633             | C.Appl (C.Meta (i, _) :: args) -> i, args
634             | _                              -> assert false
635          in
636          (* let _, upto = PEH.split_with_whd (List.nth splits pred_pos) in *)
637          let upto = List.length actual_args in
638          let rec mk_pred metasenv context' pred term' = function
639             | []           -> metasenv, pred, term'
640             | term :: tail -> 
641 (* FG: we find the predicate for the eliminator as in the rewrite tactic ****)
642                let termty, _ugraph = TC.type_of_aux' metasenv context' term ugraph in
643                let termty = CicReduction.whd context' termty in         
644                let fresh_name = 
645                   FreshNamesGenerator.mk_fresh_name 
646                   ~subst:[] metasenv context' C.Anonymous ~typ:termty
647                in
648                let hyp = Some (fresh_name, C.Decl termty) in
649                let lazy_term c m u =  
650                   let distance  = List.length c - List.length context in
651                   S.lift distance term, m, u
652                in
653                let pattern = Some lazy_term, [], Some cpatt in
654                let subst, metasenv, _ugraph, _conjecture, selected_terms =
655                   ProofEngineHelpers.select
656                   ~metasenv ~ugraph ~conjecture:(metano, context, pred) ~pattern
657                in
658                let metasenv = MS.apply_subst_metasenv subst metasenv in  
659                let map (_context_of_t, t) l = t :: l in
660                let what = List.fold_right map selected_terms [] in
661                let term' = MS.apply_subst subst term' in
662                let termty = MS.apply_subst subst termty in
663                let pred = PER.replace_with_rel_1_from ~equality:(==) ~what 1 pred in
664                let pred = MS.apply_subst subst pred in
665                let pred = C.Lambda (fresh_name, termty, pred) in
666                mk_pred metasenv (hyp :: context') pred term' tail 
667          in
668          let metasenv', pred, term = mk_pred metasenv' context ty term actual_args in
669          HLog.debug ("PRED: " ^ CicPp.ppterm ~metasenv:metasenv' pred ^ " ARGS: " ^ String.concat " " (List.map (CicPp.ppterm ~metasenv:metasenv') actual_args));
670          upto, metasenv', pred, term
671       in
672 (* FG: END OF ADDED PART ****************************************************)
673       let term_to_refine =
674          let f n =
675             if n = pred_pos then pred else
676             if n = 1 then term else C.Implicit None
677          in
678          C.Appl (eliminator_ref :: args_init args_no f)
679       in
680       let refined_term,_refined_termty,metasenv'',_ugraph = 
681          CicRefine.type_of_aux' metasenv' context term_to_refine ugraph
682       in
683       let new_goals =
684          ProofEngineHelpers.compare_metasenvs
685             ~oldmetasenv:metasenv ~newmetasenv:metasenv''
686       in
687       let proof' = curi,metasenv'',proofbo,proofty, attrs in
688       let proof'', new_goals' =
689          PET.apply_tactic (apply_tac ~term:refined_term) (proof',goal)
690       in
691       (* The apply_tactic can have closed some of the new_goals *)
692       let patched_new_goals =
693          let (_,metasenv''',_,_, _) = proof'' in
694          List.filter
695             (function i -> List.exists (function (j,_,_) -> j=i) metasenv''')
696             new_goals @ new_goals'
697       in
698       let res = proof'', patched_new_goals in
699       if upto = 0 then res else 
700       let continuation = beta_after_elim_tac upto pred in
701       let dummy_status = proof,goal in
702       PET.apply_tactic
703          (T.then_ ~start:(PET.mk_tactic (fun _ -> res)) ~continuation)
704          dummy_status
705    in
706    PET.mk_tactic elim_tac
707 ;;
708
709 let cases_intros_tac ?(mk_fresh_name_callback = FreshNamesGenerator.mk_fresh_name ~subst:[]) term =
710  let cases_tac ~term (proof, goal) =
711   let module TC = CicTypeChecker in
712   let module U = UriManager in
713   let module R = CicReduction in
714   let module C = Cic in
715   let (curi,metasenv,proofbo,proofty, attrs) = proof in
716   let metano,context,ty = CicUtil.lookup_meta goal metasenv in
717   let termty,_ = TC.type_of_aux' metasenv context term CicUniv.empty_ugraph in
718   let termty = CicReduction.whd context termty in
719   let (termty,metasenv',arguments,fresh_meta) =
720    TermUtil.saturate_term
721     (ProofEngineHelpers.new_meta_of_proof proof) metasenv context termty 0 in
722   let term = if arguments = [] then term else Cic.Appl (term::arguments) in
723   let uri,exp_named_subst,typeno,args =
724     match termty with
725     | C.MutInd (uri,typeno,exp_named_subst) -> (uri,exp_named_subst,typeno,[])
726     | C.Appl ((C.MutInd (uri,typeno,exp_named_subst))::args) ->
727         (uri,exp_named_subst,typeno,args)
728     | _ -> raise NotAnInductiveTypeToEliminate
729   in
730   let paramsno,itty,patterns,right_args =
731     match CicEnvironment.get_obj CicUniv.empty_ugraph uri with
732     | C.InductiveDefinition (tys,_,paramsno,_),_ ->
733        let _,left_parameters,right_args = 
734          List.fold_right 
735            (fun x (n,acc1,acc2) -> 
736              if n > 0 then (n-1,acc1,x::acc2) else (n,x::acc1,acc2)) 
737            args (List.length args - paramsno, [],[])
738        in
739        let _,_,itty,cl = List.nth tys typeno in
740        let rec aux left_parameters context t =
741          match left_parameters,CicReduction.whd context t with
742          | [],C.Prod (name,source,target) ->
743             let fresh_name =
744               mk_fresh_name_callback metasenv' context name ~typ:source
745             in
746              C.Lambda (fresh_name,C.Implicit None,
747              aux [] (Some (fresh_name,C.Decl source)::context) target)
748          | hd::tl,C.Prod (name,source,target) ->
749              (* left parameters instantiation *)
750              aux tl context (CicSubstitution.subst hd target)
751          | [],_ -> C.Implicit None
752          | _ -> assert false
753        in
754         paramsno,itty,
755         List.map (function (_,cty) -> aux left_parameters context cty) cl,
756         right_args
757     | _ -> assert false
758   in
759   let outtype =
760     let n_right_args = List.length right_args in
761     let n_lambdas = n_right_args + 1 in
762     let lifted_ty = CicSubstitution.lift n_lambdas ty in
763     let captured_ty = 
764       let what = 
765         List.map (CicSubstitution.lift n_lambdas) (right_args)
766       in
767       let with_what meta = 
768         let rec mkargs = function 
769           | 0 -> assert false
770           | 1 -> []
771           | n -> 
772               (if meta then Cic.Implicit None else Cic.Rel n)::(mkargs (n-1)) 
773         in
774         mkargs n_lambdas 
775       in
776       let replaced = ref false in
777       let replace = ProofEngineReduction.replace_lifting
778        ~equality:(fun _ a b -> let rc = CicUtil.alpha_equivalence a b in 
779                   if rc then replaced := true; rc)
780        ~context:[]
781       in
782       let captured = 
783         replace ~what:[CicSubstitution.lift n_lambdas term] 
784           ~with_what:[Cic.Rel 1] ~where:lifted_ty
785       in
786       if not !replaced then
787         (* this means the matched term is not there, 
788          * but maybe right params are: we user rels (to right args lambdas) *)
789         replace ~what ~with_what:(with_what false) ~where:captured
790       else
791         (* since the matched is there, rights should be inferrable *)
792         replace ~what ~with_what:(with_what true) ~where:captured
793     in
794     let captured_term_ty = 
795       let term_ty = CicSubstitution.lift n_right_args termty in
796       let rec mkrels = function 0 -> []|n -> (Cic.Rel n)::(mkrels (n-1)) in
797       let rec fstn acc l n = 
798         if n = 0 then acc else fstn (acc@[List.hd l]) (List.tl l) (n-1) 
799       in
800       match term_ty with
801       | C.MutInd _ -> term_ty
802       | C.Appl ((C.MutInd (a,b,c))::args) -> 
803            C.Appl ((C.MutInd (a,b,c))::
804                fstn [] args paramsno @ mkrels n_right_args)
805       | _ -> raise NotAnInductiveTypeToEliminate
806     in
807     let rec add_lambdas = function
808       | 0 -> captured_ty
809       | 1 -> 
810           C.Lambda (C.Name "matched", captured_term_ty, (add_lambdas 0))
811       | n -> 
812            C.Lambda (C.Name ("right_"^(string_of_int (n-1))),
813                      C.Implicit None, (add_lambdas (n-1)))
814     in
815     add_lambdas n_lambdas
816   in
817   let term_to_refine = C.MutCase (uri,typeno,outtype,term,patterns) in
818   let refined_term,_,metasenv'',_ = 
819     CicRefine.type_of_aux' metasenv' context term_to_refine
820       CicUniv.empty_ugraph
821   in
822   let new_goals =
823     ProofEngineHelpers.compare_metasenvs
824       ~oldmetasenv:metasenv ~newmetasenv:metasenv''
825   in
826   let proof' = curi,metasenv'',proofbo,proofty, attrs in
827   let proof'', new_goals' =
828     PET.apply_tactic (apply_tac ~term:refined_term) (proof',goal)
829   in
830   (* The apply_tactic can have closed some of the new_goals *)
831   let patched_new_goals =
832     let (_,metasenv''',_,_,_) = proof'' in
833       List.filter
834         (function i -> List.exists (function (j,_,_) -> j=i) metasenv''')
835         new_goals @ new_goals'
836     in
837     proof'', patched_new_goals
838   in
839   PET.mk_tactic (cases_tac ~term)
840 ;;
841
842
843 let elim_intros_tac ?(mk_fresh_name_callback = FreshNamesGenerator.mk_fresh_name ~subst:[]) 
844                     ?depth ?using ?pattern what =
845  Tacticals.then_ ~start:(elim_tac ?using ?pattern what)
846   ~continuation:(intros_tac ~mk_fresh_name_callback ?howmany:depth ())
847 ;;
848
849 (* The simplification is performed only on the conclusion *)
850 let elim_intros_simpl_tac ?(mk_fresh_name_callback = FreshNamesGenerator.mk_fresh_name ~subst:[])
851                           ?depth ?using ?pattern what =
852  Tacticals.then_ ~start:(elim_tac ?using ?pattern what)
853   ~continuation:
854    (Tacticals.thens
855      ~start:(intros_tac ~mk_fresh_name_callback ?howmany:depth ())
856      ~continuations:
857        [ReductionTactics.simpl_tac
858          ~pattern:(ProofEngineTypes.conclusion_pattern None)])
859 ;;
860
861 (* FG: insetrts a "hole" in the context (derived from letin_tac) *)
862
863 let letout_tac =
864    let mk_fresh_name_callback = FreshNamesGenerator.mk_fresh_name ~subst:[] in
865    let term = C.Sort C.Set in
866    let letout_tac (proof, goal) =
867       let curi, metasenv, pbo, pty, attrs = proof in
868       let metano, context, ty = CicUtil.lookup_meta goal metasenv in
869       let newmeta = ProofEngineHelpers.new_meta_of_proof ~proof in
870       let fresh_name = mk_fresh_name_callback metasenv context (Cic.Name "hole") ~typ:term in
871       let context_for_newmeta = None :: context in
872       let irl = CicMkImplicit.identity_relocation_list_for_metavariable context_for_newmeta in
873       let newmetaty = CicSubstitution.lift 1 ty in
874       let bo' = C.LetIn (fresh_name, term, C.Meta (newmeta,irl)) in
875       let newproof, _ = ProofEngineHelpers.subst_meta_in_proof proof metano bo'[newmeta,context_for_newmeta,newmetaty] in
876       newproof, [newmeta]
877    in
878    PET.mk_tactic letout_tac