]> matita.cs.unibo.it Git - helm.git/blob - matita/components/ng_refiner/nCicRefiner.ml
In the case type_of constructor with expected type T, T is now put in whd to
[helm.git] / matita / components / ng_refiner / nCicRefiner.ml
1 (*
2     ||M||  This file is part of HELM, an Hypertextual, Electronic        
3     ||A||  Library of Mathematics, developed at the Computer Science     
4     ||T||  Department, University of Bologna, Italy.                     
5     ||I||                                                                
6     ||T||  HELM is free software; you can redistribute it and/or         
7     ||A||  modify it under the terms of the GNU General Public License   
8     \   /  version 2 or (at your option) any later version.      
9      \ /   This software is distributed as is, NO WARRANTY.     
10       V_______________________________________________________________ *)
11
12 (* $Id$ *)
13
14 exception RefineFailure of (Stdpp.location * string) Lazy.t;;
15 exception Uncertain of (Stdpp.location * string) Lazy.t;;
16 exception AssertFailure of string Lazy.t;;
17
18 module C = NCic
19 module Ref = NReference
20
21 let debug = ref false;;
22 let indent = ref "";;
23 let times = ref [];;
24 let pp s =
25  if !debug then
26   prerr_endline (Printf.sprintf "%-20s" !indent ^ " " ^ Lazy.force s)
27 ;;
28 let inside c =
29  if !debug then
30   begin
31    let time1 = Unix.gettimeofday () in
32    indent := !indent ^ String.make 1 c;
33    times := time1 :: !times;
34    prerr_endline ("{{{" ^ !indent ^ " ")
35   end
36 ;;
37 let outside ok =
38  if !debug then
39   begin
40    let time2 = Unix.gettimeofday () in
41    let time1 =
42     match !times with time1::tl -> times := tl; time1 | [] -> assert false in
43    prerr_endline ("}}} " ^ string_of_float (time2 -. time1));
44    if not ok then prerr_endline "exception raised!";
45    try
46     indent := String.sub !indent 0 (String.length !indent -1)
47    with
48     Invalid_argument _ -> indent := "??"; ()
49  end
50 ;;
51
52
53 let wrap_exc msg = function
54   | NCicUnification.Uncertain _ -> Uncertain msg
55   | NCicUnification.UnificationFailure _ -> RefineFailure msg
56   | NCicTypeChecker.TypeCheckerFailure _ -> RefineFailure msg
57   | e -> raise e
58 ;;
59
60 let exp_implicit rdb ~localise metasenv subst context with_type t =
61  function      
62   | `Closed ->
63       let metasenv,subst,expty =
64        match with_type with
65           None -> metasenv,subst,None
66         | Some typ ->
67            let (metasenv,subst),typ =
68             try
69              NCicMetaSubst.delift
70               ~unify:(fun m s c t1 t2 ->
71                 try Some (NCicUnification.unify rdb m s c t1 t2)
72                 with NCicUnification.UnificationFailure _ | NCicUnification.Uncertain _ -> None)
73               metasenv subst context 0 (0,NCic.Irl 0) typ
74             with
75               NCicMetaSubst.MetaSubstFailure _
76             | NCicMetaSubst.Uncertain _ ->
77                raise (RefineFailure (lazy (localise t,"Trying to create a closed meta with a non closed type: " ^ NCicPp.ppterm ~metasenv ~subst ~context typ)))
78
79            in
80             metasenv,subst,Some typ
81       in
82        NCicMetaSubst.mk_meta metasenv [] ?with_type:expty `IsTerm,subst
83   | `Type -> NCicMetaSubst.mk_meta metasenv context ?with_type `IsType,subst
84   | `Term -> NCicMetaSubst.mk_meta metasenv context ?with_type `IsTerm,subst
85   | `Tagged s ->
86       NCicMetaSubst.mk_meta 
87         ~attrs:[`Name s] metasenv context ?with_type `IsTerm,subst
88   | `Vector ->
89       raise (RefineFailure (lazy (localise t, "A vector of implicit terms " ^
90        "can only be used in argument position")))
91   | _ -> assert false
92 ;;
93
94 let check_allowed_sort_elimination rdb localise r orig =
95    let mkapp he arg =
96      match he with
97      | C.Appl l -> C.Appl (l @ [arg])
98      | t -> C.Appl [t;arg] in
99    (* ctx, ind_type @ lefts, sort_of_ind_ty@lefts, outsort *)
100    let rec aux metasenv subst context ind arity1 arity2 =
101      (*D*)inside 'S'; try let rc = 
102      let arity1 = NCicReduction.whd ~subst context arity1 in
103      pp (lazy(NCicPp.ppterm ~subst ~metasenv ~context arity1 ^ "   elimsto   " ^
104         NCicPp.ppterm ~subst ~metasenv ~context arity2 ^ "\nMENV:\n"^
105         NCicPp.ppmetasenv ~subst metasenv));
106      match arity1 with
107      | C.Prod (name,so1,de1) (* , t ==?== C.Prod _ *) ->
108         let metasenv, _, meta, _ = 
109           NCicMetaSubst.mk_meta metasenv ((name,C.Decl so1)::context) `IsType 
110         in
111         let metasenv, subst = 
112           try NCicUnification.unify rdb metasenv subst context 
113                 arity2 (C.Prod (name, so1, meta)) 
114           with exc -> raise (wrap_exc (lazy (localise orig, Printf.sprintf
115             "expected %s, found %s" (* XXX localizzare meglio *)
116             (NCicPp.ppterm ~subst ~metasenv ~context (C.Prod (name, so1, meta)))
117             (NCicPp.ppterm ~subst ~metasenv ~context arity2))) exc)
118         in
119         aux metasenv subst ((name, C.Decl so1)::context)
120          (mkapp (NCicSubstitution.lift 1 ind) (C.Rel 1)) de1 meta
121      | C.Sort _ (* , t ==?== C.Prod _ *) ->
122         let metasenv, _, meta, _ = NCicMetaSubst.mk_meta metasenv [] `IsSort in
123         let metasenv, subst = 
124           try NCicUnification.unify rdb metasenv subst context 
125                 arity2 (C.Prod ("_", ind, meta)) 
126           with exc -> raise (wrap_exc (lazy (localise orig, Printf.sprintf
127             "expected %s, found %s" (* XXX localizzare meglio *)
128             (NCicPp.ppterm ~subst ~metasenv ~context (C.Prod ("_", ind, meta)))
129             (NCicPp.ppterm ~subst ~metasenv ~context arity2))) exc)
130         in
131         (try NCicTypeChecker.check_allowed_sort_elimination
132             ~metasenv ~subst r context ind arity1 arity2; 
133             metasenv, subst
134         with exc -> raise (wrap_exc (lazy (localise orig, 
135           "Sort elimination not allowed ")) exc))
136      | _ -> assert false
137      (*D*)in outside true; rc with exc -> outside false; raise exc
138    in
139     aux
140 ;;
141
142 let rec typeof rdb 
143   ?(localise=fun _ -> Stdpp.dummy_loc) 
144   metasenv subst context term expty 
145 =
146   let force_ty skip_lambda skip_appl metasenv subst context orig t infty expty =
147     (*D*)inside 'F'; try let rc = 
148     match expty with
149     | Some expty ->
150        (match t with
151        | C.Implicit _ -> assert false
152        | C.Lambda _ when skip_lambda -> metasenv, subst, t, expty
153        | C.Appl _ when skip_appl -> metasenv, subst, t, expty
154        | _ ->
155           pp (lazy ("forcing infty=expty: "^
156           (NCicPp.ppterm ~metasenv ~subst ~context infty) ^  " === " ^
157           (NCicPp.ppterm ~metasenv ~subst:[] ~context expty)));
158            try 
159              let metasenv, subst =
160     (*D*)inside 'U'; try let rc = 
161                NCicUnification.unify rdb metasenv subst context infty expty 
162     (*D*)in outside true; rc with exc -> outside false; raise exc
163              in
164              metasenv, subst, t, expty
165            with 
166            | NCicUnification.Uncertain _ 
167            | NCicUnification.UnificationFailure _ as exc -> 
168              try_coercions rdb ~localise 
169                metasenv subst context t orig infty expty true exc)
170     | None -> metasenv, subst, t, infty
171     (*D*)in outside true; rc with exc -> outside false; raise exc
172   in
173   let rec typeof_aux metasenv subst context expty = 
174     fun t as orig -> 
175     (*D*)inside 'R'; try let rc = 
176     pp (lazy (NCicPp.ppterm ~metasenv ~subst ~context t ^ " ::exp:: " ^
177        match expty with None -> "None" | Some e -> 
178        NCicPp.ppterm ~metasenv ~subst ~context e));
179     let metasenv, subst, t, infty = 
180     match t with
181     | C.Rel n ->
182         let infty = 
183          (try
184            match List.nth context (n - 1) with
185            | (_,C.Decl ty) -> NCicSubstitution.lift n ty
186            | (_,C.Def (_,ty)) -> NCicSubstitution.lift n ty
187          with Failure _ -> 
188            raise (RefineFailure (lazy (localise t,"unbound variable"))))
189         in
190         metasenv, subst, t, infty
191     | C.Sort s -> 
192          (try metasenv, subst, t, C.Sort (NCicEnvironment.typeof_sort s)
193          with 
194          | NCicEnvironment.UntypableSort msg -> 
195               raise (RefineFailure (lazy (localise t, Lazy.force msg)))
196          | NCicEnvironment.AssertFailure msg -> raise (AssertFailure msg))
197     | C.Implicit infos -> 
198          let (metasenv,_,t,ty),subst =
199            exp_implicit rdb ~localise metasenv subst context expty t infos
200          in
201          if expty = None then
202           typeof_aux metasenv subst context None t
203          else
204           metasenv, subst, t, ty 
205     | C.Meta (n,l) as t -> 
206        let metasenv, ty =
207         try
208          let _,_,_,ty = NCicUtils.lookup_subst n subst in metasenv, ty 
209         with NCicUtils.Subst_not_found _ ->
210          NCicMetaSubst.extend_meta metasenv n
211        in
212        metasenv, subst, t, NCicSubstitution.subst_meta l ty
213     | C.Const _ -> 
214        metasenv, subst, t, NCicTypeChecker.typeof ~subst ~metasenv context t
215     | C.Prod (name,(s as orig_s),(t as orig_t)) ->
216        let metasenv, subst, s, s1 = typeof_aux metasenv subst context None s in
217        let metasenv, subst, s, s1 = 
218          force_to_sort rdb 
219            metasenv subst context s orig_s localise s1 in
220        let context1 = (name,(C.Decl s))::context in
221        let metasenv, subst, t, s2 = typeof_aux metasenv subst context1 None t in
222        let metasenv, subst, t, s2 = 
223          force_to_sort rdb 
224            metasenv subst context1 t orig_t localise s2 in
225        let metasenv, subst, s, t, ty = 
226          sort_of_prod localise metasenv subst 
227            context orig_s orig_t (name,s) t (s1,s2)
228        in
229        metasenv, subst, NCic.Prod(name,s,t), ty
230     | C.Lambda (n,(s as orig_s),t) as orig ->
231        let exp_s, exp_ty_t, force_after =
232          match expty with
233          | None -> None, None, false
234          | Some expty -> 
235              match NCicReduction.whd ~subst context expty with
236              | C.Prod (_,s,t) -> Some s, Some t, false
237              | _ -> None, None, true 
238        in
239        let (metasenv,subst), s = 
240          match exp_s with 
241          | Some exp_s ->
242              (* optimized case: implicit and implicitly typed meta
243               * the optimization prevents proliferation of metas  *)
244              (match s with
245               | C.Implicit _ -> (metasenv,subst),exp_s
246               | _ ->
247                   let metasenv, subst, s = match s with 
248                     | C.Meta (n,_) 
249                         when (try (match NCicUtils.lookup_meta n metasenv with
250                               | _,_,C.Implicit (`Typeof _) -> true
251                               | _ -> false)
252                         with 
253                           | _ -> false) -> metasenv, subst, s
254                     | _ ->  check_type rdb ~localise metasenv subst context s in
255                   (try 
256                     pp(lazy("Force source to: "^NCicPp.ppterm ~metasenv ~subst
257                        ~context exp_s));
258                     NCicUnification.unify ~test_eq_only:true rdb metasenv subst context s exp_s,s
259                   with exc -> raise (wrap_exc (lazy (localise orig_s, Printf.sprintf
260                     "Source type %s was expected to be %s" (NCicPp.ppterm ~metasenv
261                     ~subst ~context s) (NCicPp.ppterm ~metasenv ~subst ~context
262                     exp_s))) exc)))
263          | None -> 
264              let metasenv, subst, s = 
265                check_type rdb ~localise metasenv subst context s in
266              (metasenv, subst), s
267        in
268        let context_for_t = (n,C.Decl s) :: context in
269        let metasenv, subst, t, ty_t = 
270          typeof_aux metasenv subst context_for_t exp_ty_t t 
271        in
272        if force_after then
273          force_ty false true metasenv subst context orig 
274            (C.Lambda(n,s,t)) (C.Prod (n,s,ty_t)) expty
275        else 
276          metasenv, subst, C.Lambda(n,s,t), C.Prod (n,s,ty_t)
277     | C.LetIn (n,ty,t,bo) ->
278        let metasenv, subst, ty = 
279          check_type rdb ~localise metasenv subst context ty in
280        let metasenv, subst, t, _ = 
281          typeof_aux metasenv subst context (Some ty) t in
282        let context1 = (n, C.Def (t,ty)) :: context in
283        let metasenv, subst, expty1 = 
284          match expty with 
285          | None -> metasenv, subst, None 
286          | Some x -> 
287              let m, s, x = 
288                NCicUnification.delift_type_wrt_terms 
289                 rdb metasenv subst context1 (NCicSubstitution.lift 1 x)
290                 [NCicSubstitution.lift 1 t]
291              in
292                m, s, Some x
293        in
294        let metasenv, subst, bo, bo_ty = 
295          typeof_aux metasenv subst context1 expty1 bo 
296        in
297        let bo_ty = NCicSubstitution.subst ~avoid_beta_redexes:true t bo_ty in
298        metasenv, subst, C.LetIn (n, ty, t, bo), bo_ty
299     | C.Appl ((he as orig_he)::(_::_ as args)) ->
300        let upto = match orig_he with C.Meta _ -> List.length args | _ -> 0 in
301        let hbr t = 
302          if upto > 0 then NCicReduction.head_beta_reduce ~upto t else t 
303        in
304        let refine_appl metasenv subst args =
305          let metasenv, subst, he, ty_he = 
306             typeof_aux metasenv subst context None he in
307          let metasenv, subst, t, ty = 
308            eat_prods rdb ~localise force_ty metasenv subst context expty t
309             orig_he he ty_he args in
310          metasenv, subst, hbr t, ty
311        in
312        if args = [C.Implicit `Vector] && expty <> None then
313          (* we try here to expand the vector a 0 implicits, but we use
314           * the expected type *)
315          try
316            let metasenv, subst, he, ty_he = 
317               typeof_aux metasenv subst context expty he in
318            metasenv, subst, hbr he, ty_he
319          with Uncertain _ | RefineFailure _ -> refine_appl metasenv subst args
320        else
321         (* CSC: whd can be useful on he and expty ... *)
322         (match he with
323             C.Const (Ref.Ref (uri1,Ref.Con _)) -> (
324              match
325               HExtlib.map_option (NCicReduction.whd ~subst context) expty
326              with
327                 Some (C.Appl(C.Const(Ref.Ref (uri2,Ref.Ind _) as ref)::expargs))
328                 when NUri.eq uri1 uri2 ->
329                 (try
330                  let _,leftno,_,_,_ = NCicEnvironment.get_checked_indtys ref in
331                  let leftargs,rightargs = HExtlib.split_nth leftno args in
332                  let leftexpargs,_ = HExtlib.split_nth leftno expargs in
333                  let metasenv,subst,leftargs_rev =
334                   List.fold_left2
335                    (fun (metasenv,subst,args) arg exparg ->
336                      let metasenv,subst,arg,_ =
337                       typeof_aux metasenv subst context None arg in
338                      let metasenv,subst =
339                       NCicUnification.unify rdb metasenv subst context arg exparg
340                      in
341                       metasenv,subst,arg::args
342                    ) (metasenv,subst,[]) leftargs leftexpargs
343                  in
344                   (* some checks are re-done here, but it would be complex to
345                    * avoid them (e.g. we did not check yet that the constructor
346                    * is a construtor for that inductive type) *)
347                   refine_appl metasenv subst ((List.rev leftargs_rev)@rightargs)
348                 with
349                  | Sys.Break as exn -> raise exn
350                  | _ -> refine_appl metasenv subst args (* to try coercions *))
351               | _ -> refine_appl metasenv subst args)
352           | _ -> refine_appl metasenv subst args)
353    | C.Appl _ -> raise (AssertFailure (lazy "Appl of length < 2"))
354    | C.Match (Ref.Ref (_,Ref.Ind (_,tyno,_)) as r,
355           outtype,(term as orig_term),pl) as orig ->
356       let _, leftno, itl, _, _ = NCicEnvironment.get_checked_indtys r in
357       let _, _, arity, cl = List.nth itl tyno in
358       let constructorsno = List.length cl in
359       let _, metasenv, args = 
360         NCicMetaSubst.saturate metasenv subst context arity 0 in
361       let ind = if args = [] then C.Const r else C.Appl (C.Const r::args) in
362       let metasenv, subst, term, _ = 
363         typeof_aux metasenv subst context (Some ind) term in
364       let parameters, arguments = HExtlib.split_nth leftno args in
365       let outtype =  
366         match outtype with
367         | C.Implicit _ as ot -> 
368              let rec aux = function
369                | [] -> NCic.Lambda ("_",NCic.Implicit `Type,ot)
370                | _::tl -> NCic.Lambda ("_",NCic.Implicit `Type,aux tl)
371              in
372                aux arguments
373         | _ -> outtype
374       in 
375       let metasenv, subst, outtype, outsort = 
376         typeof_aux metasenv subst context None outtype in
377
378       (* next lines are to do a subst-expansion of the outtype, so
379          that when it becomes a beta-abstraction, the beta-redex is
380          fired during substitution *)
381       let _,fresh_metanoouttype,newouttype,_ =
382        NCicMetaSubst.mk_meta metasenv context `IsTerm in
383       let subst =
384        (fresh_metanoouttype,([`Name "outtype"],context,outtype,outsort))
385          ::subst in
386       let outtype = newouttype in
387
388       (* let's control if the sort elimination is allowed: [(I q1 ... qr)|B] *)
389       let ind =
390         if parameters = [] then C.Const r
391         else C.Appl ((C.Const r)::parameters) in
392       let metasenv, subst, ind, ind_ty = 
393         typeof_aux metasenv subst context None ind in
394       let metasenv, subst = 
395          check_allowed_sort_elimination rdb localise r orig_term metasenv subst 
396            context ind ind_ty outsort 
397       in
398       (* let's check if the type of branches are right *)
399       if List.length pl <> constructorsno then
400        raise (RefineFailure (lazy (localise orig, 
401          "Wrong number of cases in a match")));
402 (*
403       let metasenv, subst =
404         match expty with
405         | None -> metasenv, subst
406         | Some expty -> 
407            NCicUnification.unify rdb metasenv subst context resty expty 
408       in
409 *)
410       let _, metasenv, subst, pl =
411         List.fold_right
412           (fun p (j, metasenv, subst, branches) ->
413               let cons = 
414                 let cons = Ref.mk_constructor j r in
415                 if parameters = [] then C.Const cons
416                 else C.Appl (C.Const cons::parameters)
417               in
418               let metasenv, subst, cons, ty_cons = 
419                 typeof_aux metasenv subst context None cons in
420               let ty_branch = 
421                 NCicTypeChecker.type_of_branch 
422                   ~subst context leftno outtype cons ty_cons in
423               pp (lazy ("TYPEOFBRANCH: " ^
424                NCicPp.ppterm ~metasenv ~subst ~context p ^ " ::inf:: " ^
425                NCicPp.ppterm ~metasenv ~subst ~context ty_branch ));
426               let metasenv, subst, p, _ = 
427                 typeof_aux metasenv subst context (Some ty_branch) p in
428               j-1, metasenv, subst, p :: branches)
429           pl (List.length pl, metasenv, subst, [])
430       in
431       let resty = C.Appl (outtype::arguments@[term]) in
432       let resty = NCicReduction.head_beta_reduce ~subst resty in
433       metasenv, subst, C.Match (r, outtype, term, pl),resty
434     | C.Match _ -> assert false
435     in
436     pp (lazy (NCicPp.ppterm ~metasenv ~subst ~context t ^ " ::inf:: "^
437          NCicPp.ppterm ~metasenv ~subst ~context infty ));
438       force_ty true true metasenv subst context orig t infty expty
439     (*D*)in outside true; rc with exc -> outside false; raise exc
440   in
441     typeof_aux metasenv subst context expty term
442
443 and check_type rdb ~localise metasenv subst context (ty as orig_ty) =
444   let metasenv, subst, ty, sort = 
445     typeof rdb ~localise metasenv subst context ty None
446   in
447   let metasenv, subst, ty, _ = 
448     force_to_sort rdb metasenv subst context ty orig_ty localise sort
449   in
450     metasenv, subst, ty
451
452 and try_coercions rdb 
453   ~localise 
454   metasenv subst context t orig_t infty expty perform_unification exc 
455 =
456   let cs_subst = NCicSubstitution.subst ~avoid_beta_redexes:true in
457   try 
458     pp (lazy "WWW: trying coercions -- preliminary unification");
459     let metasenv, subst = 
460       NCicUnification.unify rdb metasenv subst context infty expty
461     in metasenv, subst, t, expty
462   with
463   | exn ->
464   (* we try with a coercion *)
465   let rec first exc = function
466   | [] ->   
467       pp (lazy "WWW: no more coercions!");      
468       raise (wrap_exc (lazy (localise orig_t, Printf.sprintf
469         "The term\n%s\nhas type\n%s\nbut is here used with type\n%s"
470         (NCicPp.ppterm ~metasenv ~subst ~context t)
471         (NCicPp.ppterm ~metasenv ~subst ~context infty)
472         (NCicPp.ppterm ~metasenv ~subst ~context expty))) exc)
473   | (_,metasenv, newterm, newtype, meta)::tl ->
474       try 
475           pp (lazy("K=" ^ NCicPp.ppterm ~metasenv ~subst ~context newterm));
476           pp (lazy ( "UNIFICATION in CTX:\n"^ 
477             NCicPp.ppcontext ~metasenv ~subst context
478             ^ "\nMENV: " ^
479             NCicPp.ppmetasenv metasenv ~subst
480             ^ "\nOF: " ^
481             NCicPp.ppterm ~metasenv ~subst ~context t ^  " === " ^
482             NCicPp.ppterm ~metasenv ~subst ~context meta ^ "\n"));
483         let metasenv, subst = 
484           NCicUnification.unify rdb metasenv subst context t meta
485         in
486           pp (lazy ( "UNIFICATION in CTX:\n"^ 
487             NCicPp.ppcontext ~metasenv ~subst context
488             ^ "\nMENV: " ^
489             NCicPp.ppmetasenv metasenv ~subst
490             ^ "\nOF: " ^
491             NCicPp.ppterm ~metasenv ~subst ~context newtype ^  " === " ^
492             NCicPp.ppterm ~metasenv ~subst ~context expty ^ "\n"));
493         let metasenv, subst = 
494           if perform_unification then
495             NCicUnification.unify rdb metasenv subst context newtype expty
496           else metasenv, subst
497         in
498         metasenv, subst, newterm, newtype
499       with 
500       | NCicUnification.UnificationFailure _ -> first exc tl
501       | NCicUnification.Uncertain _ as exc -> first exc tl
502   in
503   
504   try 
505     pp (lazy "WWW: trying coercions -- inner check");
506     match infty, expty, t with
507     | _,_, NCic.Match (Ref.Ref (_,Ref.Ind (_,tyno,leftno)) as r,outty,m,pl) ->
508         (*{{{*) pp (lazy "CASE");
509         (* {{{ helper functions *)
510         let get_cl_and_left_p refit outty =
511           match refit with
512           | NReference.Ref (uri, NReference.Ind (_,tyno,leftno)) ->
513              let _, leftno, itl, _, _ = NCicEnvironment.get_checked_indtys r in
514              let _, _, ty, cl = List.nth itl tyno in
515              let constructorsno = List.length cl in
516               let count_pis t =
517                 let rec aux ctx t = 
518                   match NCicReduction.whd ~subst ~delta:max_int ctx t with
519                   | NCic.Prod (name,src,tgt) ->
520                       let ctx = (name, (NCic.Decl src)) :: ctx in
521                       1 + aux ctx tgt
522                   | _ -> 0
523                 in
524                   aux [] t
525               in
526               let rec skip_lambda_delifting t n =
527                 match t,n with
528                 | _,0 -> t
529                 | NCic.Lambda (_,_,t),n ->
530                     pp (lazy ("WWW: failing term? "^ NCicPp.ppterm ~subst:[] ~metasenv:[] ~context:[] t)); 
531                     skip_lambda_delifting
532                       (* substitute dangling indices with a dummy *)
533                       (NCicSubstitution.subst (NCic.Sort NCic.Prop) t) (n - 1)
534                 | _ -> assert false
535               in
536               let get_l_r_p n = function
537                 | NCic.Lambda (_,NCic.Const (Ref.Ref (_,Ref.Ind (_,_,_))),_) -> [],[]
538                 | NCic.Lambda (_,NCic.Appl 
539                     (NCic.Const (Ref.Ref (_,Ref.Ind (_,_,_))) :: args),_) ->
540                     HExtlib.split_nth n args
541                 | _ -> assert false
542               in
543               let pis = count_pis ty in
544               let rno = pis - leftno in
545               let t = skip_lambda_delifting outty rno in
546               let left_p, _ = get_l_r_p leftno t in
547               let instantiate_with_left cl =
548                 List.map 
549                   (fun ty -> 
550                     List.fold_left 
551                       (fun t p -> match t with
552                         | NCic.Prod (_,_,t) ->
553                             cs_subst p t
554                         | _-> assert false)
555                       ty left_p) 
556                   cl 
557               in
558               let cl = instantiate_with_left (List.map (fun (_,_,x) -> x) cl) in
559               cl, left_p, leftno, rno
560           | _ -> raise exn
561         in
562         let rec keep_lambdas_and_put_expty ctx t bo right_p matched n =
563           match t,n with
564           | _,0 ->
565             let rec mkr n = function 
566               | [] -> [] | _::tl -> NCic.Rel n :: mkr (n+1) tl
567             in
568             pp (lazy ("input replace: " ^ NCicPp.ppterm ~context:ctx ~metasenv:[] ~subst:[] bo));
569             let bo =
570             NCicRefineUtil.replace_lifting
571               ~equality:(fun _ -> NCicRefineUtil.alpha_equivalence)
572               ~context:ctx
573               ~what:(matched::right_p)
574               ~with_what:(NCic.Rel 1::List.rev (mkr 2 right_p))
575               ~where:bo
576             in
577             pp (lazy ("output replace: " ^ NCicPp.ppterm ~context:ctx ~metasenv:[] ~subst:[] bo));
578             bo
579           | NCic.Lambda (name, src, tgt),_ ->
580               NCic.Lambda (name, src,
581                keep_lambdas_and_put_expty 
582                 ((name, NCic.Decl src)::ctx) tgt (NCicSubstitution.lift 1 bo)
583                 (List.map (NCicSubstitution.lift 1) right_p) (NCicSubstitution.lift 1 matched) (n-1))
584           | _ -> assert false
585         in
586         (*let eq = NCic.Const (NReference.reference_of_string ("cic:/matita/ng/Plogic/equality/eq.ind(1,0,2)")) in
587         let eq_refl = NCic.Const (NReference.reference_of_string ("cic:/matita/ng/Plogic/equality/eq.con(0,1,2)")) in*)
588         let eq = NCic.Const (NReference.reference_of_string ("cic:/matita/ng/Plogic/jmeq/jmeq.ind(1,0,2)")) in
589         let eq_refl = NCic.Const (NReference.reference_of_string ("cic:/matita/ng/Plogic/jmeq/jmeq.con(0,1,2)")) in
590         let add_params 
591           metasenv subst context cty outty mty m i 
592         =
593           let rec aux context outty par j mty m = function
594             | NCic.Prod (name, src, tgt) ->
595                 let t,k = 
596                   aux 
597                     ((name, NCic.Decl src) :: context)
598                     (NCicSubstitution.lift 1 outty) (NCic.Rel j::par) (j+1) 
599                     (NCicSubstitution.lift 1 mty) (NCicSubstitution.lift 1 m) tgt
600                 in
601                 NCic.Prod (name, src, t), k
602             | NCic.Const (Ref.Ref (_,Ref.Ind (_,_,leftno)) as r) ->
603                 let k = 
604                   let k = NCic.Const(Ref.mk_constructor i r) in
605                   NCicUntrusted.mk_appl k par
606                 in
607                 (* the type has no parameters, so kty = mty
608                 let kty = 
609                   try NCicTypechecker.typeof ~subst ~metasenv context k
610                   with NCicTypeChecker.TypeCheckerFailure _ -> assert false
611                 in *)
612                 NCic.Prod ("p", 
613                  NCic.Appl [eq; mty; m; mty; k],
614                  (NCicSubstitution.lift 1
615                   (NCicReduction.head_beta_reduce ~delta:max_int 
616                    (NCicUntrusted.mk_appl outty [k])))),[mty,m,mty,k]
617             | NCic.Appl (NCic.Const (Ref.Ref (_,Ref.Ind (_,_,leftno)) as r)::pl) ->
618                 let left_p,right_p = HExtlib.split_nth leftno pl in
619                 let has_rights = right_p <> [] in
620                 let k = 
621                   let k = NCic.Const(Ref.mk_constructor i r) in
622                   NCicUntrusted.mk_appl k (left_p@par)
623                 in
624                 let right_p = 
625                   try match 
626                     NCicTypeChecker.typeof ~subst ~metasenv context k
627                   with
628                   | NCic.Appl (NCic.Const (Ref.Ref (_,Ref.Ind (_,_,_)))::args) ->
629                       snd (HExtlib.split_nth leftno args)
630                   | _ -> assert false
631                   with NCicTypeChecker.TypeCheckerFailure _-> assert false
632                 in
633                 let orig_right_p =
634                   match mty with
635                   | NCic.Appl (NCic.Const (Ref.Ref (_,Ref.Ind (_,_,_)))::args) ->
636                       snd (HExtlib.split_nth leftno args)
637                   | _ -> assert false
638                 in
639                 List.fold_right2 
640                   (fun x y (tacc,pacc) ->
641                     let xty = 
642                     try NCicTypeChecker.typeof ~subst ~metasenv context x
643                       with NCicTypeChecker.TypeCheckerFailure _ -> assert false
644                     in
645                     let yty = 
646                     try NCicTypeChecker.typeof ~subst ~metasenv context y
647                       with NCicTypeChecker.TypeCheckerFailure _ -> assert false
648                     in
649                     NCic.Prod ("_", NCic.Appl [eq;xty;x;yty;y],
650                      NCicSubstitution.lift 1 tacc), (xty,x,yty,y)::pacc) 
651                   (orig_right_p @ [m]) (right_p @ [k]) 
652                   (NCicReduction.head_beta_reduce ~delta:max_int
653                       (NCicUntrusted.mk_appl outty (right_p@[k])), [])  
654
655   (*              if has_rights then
656                   NCicReduction.head_beta_reduce ~delta:max_int
657                     (NCic.Appl (outty ::right_p @ [k])),k
658                 else
659                   NCic.Prod ("p", 
660                    NCic.Appl [eq; mty; m; k],
661                    (NCicSubstitution.lift 1
662                     (NCicReduction.head_beta_reduce ~delta:max_int 
663                      (NCic.Appl (outty ::right_p @ [k]))))),k*)
664             | _ -> assert false
665           in
666             aux context outty [] 1 mty m cty
667         in
668         let add_lambda_for_refl_m pbo eqs cty =
669           (* k lives in the right context *)
670           (* if rno <> 0 then pbo else *)
671           let rec aux = function 
672             | NCic.Lambda (name,src,bo), NCic.Prod (_,_,ty) ->
673                 NCic.Lambda (name,src,
674                 (aux (bo,ty)))
675             | t,_ -> snd (List.fold_right
676                 (fun (xty,x,yty,y) (n,acc) -> n+1, NCic.Lambda ("p" ^ string_of_int n,
677                   NCic.Appl [eq; xty; x; yty; y],
678                   NCicSubstitution.lift 1 acc)) eqs (1,t))
679                 (*NCic.Lambda ("p",
680                   NCic.Appl [eq; mty; m; k],NCicSubstitution.lift 1 t)*)
681           in
682           aux (pbo,cty)
683         in
684         let add_pi_for_refl_m context new_outty mty m lno rno =
685           (*if rno <> 0 then new_outty else*)
686             let rec aux context mty m = function
687               | NCic.Lambda (name, src, tgt) ->
688                   let context = (name, NCic.Decl src)::context in
689                   NCic.Lambda (name, src, aux context (NCicSubstitution.lift 1 mty) (NCicSubstitution.lift 1 m) tgt)
690               | t -> 
691                   let lhs =
692                     match mty with
693                     | NCic.Appl (_::params) -> (snd (HExtlib.split_nth lno params))@[m]
694                     | _ -> [m]
695                   in
696                   let rhs = 
697                     List.map (fun x -> NCic.Rel x) 
698                       (List.rev (HExtlib.list_seq 1 (rno+2))) in
699                   List.fold_right2
700                     (fun x y acc ->
701                       let xty = 
702                     try NCicTypeChecker.typeof ~subst ~metasenv context x
703                         with NCicTypeChecker.TypeCheckerFailure _ -> assert false
704                       in
705                       let yty = 
706                     try NCicTypeChecker.typeof ~subst ~metasenv context y
707                         with NCicTypeChecker.TypeCheckerFailure _ -> assert false
708                       in
709                       NCic.Prod
710                       ("_", NCic.Appl [eq;xty;x;yty;y],
711                        (NCicSubstitution.lift 1 acc)))
712                     lhs rhs t
713   (*                NCic.Prod 
714                     ("_", NCic.Appl [eq;mty;m;NCic.Rel 1],
715                     NCicSubstitution.lift 1 t)*)
716             in
717               aux context mty m new_outty
718         in (* }}} end helper functions *)
719         (* constructors types with left params already instantiated *)
720         let outty = NCicUntrusted.apply_subst subst context outty in
721         let cl, left_p, leftno,rno = 
722           get_cl_and_left_p r outty
723         in
724         let right_p, mty = 
725           try
726             match 
727               NCicTypeChecker.typeof ~subst ~metasenv context m
728             with
729             | (NCic.Const (Ref.Ref (_,Ref.Ind (_,_,_))) | NCic.Meta _) as mty -> [], mty
730             | NCic.Appl ((NCic.Const (Ref.Ref (_,Ref.Ind (_,_,_)))|NCic.Meta _)::args) as mty ->
731                 snd (HExtlib.split_nth leftno args), mty
732             | _ -> assert false
733           with NCicTypeChecker.TypeCheckerFailure _ -> 
734              raise (AssertFailure(lazy "already ill-typed matched term"))
735         in
736         let mty = NCicUntrusted.apply_subst subst context mty in
737         let outty = NCicUntrusted.apply_subst subst context outty in
738         let expty = NCicUntrusted.apply_subst subst context expty in
739         let new_outty =
740          keep_lambdas_and_put_expty context outty expty right_p m (rno+1)
741         in
742         pp 
743           (lazy ("CASE: new_outty: " ^ NCicPp.ppterm 
744            ~context ~metasenv ~subst new_outty));
745         let _,pl,subst,metasenv = 
746           List.fold_right2
747             (fun cty pbo (i, acc, s, menv) -> 
748                pp (lazy ("begin iteration"));
749               (* Pi k_par, (Pi H:m=(K_i left_par k_par)), 
750                *   (new_)outty right_par (K_i left_par k_par) *)
751                let infty_pbo, _ = 
752                  add_params menv s context cty outty mty m i in
753                pp 
754                 (lazy ("CASE: infty_pbo: "^ NCicPp.ppterm
755                  ~context ~metasenv ~subst infty_pbo));
756                let expty_pbo, eqs = (* k is (K_i left_par k_par) *)
757                  add_params menv s context cty new_outty mty m i in
758                pp 
759                 (lazy ("CASE: expty_pbo: "^ NCicPp.ppterm
760                  ~context ~metasenv ~subst expty_pbo));
761                let pbo = add_lambda_for_refl_m pbo eqs cty in
762                pp 
763                  (lazy ("CASE: pbo: " ^ NCicPp.ppterm
764                  ~context ~metasenv ~subst pbo));
765                let metasenv, subst, pbo, _ =
766                  try_coercions rdb ~localise menv s context pbo 
767                  orig_t (*??*) infty_pbo expty_pbo perform_unification exc
768                in
769                pp 
770                  (lazy ("CASE: pbo2: " ^ NCicPp.ppterm 
771                  ~context ~metasenv ~subst pbo));
772                (i-1, pbo::acc, subst, metasenv))
773             cl pl (List.length pl, [], subst, metasenv)
774         in
775         (*let metasenv, subst = 
776           try 
777             NCicUnification.unify rdb metasenv subst context outty new_outty 
778           with _ -> raise (RefineFailure (lazy (localise orig_t, "try_coercions")))
779         in*)
780         let new_outty = add_pi_for_refl_m context new_outty mty m leftno rno in
781         pp (lazy ("CASE: new_outty: " ^ (NCicPp.ppterm 
782            ~metasenv ~subst ~context new_outty)));
783         let right_tys = 
784           List.map 
785             (fun t -> NCicTypeChecker.typeof ~subst ~metasenv context t) right_p in
786         let eqs = 
787           List.map2 (fun x y -> NCic.Appl[eq_refl;x;y]) (right_tys@[mty]) 
788             (right_p@[m]) in
789         let t = NCic.Appl (NCic.Match(r,new_outty,m,pl) :: eqs) 
790         in
791         metasenv, subst, t, expty (*}}}*)
792     | NCic.Prod (nameprod, src, ty),NCic.Prod (_, src2, ty2), _ -> 
793         let rec mk_fresh_name ctx firstch n =
794           let candidate = (String.make 1 firstch) ^ (string_of_int n) in
795           if (List.for_all (fun (s,_) -> s <> candidate) ctx) then candidate
796           else mk_fresh_name ctx firstch (n+1)
797         in
798         (*{{{*) pp (lazy "LAM");
799         pp (lazy ("LAM: t = " ^ NCicPp.ppterm ~metasenv ~subst ~context t));
800         let name_con = mk_fresh_name context 'c' 0
801           (*FreshNamesGenerator.mk_fresh_name 
802             ~subst metasenv context ~typ:src2 Cic.Anonymous*)
803         in
804         let context_src2 = ((name_con, NCic.Decl src2) :: context) in
805         (* contravariant part: the argument of f:src->ty *)
806         let metasenv, subst, rel1, _ = 
807           try_coercions rdb ~localise metasenv subst context_src2
808            (NCic.Rel 1) orig_t (NCicSubstitution.lift 1 src2) 
809            (NCicSubstitution.lift 1 src) perform_unification exc
810         in
811         (* covariant part: the result of f(c x); x:src2; (c x):src *)
812         let name_t, bo = 
813           match t with
814           | NCic.Lambda (n,_,bo) -> n, cs_subst rel1 (NCicSubstitution.lift_from 2 1 bo)
815           | _ -> name_con, NCicUntrusted.mk_appl (NCicSubstitution.lift 1 t) [rel1]
816         in
817         (* we fix the possible dependency problem in the source ty *)
818         let ty = cs_subst rel1 (NCicSubstitution.lift_from 2 1 ty) in
819         let metasenv, subst, bo, _ = 
820           try_coercions rdb ~localise metasenv subst context_src2
821             bo orig_t ty ty2 perform_unification exc
822         in
823         let coerced = NCic.Lambda (name_t,src2, bo) in
824         pp (lazy ("LAM: coerced = " ^ NCicPp.ppterm ~metasenv ~subst ~context coerced));
825         metasenv, subst, coerced, expty (*}}}*)
826     | _ -> raise exc
827   with exc2 ->    
828   pp(lazy("try_coercion " ^ 
829     NCicPp.ppterm ~metasenv ~subst ~context infty ^ " |---> " ^
830     NCicPp.ppterm ~metasenv ~subst ~context expty));
831     first exc
832       (NCicCoercion.look_for_coercion 
833         rdb metasenv subst context infty expty)
834
835 and force_to_sort rdb metasenv subst context t orig_t localise ty =
836   try 
837     let metasenv, subst, ty = 
838       NCicUnification.sortfy (Failure "sortfy") metasenv subst context ty in
839       metasenv, subst, t, ty
840   with
841       Failure _ -> 
842         let ty = NCicReduction.whd ~subst context ty in
843           try_coercions rdb ~localise metasenv subst context
844             t orig_t ty (NCic.Sort (NCic.Type 
845               (match NCicEnvironment.get_universes () with
846                | x::_ -> x 
847                | _ -> assert false))) false 
848              (Uncertain (lazy (localise orig_t, 
849              "The type of " ^ NCicPp.ppterm ~metasenv ~subst ~context t ^
850              " is not a sort: " ^ NCicPp.ppterm ~metasenv ~subst ~context ty)))
851
852 and sort_of_prod 
853   localise metasenv subst context orig_s orig_t (name,s) t (t1, t2) 
854 =
855    (* force to sort is done in the Prod case in typeof *)
856    match t1, t2 with
857    | C.Sort _, C.Sort C.Prop -> metasenv, subst, s, t, t2
858    | C.Sort (C.Type u1), C.Sort (C.Type u2) ->
859         metasenv, subst, s, t, C.Sort (C.Type (NCicEnvironment.max u1 u2)) 
860    | C.Sort C.Prop,C.Sort (C.Type _) -> metasenv, subst, s, t, t2
861    | C.Meta _, C.Sort _ 
862    | C.Meta _, C.Meta (_,(_,_))
863    | C.Sort _, C.Meta (_,(_,_)) -> metasenv, subst, s, t, t2 
864    | x, (C.Sort _ | C.Meta _) | _, x -> 
865       let y, context, orig = 
866         if x == t1 then s, context, orig_s 
867         else t, ((name,C.Decl s)::context), orig_t
868       in
869       raise (RefineFailure (lazy (localise orig,Printf.sprintf
870         "%s is expected to be a type, but its type is %s that is not a sort" 
871          (NCicPp.ppterm ~subst ~metasenv ~context y) 
872          (NCicPp.ppterm ~subst ~metasenv ~context x))))
873
874 and guess_name subst ctx ty = 
875   let aux initial = "#" ^ String.make 1 initial in
876   match ty with
877   | C.Const (NReference.Ref (u,_))
878   | C.Appl (C.Const (NReference.Ref (u,_)) :: _) ->
879       aux (String.sub (NUri.name_of_uri u) 0 1).[0] 
880   | C.Prod _ -> aux 'f' 
881   | C.Meta (n,lc) -> 
882       (try
883          let _,_,t,_ = NCicUtils.lookup_subst n subst in
884          guess_name subst ctx (NCicSubstitution.subst_meta lc t)
885       with NCicUtils.Subst_not_found _ -> aux 'M')
886   | _ -> aux 'H' 
887
888 and eat_prods rdb ~localise force_ty metasenv subst context expty orig_t orig_he he ty_he args =
889   (*D*)inside 'E'; try let rc = 
890   let rec aux metasenv subst args_so_far he ty_he xxx =
891   (*D*)inside 'V'; try let rc = 
892    match xxx with
893   | [] ->
894      let res = NCicUntrusted.mk_appl he (List.rev args_so_far) in
895      pp(lazy("FORCE FINAL APPL: " ^ 
896        NCicPp.ppterm ~metasenv ~subst ~context res ^
897        " of type " ^ NCicPp.ppterm ~metasenv ~subst ~context ty_he
898        ^ " to type " ^ match expty with None -> "None" | Some x -> 
899        NCicPp.ppterm ~metasenv ~subst ~context x));
900      (* whatever the term is, we force the type. in case of ((Lambda..) ?...)
901       * the application may also be a lambda! *)
902      force_ty false false metasenv subst context orig_t res ty_he expty
903   | NCic.Implicit `Vector::tl ->
904       let has_some_more_pis x =
905         match NCicReduction.whd ~subst context x with
906         |  NCic.Meta _ | NCic.Appl (NCic.Meta _::_) -> false
907         | _ -> true
908       in
909      (try
910        aux metasenv subst args_so_far he ty_he tl
911       with
912       | Uncertain _
913       | RefineFailure _ as exc when has_some_more_pis ty_he ->
914           (try
915            aux metasenv subst args_so_far he ty_he
916             (NCic.Implicit `Term :: NCic.Implicit `Vector :: tl)
917           with
918            Uncertain msg | RefineFailure msg -> raise (wrap_exc msg exc))
919      | RefineFailure msg when not (has_some_more_pis ty_he) ->
920         (* instantiating the head could change the has_some_more_pis flag *)
921         raise (Uncertain msg))
922   | arg::tl ->
923       match NCicReduction.whd ~subst context ty_he with 
924       | C.Prod (_,s,t) ->
925           let metasenv, subst, arg, _ = 
926             typeof rdb ~localise metasenv subst context arg (Some s) in
927           let t = NCicSubstitution.subst ~avoid_beta_redexes:true arg t in
928           aux metasenv subst (arg :: args_so_far) he t tl
929       | C.Meta _
930       | C.Appl (C.Meta _ :: _) as t ->
931           let metasenv, subst, arg, ty_arg = 
932             typeof rdb ~localise metasenv subst context arg None in
933           let name = guess_name subst context ty_arg in
934           let metasenv, _, meta, _ = 
935             NCicMetaSubst.mk_meta metasenv 
936               ((name,C.Decl ty_arg) :: context) `IsType
937           in
938           let flex_prod = C.Prod (name, ty_arg, meta) in
939           (* next line grants that ty_args is a type *)
940           let metasenv,subst, flex_prod, _ =
941            typeof rdb ~localise metasenv subst context flex_prod None in
942 (*
943           pp (lazy ( "UNIFICATION in CTX:\n"^ 
944             NCicPp.ppcontext ~metasenv ~subst context
945             ^ "\nOF: " ^
946             NCicPp.ppterm ~metasenv ~subst ~context t ^  " === " ^
947             NCicPp.ppterm ~metasenv ~subst ~context flex_prod ^ "\n"));
948 *)
949           let metasenv, subst =
950              try NCicUnification.unify rdb metasenv subst context t flex_prod 
951              with exc -> raise (wrap_exc (lazy (localise orig_he, Printf.sprintf
952               ("The term %s has an inferred type %s, but is applied to the" ^^
953                " argument %s of type %s")
954               (NCicPp.ppterm ~metasenv ~subst ~context he)
955               (NCicPp.ppterm ~metasenv ~subst ~context t)
956               (NCicPp.ppterm ~metasenv ~subst ~context arg)
957               (NCicPp.ppterm ~metasenv ~subst ~context ty_arg))) 
958                  (match exc with
959                  | NCicUnification.UnificationFailure m -> 
960                      NCicUnification.Uncertain m
961                  | x -> x))
962               (* XXX coerce to funclass *)
963           in
964           let meta = NCicSubstitution.subst ~avoid_beta_redexes:true arg meta in
965           aux metasenv subst (arg :: args_so_far) he meta tl
966       | C.Match (_,_,C.Meta _,_) 
967       | C.Match (_,_,C.Appl (C.Meta _ :: _),_) 
968       | C.Appl (C.Const (NReference.Ref (_, NReference.Fix _)) :: _) ->
969           raise (Uncertain (lazy (localise orig_he, Printf.sprintf
970             ("The term %s is here applied to %d arguments but expects " ^^
971             "only %d arguments") (NCicPp.ppterm ~metasenv ~subst ~context he)
972             (List.length args) (List.length args_so_far))))
973       | ty ->
974           let metasenv, subst, newhead, newheadty = 
975             try_coercions rdb ~localise metasenv subst context
976               (NCicUntrusted.mk_appl he (List.rev args_so_far)) orig_he ty
977               (NCic.Prod ("_",NCic.Implicit `Term,NCic.Implicit `Term))
978               false
979               (RefineFailure (lazy (localise orig_he, Printf.sprintf
980                ("The term %s is here applied to %d arguments but expects " ^^
981                "only %d arguments") (NCicPp.ppterm ~metasenv ~subst ~context he)
982                (List.length args) (List.length args_so_far))))
983           in
984            aux metasenv subst [] newhead newheadty (arg :: tl)
985   (*D*)in outside true; rc with exc -> outside false; raise exc
986   in
987    (* We need to reverse the order of the new created metas since they
988       are pushed on top of the metasenv in the wrong order *)
989    let highest_meta = NCicMetaSubst.maxmeta () in
990    let metasenv, subst, newhead, newheadty = 
991     aux metasenv subst [] he ty_he args in
992    let metasenv_old,metasenv_new =
993     List.partition (fun (i,_) -> i <= highest_meta) metasenv
994    in
995     (List.rev metasenv_new) @ metasenv_old, subst, newhead, newheadty
996   (*D*)in outside true; rc with exc -> outside false; raise exc
997 ;;
998
999 let rec first f l1 l2 =
1000   match l1,l2 with
1001   | x1::tl1, x2::tl2 -> 
1002       (try f x1 x2 with Not_found -> first f tl1 tl2)
1003   | _ -> raise Not_found
1004 ;;
1005
1006 let rec find add dt t =
1007   if dt == add then t
1008   else
1009     let dl, l = 
1010       match dt, t with
1011       | C.Meta (_,(_,C.Ctx dl)), C.Meta (_,(_,C.Ctx l))
1012       | C.Appl dl,C.Appl l -> dl,l
1013       | C.Lambda (_,ds,dt), C.Lambda (_,s,t) 
1014       | C.Prod (_,ds,dt), C.Prod (_,s,t) -> [ds;dt],[s;t]
1015       | C.LetIn (_,ds,db,dt), C.LetIn (_,s,b,t) -> [ds;db;dt],[s;b;t] 
1016       | C.Match (_,dot,dt,dl),  C.Match (_,ot,t,l) -> (dot::dt::dl),(ot::t::l)
1017       | _ -> raise Not_found
1018     in
1019       first (find add) dl l
1020 ;;
1021
1022 let relocalise old_localise dt t add = 
1023   old_localise 
1024     (try find add dt t with Not_found -> assert false)
1025 ;;
1026
1027 let undebruijnate inductive ref t rev_fl =
1028   let len = List.length rev_fl in 
1029   NCicSubstitution.psubst (fun x -> x) 
1030    (HExtlib.list_mapi 
1031       (fun (_,_,rno,_,_,_) i -> 
1032          let i = len - i - 1 in
1033          NCic.Const 
1034            (if inductive then NReference.mk_fix i rno ref
1035             else NReference.mk_cofix i ref))
1036       rev_fl)
1037     t
1038 ;; 
1039
1040
1041 let typeof_obj 
1042   rdb ?(localise=fun _ -> Stdpp.dummy_loc) (uri,height,metasenv,subst,obj)
1043
1044   match obj with 
1045   | C.Constant (relevance, name, bo, ty, attr) ->
1046        let metasenv, subst, ty = 
1047          check_type rdb ~localise metasenv subst [] ty in
1048        let metasenv, subst, bo, ty, height = 
1049          match bo with
1050          | Some bo ->
1051              let metasenv, subst, bo, ty = 
1052                typeof rdb ~localise metasenv subst [] bo (Some ty) in
1053              let height = (* XXX recalculate *) height in
1054                metasenv, subst, Some bo, ty, height
1055          | None -> metasenv, subst, None, ty, 0
1056        in
1057        uri, height, metasenv, subst, 
1058          C.Constant (relevance, name, bo, ty, attr)
1059   | C.Fixpoint (inductive, fl, attr) -> 
1060       let len = List.length fl in
1061       let types, metasenv, subst, rev_fl =
1062         List.fold_left
1063          (fun (types, metasenv, subst, fl) (relevance,name,k,ty,bo) ->
1064            let metasenv, subst, ty = 
1065              check_type rdb ~localise metasenv subst [] ty in
1066            let dbo = NCicTypeChecker.debruijn uri len [] ~subst bo in
1067            let localise = relocalise localise dbo bo in
1068             (name,C.Decl ty)::types,
1069               metasenv, subst, (relevance,name,k,ty,dbo,localise)::fl
1070          ) ([], metasenv, subst, []) fl (* XXX kl rimosso nel nucleo *)
1071       in
1072       let metasenv, subst, fl = 
1073         List.fold_left 
1074           (fun (metasenv,subst,fl) (relevance,name,k,ty,dbo,localise) ->
1075             let metasenv, subst, dbo, ty = 
1076               typeof rdb ~localise metasenv subst types dbo (Some ty)
1077             in
1078             metasenv, subst, (relevance,name,k,ty,dbo)::fl)
1079           (metasenv, subst, []) rev_fl
1080       in
1081       let height = (* XXX recalculate *) height in
1082       let fl =
1083         List.map 
1084           (fun (relevance,name,k,ty,dbo) ->
1085             let bo = 
1086               undebruijnate inductive 
1087                (NReference.reference_of_spec uri 
1088                  (if inductive then NReference.Fix (0,k,0) 
1089                   else NReference.CoFix 0)) dbo rev_fl
1090             in
1091               relevance,name,k,ty,bo)
1092           fl
1093       in
1094        uri, height, metasenv, subst, 
1095          C.Fixpoint (inductive, fl, attr)
1096   | C.Inductive (ind, leftno, itl, attr) ->
1097      let len = List.length itl in
1098      let metasenv,subst,rev_itl,tys =
1099       List.fold_left
1100        (fun (metasenv,subst,res,ctx) (relevance,n,ty,cl) ->
1101           let metasenv, subst, ty = 
1102             check_type rdb ~localise metasenv subst [] ty in
1103           metasenv,subst,(relevance,n,ty,cl)::res,(n,NCic.Decl ty)::ctx
1104        ) (metasenv,subst,[],[]) itl in
1105      let metasenv,subst,itl,_ =
1106       List.fold_left
1107        (fun (metasenv,subst,res,i) (it_relev,n,ty,cl) ->
1108          let context,ty_sort = NCicReduction.split_prods ~subst [] ~-1 ty in
1109          let sx_context_ty_rev,_= HExtlib.split_nth leftno (List.rev context) in
1110          let metasenv,subst,cl =
1111           List.fold_right
1112            (fun (k_relev,n,te) (metasenv,subst,res) ->
1113              let k_relev =
1114               try snd (HExtlib.split_nth leftno k_relev)
1115               with Failure _ -> k_relev in
1116              let te = NCicTypeChecker.debruijn uri len [] ~subst te in
1117              let metasenv, subst, te = 
1118                check_type rdb ~localise metasenv subst tys te in
1119              let context,te = NCicReduction.split_prods ~subst tys leftno te in
1120              let _,chopped_context_rev =
1121               HExtlib.split_nth (List.length tys) (List.rev context) in
1122              let sx_context_te_rev,_ =
1123               HExtlib.split_nth leftno chopped_context_rev in
1124              let metasenv,subst,_ =
1125               try
1126                List.fold_left2
1127                 (fun (metasenv,subst,context) item1 item2 ->
1128                   let (metasenv,subst),convertible =
1129                    try
1130                     match item1,item2 with
1131                       (n1,C.Decl ty1),(n2,C.Decl ty2) ->
1132                         if n1 = n2 then
1133                          NCicUnification.unify rdb ~test_eq_only:true metasenv
1134                           subst context ty1 ty2,true
1135                         else
1136                          (metasenv,subst),false
1137                     | (n1,C.Def (bo1,ty1)),(n2,C.Def (bo2,ty2)) ->
1138                         if n1 = n2 then
1139                          let metasenv,subst =
1140                           NCicUnification.unify rdb ~test_eq_only:true metasenv
1141                            subst context ty1 ty2
1142                          in
1143                           NCicUnification.unify rdb ~test_eq_only:true metasenv
1144                            subst context bo1 bo2,true
1145                         else
1146                          (metasenv,subst),false
1147                     | _,_ -> (metasenv,subst),false
1148                    with
1149                    | NCicUnification.Uncertain _
1150                    | NCicUnification.UnificationFailure _ ->
1151                       (metasenv,subst),false
1152                   in
1153                    let term2 =
1154                     match item2 with
1155                        _,C.Decl t -> t
1156                      | _,C.Def (b,_) -> b in
1157                    if not convertible then
1158                     raise (RefineFailure (lazy (localise term2,
1159                      ("Mismatch between the left parameters of the constructor " ^
1160                       "and those of its inductive type"))))
1161                    else
1162                     metasenv,subst,item1::context
1163                 ) (metasenv,subst,tys) sx_context_ty_rev sx_context_te_rev
1164               with Invalid_argument "List.fold_left2" -> assert false in
1165              let metasenv, subst =
1166                 let rec aux context (metasenv,subst) = function
1167                   | NCic.Meta _ -> metasenv, subst
1168                   | NCic.Implicit _ -> metasenv, subst
1169                   | NCic.Appl (NCic.Rel i :: args) as t
1170                       when i > List.length context - len ->
1171                       let lefts, _ = HExtlib.split_nth leftno args in
1172                       let ctxlen = List.length context in
1173                       let (metasenv, subst), _ = 
1174                         List.fold_left
1175                         (fun ((metasenv, subst),l) arg ->
1176                           NCicUnification.unify rdb 
1177                            ~test_eq_only:true metasenv subst context arg 
1178                              (NCic.Rel (ctxlen - len - l)), l+1
1179                           )
1180                         ((metasenv, subst), 0) lefts
1181                       in
1182                       metasenv, subst
1183                   | t -> NCicUtils.fold (fun e c -> e::c) context aux
1184                   (metasenv,subst) t
1185                 in
1186                aux context (metasenv,subst) te
1187              in
1188              let con_sort= NCicTypeChecker.typeof ~subst ~metasenv context te in
1189               (match
1190                 NCicReduction.whd ~subst context con_sort,
1191                 NCicReduction.whd ~subst [] ty_sort
1192                with
1193                   (C.Sort (C.Type u1) as s1), (C.Sort (C.Type u2) as s2) ->
1194                    if not (NCicEnvironment.universe_leq u1 u2) then
1195                     raise
1196                      (RefineFailure
1197                        (lazy(localise te, "The type " ^
1198                          NCicPp.ppterm ~metasenv ~subst ~context s1 ^
1199                          " of the constructor is not included in the inductive"^
1200                          " type sort " ^
1201                          NCicPp.ppterm ~metasenv ~subst ~context s2)))
1202                 | C.Sort _, C.Sort C.Prop
1203                 | C.Sort _, C.Sort C.Type _ -> ()
1204                 | _, _ ->
1205                    raise
1206                     (RefineFailure
1207                       (lazy (localise te,
1208                         "Wrong constructor or inductive arity shape"))));
1209               (* let's check also the positivity conditions *)
1210               if 
1211                not
1212                (NCicTypeChecker.are_all_occurrences_positive
1213                  ~subst context uri leftno (i+leftno) leftno (len+leftno) te) 
1214               then
1215                 raise
1216                   (RefineFailure
1217                     (lazy (localise te,
1218                       "Non positive occurence in " ^
1219                         NCicPp.ppterm ~metasenv ~subst ~context te)))
1220               else
1221                let relsno = List.length itl + leftno in
1222                let te = 
1223                  NCicSubstitution.psubst 
1224                   (fun i ->
1225                     if i <= leftno  then
1226                      NCic.Rel i
1227                     else
1228                      NCic.Const (NReference.reference_of_spec uri
1229                       (NReference.Ind (ind,relsno - i,leftno))))
1230                   (HExtlib.list_seq 1 (relsno+1))
1231                    te in
1232                let te =
1233                 List.fold_right
1234                  (fun (name,decl) te ->
1235                    match decl with
1236                       NCic.Decl ty -> NCic.Prod (name,ty,te)
1237                     | NCic.Def (bo,ty) -> NCic.LetIn (name,ty,bo,te)
1238                  ) sx_context_te_rev te
1239                in
1240                 metasenv,subst,(k_relev,n,te)::res
1241               ) cl (metasenv,subst,[])
1242          in
1243           metasenv,subst,(it_relev,n,ty,cl)::res,i+1
1244        ) (metasenv,subst,[],1) rev_itl
1245      in
1246       uri, height, metasenv, subst, C.Inductive (ind, leftno, itl, attr)
1247 ;;
1248
1249 (* vim:set foldmethod=marker: *)