]> matita.cs.unibo.it Git - helm.git/commitdiff
initial refiner ....
authorEnrico Tassi <enrico.tassi@inria.fr>
Mon, 13 Oct 2008 17:42:30 +0000 (17:42 +0000)
committerEnrico Tassi <enrico.tassi@inria.fr>
Mon, 13 Oct 2008 17:42:30 +0000 (17:42 +0000)
helm/software/components/ng_kernel/nCic.ml
helm/software/components/ng_kernel/nCicTypeChecker.ml
helm/software/components/ng_kernel/nCicUntrusted.ml
helm/software/components/ng_kernel/nCicUntrusted.mli
helm/software/components/ng_refiner/check.ml
helm/software/components/ng_refiner/nCicMetaSubst.ml
helm/software/components/ng_refiner/nCicMetaSubst.mli
helm/software/components/ng_refiner/nCicRefiner.ml
helm/software/components/ng_refiner/nCicRefiner.mli

index 1f02ee7d8e5406c713f32cff20d1e2926699a436..97733891e8745b4e3251500878f4804abe1743ef 100644 (file)
 (********************************* TERMS ************************************)
 
 type universe = (bool * NUri.uri) list 
-  (* Max of non-empty list of named universes, or their successor (when true) *)
+  (* Max of non-empty list of named universes, or their successor (when true) 
+   * The empty list represents type0 *)
 
 type sort = Prop | Type of universe
 
-type implicit_annotation = [ `Closed | `Type | `Hole | `Term ]
+type implicit_annotation = [ `Closed | `Type | `Hole | `Term | `Typeof of int ]
 
 type lc_kind = Irl of int | Ctx of term list
 
index 806238ce3580330f7409289805b4bf655865334c..1d53f36f40fa39a3ff717b2f0bc543c138b0e417 100644 (file)
@@ -121,7 +121,7 @@ let debruijn uri number_of_types context =
   aux (List.length context)
 ;;
 
-let sort_of_prod ~metasenv ~subst context (name,s) (t1, t2) =
+let sort_of_prod ~metasenv ~subst context (name,s) (t1, t2) =
    let t1 = R.whd ~subst context t1 in
    let t2 = R.whd ~subst ((name,C.Decl s)::context) t2 in
    match t1, t2 with
@@ -131,11 +131,15 @@ let sort_of_prod ~metasenv ~subst context (name,s) (t1, t2) =
    | C.Meta (_,(_,(C.Irl 0 | C.Ctx []))), C.Sort _
    | C.Meta (_,(_,(C.Irl 0 | C.Ctx []))), C.Meta (_,(_,(C.Irl 0 | C.Ctx [])))
    | C.Sort _, C.Meta  (_,(_,(C.Irl 0 | C.Ctx []))) -> t2
-   | _ -> 
+   | x, (C.Sort _ | C.Meta (_,(_,(C.Irl 0 | C.Ctx []))))
+   | _, x -> 
+      let y, context = 
+        if x == t1 then s, context else t, ((name,C.Decl s)::context)
+      in
       raise (TypeCheckerFailure (lazy (Printf.sprintf
-        "Prod: expected two sorts, found = %s, %s
-         (PP.ppterm ~subst ~metasenv ~context t1
-         (PP.ppterm ~subst ~metasenv ~context t2))))
+        "%s is expected to be a type, but its type is %s that is not a sort
+         (PP.ppterm ~subst ~metasenv ~context y
+         (PP.ppterm ~subst ~metasenv ~context x))))
 ;;
 
 (* instantiate_parameters ps (x1:T1)...(xn:Tn)C                             *)
@@ -361,7 +365,8 @@ let rec typeof ~subst ~metasenv context term =
         try
          let _,c,_,ty = U.lookup_subst n subst in c,ty
         with U.Subst_not_found _ -> try
-         let _,c,ty = U.lookup_meta n metasenv in c,ty
+         let _,c,ty = U.lookup_meta n metasenv in c, ty
+(*          match ty with C.Implicit _ -> assert false | _ -> c,ty *)
         with U.Meta_not_found _ ->
          raise (AssertFailure (lazy (Printf.sprintf
           "%s not found" (PP.ppterm ~subst ~metasenv ~context t))))
@@ -372,7 +377,7 @@ let rec typeof ~subst ~metasenv context term =
     | C.Prod (name,s,t) ->
        let sort1 = typeof_aux context s in
        let sort2 = typeof_aux ((name,(C.Decl s))::context) t in
-       sort_of_prod ~metasenv ~subst context (name,s) (sort1,sort2)
+       sort_of_prod ~metasenv ~subst context (name,s) (sort1,sort2)
     | C.Lambda (n,s,t) ->
        let sort = typeof_aux context s in
        (match R.whd ~subst context sort with
index 167fd47c9e313fa5e574a9f3794d25d47cf12eb0..57a2a4f178d5056c6131778a9458938c1b25afd1 100644 (file)
@@ -72,3 +72,12 @@ let metas_of_term subst context term =
     HExtlib.list_uniq (List.sort Pervasives.compare (aux context [] term))
 ;;
 
+module NCicHash =
+ Hashtbl.Make
+  (struct
+    type t = C.term
+    let equal = (==)
+    let hash = Hashtbl.hash_param 100 1000 
+   end)
+;;
+
index 72e512784cdee4ea5129ac92671ea4ac51003b9c..f752c01dc9a668c23cdd0b3b715b4c68081d993a 100644 (file)
@@ -17,3 +17,5 @@ val map_term_fold_a:
 
 val metas_of_term : NCic.substitution -> NCic.context -> NCic.term -> int list
 
+module NCicHash : Hashtbl.S
+
index 82aff623b84c66a5919ac6ffc149c4aa94ca033b..58428dca758b7c4c4b11c9786b56973b3be200fb 100644 (file)
@@ -209,7 +209,7 @@ let _ =
             | NCic.Appl (NCic.Const (NReference.Ref (u,_))::ty::_)
               when NUri.string_of_uri u = "cic:/matita/tests/hole.con" ->
                 let menv, ty =  perforate ctx menv ty in
-                NCicMetaSubst.mk_meta menv ctx ty
+                NCicMetaSubst.mk_meta menv ctx (Some ty)
             | t -> 
                 NCicUntrusted.map_term_fold_a
                  (fun e ctx -> e::ctx) ctx perforate menv t
index 967074a564fd0aaca79b44ea771cd2e0eae74fbf..102958869817e6123245a4f299a0b3f0b91ad4c6 100644 (file)
@@ -571,17 +571,27 @@ let delift_rels subst metasenv n t =
 *) 
 
 let mk_meta ?name metasenv context ty = 
-  let n = newmeta () in
-  let len = List.length context in
-  let menv_entry = (n, (name, context, ty)) in
-  menv_entry :: metasenv, NCic.Meta (n, (0,NCic.Irl len))
+  match ty with
+  | None ->
+    let len = List.length context in
+    let n = newmeta () in
+    let ty_menv_entry = (n, (name, context, NCic.Implicit (`Typeof n))) in
+    let m = newmeta () in
+    let ty = NCic.Meta (n, (0,NCic.Irl len)) in
+    let menv_entry = (m, (name, context, ty)) in
+    menv_entry :: ty_menv_entry :: metasenv, NCic.Meta (m, (0,NCic.Irl len)), ty
+  | Some ty ->
+    let n = newmeta () in
+    let len = List.length context in
+    let menv_entry = (n, (name, context, ty)) in
+    menv_entry :: metasenv, NCic.Meta (n, (0,NCic.Irl len)), ty
 ;;
 
 let saturate ?(delta=0) metasenv context ty goal_arity =
  assert (goal_arity >= 0);
   let rec aux metasenv = function
    | NCic.Prod (name,s,t) ->
-       let metasenv1, arg = mk_meta ~name:name metasenv context s in            
+       let metasenv1, arg,_ = mk_meta ~name:name metasenv context (Some s) in            
        let t, metasenv1, args, pno = 
            aux metasenv1 (NCicSubstitution.subst arg t) 
        in
index 6c6ddf554e55def7d020a33a32bffba440e39894..f8af0944e6a95ca0b1cbd91f2c339ebb1cd77ddb 100644 (file)
@@ -53,8 +53,9 @@ val restrict:
       int -> int list -> NCic.metasenv * NCic.substitution * int
 
 val mk_meta: 
-    ?name:string ->
-    NCic.metasenv -> NCic.context -> NCic.term -> NCic.metasenv * NCic.term
+    ?name:string -> 
+    NCic.metasenv -> NCic.context -> NCic.term option -> 
+      NCic.metasenv * NCic.term * NCic.term (* menv, instance, type *)
 
 val saturate:
     ?delta:int -> NCic.metasenv -> NCic.context -> NCic.term -> int ->
index eb084204cc3adca745d143fe3f591d39e1fd4adb..de8b98cfffdfe04b2df88c4c95d6af533a3b3800 100644 (file)
 
 (* $Id$ *)
 
+let exp_implicit metasenv subst context expty =
+ function      
+  | Some `Closed -> NCicMetaSubst.mk_meta metasenv [] expty
+  | Some `Type | None -> NCicMetaSubst.mk_meta metasenv context expty
+  | _ -> assert false
+;;
+
+let sort_of_prod localise metasenv subst context (name,s) t (t1, t2) =
+   let restrict metasenv subst = function
+     | NCic.Meta (i,(_,(NCic.Irl 0 | NCic.Ctx []))) as t -> 
+        metasenv, subst, t
+     | NCic.Meta (i,(_,lc)) as t ->
+        let len = match lc with NCic.Irl len->len | NCic.Ctx l->List.lenght l in
+        let metasenv, subst, _ = 
+          NCicMetaSubst.restrict metasenv subst i (HExtlib.seq 1 len) 
+        in
+        metasenv, subst, t
+     | t -> metasenv, subst, t 
+   in
+   let t1 = R.whd ~subst context t1 in
+   let t2 = R.whd ~subst ((name,C.Decl s)::context) t2 in
+   let metasenv, subst, t1 = restrict metasenv subst t1 in
+   let metasenv, subst, t2 = restrict metasenv subst t2 in
+   match t1, t2 with
+   | C.Sort _, C.Sort C.Prop -> metasenv, subst, t2
+   | C.Sort (C.Type u1), C.Sort (C.Type u2) ->
+        metasenv, subst, C.Sort (C.Type (u1@u2)) 
+   | C.Sort C.Prop,C.Sort (C.Type _) -> metasenv, subst, t2
+   | C.Meta _, C.Sort _
+   | C.Meta _, C.Meta _
+   | C.Sort _, C.Meta  _ -> metasenv, subst, t2
+   | x, _ | _, x -> 
+      let y, context = 
+        if x == t1 then s, context else t, ((name,C.Decl s)::context)
+      in
+      raise (TypeCheckerFailure (lazy (Printf.sprintf
+        "%s is expected to be a type, but its type is %s that is not a sort" 
+         (PP.ppterm ~subst ~metasenv ~context y) 
+         (PP.ppterm ~subst ~metasenv ~context x))))
+;;
+
+let rec typeof 
+  ?(localize=fun _ -> Stdpp.dummy) metasenv subst context term expty 
+=
+  let force_ty metasenv subst context t infty = function
+    | Some expty ->
+       (match t with
+       | NCic.Implicit _ -> metasenv, subst, t, expty
+       | _ ->
+         let metasenv, subst = 
+           NCicUnification.unify metasenv subst context infty expty in
+         metasenv, subst, t, expty)
+    | None -> metasenv, subst, t, infty
+  in
+  let rec typeof_aux metasenv subst context expty = 
+    fun t -> (*prerr_endline (PP.ppterm ~metasenv ~subst ~context t);*)
+    let metasenv, subst, t, infty = 
+    match t with
+    | C.Rel n ->
+        let infty = 
+         (try
+           match List.nth context (n - 1) with
+           | (_,C.Decl ty) -> S.lift n ty
+           | (_,C.Def (_,ty)) -> S.lift n ty
+         with Failure _ -> 
+           raise (RefineFailure (lazy (localize t,"unbound variable"))))
+        in
+        metasenv, subst, t, infty
+    | C.Sort (C.Type [false,u]) -> metasenv,subst,t,(C.Sort (C.Type [true, u]))
+    | C.Sort (C.Type _) -> 
+        raise (AssertFailure (lazy ("Cannot type an inferred type: "^
+          NCicPp.ppterm ~subst ~metasenv ~context t)))
+    | C.Sort _ -> metasenv,subst,t,(C.Sort (C.Type NCicEnvironment.type0))
+    | C.Implicit infos -> 
+         let metasenv,t,ty = exp_implicit metasenv subst context infos expty in
+         metasenv, subst, t, ty 
+    | C.Meta (n,l) as t -> 
+       let ty =
+        try
+         let _,_,_,ty = U.lookup_subst n subst in 
+        with U.Subst_not_found _ -> try
+         let _,_,ty = U.lookup_meta n metasenv in 
+         match ty with C.Implicit _ -> assert false | _ -> c,ty 
+        with U.Meta_not_found _ ->
+         raise (AssertFailure (lazy (Printf.sprintf
+          "%s not found" (PP.ppterm ~subst ~metasenv ~context t))))
+       in
+       metasenv, subst, t, S.subst_meta l ty
+    | C.Const _ -> 
+       metasenv, subst, t, NCicTypeChecker.typeof ~subst ~metasenv context t
+    | C.Prod (name,s,t) as orig ->
+       let metasenv, subst, s, s1 = typeof_aux metasenv subst context None s in
+       let context = (name,(NCic.Decl s))::context in
+       let metasenv, subst, t, s2 = typeof_aux metasenv subst context None t in
+       let metasenv, subst, ty = 
+         sort_of_prod localize metasenv subst context (name,s) t (s1,s2)
+       in
+       metasenv, subst, orig, ty
+    | C.Lambda (n,s,t) ->
+       let sort = typeof_aux context s in
+       (match R.whd ~subst context sort with
+       | C.Meta _ | C.Sort _ -> ()
+       | _ ->
+         raise
+           (TypeCheckerFailure (lazy (Printf.sprintf
+             ("Not well-typed lambda-abstraction: " ^^
+             "the source %s should be a type; instead it is a term " ^^ 
+             "of type %s") (PP.ppterm ~subst ~metasenv ~context s)
+             (PP.ppterm ~subst ~metasenv ~context sort)))));
+       let ty = typeof_aux ((n,(C.Decl s))::context) t in
+         C.Prod (n,s,ty)
+    | C.LetIn (n,ty,t,bo) ->
+       let ty_t = typeof_aux context t in
+       let _ = typeof_aux context ty in
+       if not (R.are_convertible ~subst context ty_t ty) then
+         raise 
+          (TypeCheckerFailure 
+            (lazy (Printf.sprintf
+              "The type of %s is %s but it is expected to be %s" 
+                (PP.ppterm ~subst ~metasenv ~context t) 
+                (PP.ppterm ~subst ~metasenv ~context ty_t) 
+                (PP.ppterm ~subst ~metasenv ~context ty))))
+       else
+         let ty_bo = typeof_aux  ((n,C.Def (t,ty))::context) bo in
+         S.subst ~avoid_beta_redexes:true t ty_bo
+    | C.Appl (he::(_::_ as args)) ->
+       let ty_he = typeof_aux context he in
+       let args_with_ty = List.map (fun t -> t, typeof_aux context t) args in
+       eat_prods ~subst ~metasenv context he ty_he args_with_ty
+   | C.Appl _ -> raise (AssertFailure (lazy "Appl of length < 2"))
+   | C.Match (Ref.Ref (_,Ref.Ind (_,tyno,_)) as r,outtype,term,pl) ->
+      let outsort = typeof_aux context outtype in
+      let _,leftno,itl,_,_ = E.get_checked_indtys r in
+      let constructorsno =
+        let _,_,_,cl = List.nth itl tyno in List.length cl
+      in
+      let parameters, arguments =
+        let ty = R.whd ~subst context (typeof_aux context term) in
+        let r',tl =
+         match ty with
+            C.Const (Ref.Ref (_,Ref.Ind _) as r') -> r',[]
+          | C.Appl (C.Const (Ref.Ref (_,Ref.Ind _) as r') :: tl) -> r',tl
+          | _ ->
+             raise 
+               (TypeCheckerFailure (lazy (Printf.sprintf
+                 "Case analysis: analysed term %s is not an inductive one"
+                 (PP.ppterm ~subst ~metasenv ~context term)))) in
+        if not (Ref.eq r r') then
+         raise
+          (TypeCheckerFailure (lazy (Printf.sprintf
+            ("Case analysys: analysed term type is %s, but is expected " ^^
+             "to be (an application of) %s")
+            (PP.ppterm ~subst ~metasenv ~context ty) 
+            (PP.ppterm ~subst ~metasenv ~context (C.Const r')))))
+        else
+         try HExtlib.split_nth leftno tl
+         with
+          Failure _ ->
+           raise (TypeCheckerFailure (lazy (Printf.sprintf 
+           "%s is partially applied" 
+           (PP.ppterm  ~subst ~metasenv ~context ty)))) in
+      (* let's control if the sort elimination is allowed: [(I q1 ... qr)|B] *)
+      let sort_of_ind_type =
+        if parameters = [] then C.Const r
+        else C.Appl ((C.Const r)::parameters) in
+      let type_of_sort_of_ind_ty = typeof_aux context sort_of_ind_type in
+      check_allowed_sort_elimination ~subst ~metasenv r context
+       sort_of_ind_type type_of_sort_of_ind_ty outsort;
+      (* let's check if the type of branches are right *)
+      if List.length pl <> constructorsno then
+       raise (TypeCheckerFailure (lazy ("Wrong number of cases in a match")));
+      let j,branches_ok,p_ty, exp_p_ty =
+        List.fold_left
+          (fun (j,b,old_p_ty,old_exp_p_ty) p ->
+            if b then
+              let cons = 
+                let cons = Ref.mk_constructor j r in
+                if parameters = [] then C.Const cons
+                else C.Appl (C.Const cons::parameters)
+              in
+              let ty_p = typeof_aux context p in
+              let ty_cons = typeof_aux context cons in
+              let ty_branch = 
+                type_of_branch ~subst context leftno outtype cons ty_cons 0 
+              in
+              j+1, R.are_convertible ~subst context ty_p ty_branch,
+              ty_p, ty_branch
+            else
+              j,false,old_p_ty,old_exp_p_ty
+          ) (1,true,C.Sort C.Prop,C.Sort C.Prop) pl
+      in
+      if not branches_ok then
+        raise
+         (TypeCheckerFailure 
+          (lazy (Printf.sprintf ("Branch for constructor %s :=\n%s\n"^^
+          "has type %s\nnot convertible with %s") 
+          (PP.ppterm  ~subst ~metasenv ~context
+            (C.Const (Ref.mk_constructor (j-1) r)))
+          (PP.ppterm ~metasenv ~subst ~context (List.nth pl (j-2)))
+          (PP.ppterm ~metasenv ~subst ~context p_ty) 
+          (PP.ppterm ~metasenv ~subst ~context exp_p_ty)))); 
+      let res = outtype::arguments@[term] in
+      R.head_beta_reduce (C.Appl res)
+    | C.Match _ -> assert false
+
+  and type_of_branch ~subst context leftno outty cons tycons liftno = 
+    match R.whd ~subst context tycons with
+    | C.Const (Ref.Ref (_,Ref.Ind _)) -> C.Appl [S.lift liftno outty ; cons]
+    | C.Appl (C.Const (Ref.Ref (_,Ref.Ind _))::tl) ->
+        let _,arguments = HExtlib.split_nth leftno tl in
+        C.Appl (S.lift liftno outty::arguments@[cons])
+    | C.Prod (name,so,de) ->
+        let cons =
+         match S.lift 1 cons with
+         | C.Appl l -> C.Appl (l@[C.Rel 1])
+         | t -> C.Appl [t ; C.Rel 1]
+        in
+         C.Prod (name,so,
+           type_of_branch ~subst ((name,(C.Decl so))::context) 
+            leftno outty cons de (liftno+1))
+    | _ -> raise (AssertFailure (lazy "type_of_branch"))
+
+  (* check_metasenv_consistency checks that the "canonical" context of a
+     metavariable is consitent - up to relocation via the relocation list l -
+     with the actual context *)
+  and check_metasenv_consistency 
+    ~subst ~metasenv term context canonical_context l 
+  =
+   match l with
+    | shift, C.Irl n ->
+       let context = snd (HExtlib.split_nth shift context) in
+        let rec compare = function
+         | 0,_,[] -> ()
+         | 0,_,_::_
+         | _,_,[] ->
+            raise (AssertFailure (lazy (Printf.sprintf
+             "Local and canonical context %s have different lengths"
+             (PP.ppterm ~subst ~context ~metasenv term))))
+         | m,[],_::_ ->
+            raise (TypeCheckerFailure (lazy (Printf.sprintf
+             "Unbound variable -%d in %s" m 
+             (PP.ppterm ~subst ~metasenv ~context term))))
+         | m,t::tl,ct::ctl ->
+            (match t,ct with
+                (_,C.Decl t1), (_,C.Decl t2)
+              | (_,C.Def (t1,_)), (_,C.Def (t2,_))
+              | (_,C.Def (_,t1)), (_,C.Decl t2) ->
+                 if not (R.are_convertible ~subst tl t1 t2) then
+                  raise 
+                      (TypeCheckerFailure 
+                        (lazy (Printf.sprintf 
+                      ("Not well typed metavariable local context for %s: " ^^ 
+                       "%s expected, which is not convertible with %s")
+                      (PP.ppterm ~subst ~metasenv ~context term) 
+                      (PP.ppterm ~subst ~metasenv ~context t2) 
+                      (PP.ppterm ~subst ~metasenv ~context t1))))
+              | _,_ ->
+               raise 
+                   (TypeCheckerFailure (lazy (Printf.sprintf 
+                    ("Not well typed metavariable local context for %s: " ^^ 
+                     "a definition expected, but a declaration found")
+                    (PP.ppterm ~subst ~metasenv ~context term)))));
+            compare (m - 1,tl,ctl)
+        in
+         compare (n,context,canonical_context)
+    | shift, lc_kind ->
+       (* we avoid useless lifting by shortening the context*)
+       let l,context = (0,lc_kind), snd (HExtlib.split_nth shift context) in
+       let lifted_canonical_context = 
+         let rec lift_metas i = function
+           | [] -> []
+           | (n,C.Decl t)::tl ->
+               (n,C.Decl (S.subst_meta l (S.lift i t)))::(lift_metas (i+1) tl)
+           | (n,C.Def (t,ty))::tl ->
+               (n,C.Def ((S.subst_meta l (S.lift i t)),
+                          S.subst_meta l (S.lift i ty)))::(lift_metas (i+1) tl)
+         in
+          lift_metas 1 canonical_context in
+       let l = U.expand_local_context lc_kind in
+       try
+        List.iter2 
+        (fun t ct -> 
+          match (t,ct) with
+          | t, (_,C.Def (ct,_)) ->
+             (*CSC: the following optimization is to avoid a possibly expensive
+                    reduction that can be easily avoided and that is quite
+                    frequent. However, this is better handled using levels to
+                    control reduction *)
+             let optimized_t =
+              match t with
+              | C.Rel n ->
+                  (try
+                    match List.nth context (n - 1) with
+                    | (_,C.Def (te,_)) -> S.lift n te
+                    | _ -> t
+                    with Failure _ -> t)
+              | _ -> t
+             in
+             if not (R.are_convertible ~subst context optimized_t ct)
+             then
+               raise 
+                 (TypeCheckerFailure 
+                   (lazy (Printf.sprintf 
+                     ("Not well typed metavariable local context: " ^^ 
+                      "expected a term convertible with %s, found %s")
+                     (PP.ppterm ~subst ~metasenv ~context ct) 
+                     (PP.ppterm ~subst ~metasenv ~context t))))
+          | t, (_,C.Decl ct) ->
+              let type_t = typeof_aux context t in
+              if not (R.are_convertible ~subst context type_t ct) then
+                raise (TypeCheckerFailure 
+                 (lazy (Printf.sprintf 
+                  ("Not well typed metavariable local context: "^^
+                  "expected a term of type %s, found %s of type %s") 
+                  (PP.ppterm ~subst ~metasenv ~context ct) 
+                  (PP.ppterm ~subst ~metasenv ~context t) 
+                  (PP.ppterm ~subst ~metasenv ~context type_t))))
+        ) l lifted_canonical_context 
+       with
+        Invalid_argument _ ->
+          raise (AssertFailure (lazy (Printf.sprintf
+           "Local and canonical context %s have different lengths"
+           (PP.ppterm ~subst ~metasenv ~context term))))
+
+  and check_allowed_sort_elimination ~subst ~metasenv r =
+   let mkapp he arg =
+     match he with
+     | C.Appl l -> C.Appl (l @ [arg])
+     | t -> C.Appl [t;arg] in
+   let rec aux context ind arity1 arity2 =
+    let arity1 = R.whd ~subst context arity1 in
+    let arity2 = R.whd ~subst context arity2 in
+      match arity1,arity2 with
+       | C.Prod (name,so1,de1), C.Prod (_,so2,de2) ->
+          if not (R.are_convertible ~subst context so1 so2) then
+           raise (TypeCheckerFailure (lazy (Printf.sprintf
+            "In outtype: expected %s, found %s"
+            (PP.ppterm ~subst ~metasenv ~context so1)
+            (PP.ppterm ~subst ~metasenv ~context so2)
+            )));
+          aux ((name, C.Decl so1)::context)
+           (mkapp (S.lift 1 ind) (C.Rel 1)) de1 de2
+       | C.Sort _, C.Prod (name,so,ta) ->
+          if not (R.are_convertible ~subst context so ind) then
+           raise (TypeCheckerFailure (lazy (Printf.sprintf
+            "In outtype: expected %s, found %s"
+            (PP.ppterm ~subst ~metasenv ~context ind)
+            (PP.ppterm ~subst ~metasenv ~context so)
+            )));
+          (match arity1, R.whd ~subst ((name,C.Decl so)::context) ta with
+            | (C.Sort C.Type _, C.Sort _)
+            | (C.Sort C.Prop, C.Sort C.Prop) -> ()
+            | (C.Sort C.Prop, C.Sort C.Type _) ->
+        (* TODO: we should pass all these parameters since we
+         * have them already *)
+                let _,leftno,itl,_,i = E.get_checked_indtys r in
+                let itl_len = List.length itl in
+                let _,itname,ittype,cl = List.nth itl i in
+                let cl_len = List.length cl in
+                 (* is it a singleton, non recursive and non informative
+                    definition or an empty one? *)
+                 if not
+                  (cl_len = 0 ||
+                   (itl_len = 1 && cl_len = 1 &&
+                    let _,_,constrty = List.hd cl in
+                      is_non_recursive_singleton r itname ittype constrty &&
+                      is_non_informative leftno constrty))
+                 then
+                  raise (TypeCheckerFailure (lazy
+                   ("Sort elimination not allowed")));
+          | _,_ -> ())
+       | _,_ -> ()
+   in
+    aux 
+
+ in 
+   typeof_aux context term
+
+and eat_prods ~subst ~metasenv context he ty_he args_with_ty = 
+  let rec aux ty_he = function 
+  | [] -> ty_he
+  | (arg, ty_arg)::tl ->
+      match R.whd ~subst context ty_he with 
+      | C.Prod (_,s,t) ->
+          if R.are_convertible ~subst context ty_arg s then
+            aux (S.subst ~avoid_beta_redexes:true arg t) tl
+          else
+            raise 
+              (TypeCheckerFailure 
+                (lazy (Printf.sprintf
+                  ("Appl: wrong application of %s: the parameter %s has type"^^
+                   "\n%s\nbut it should have type \n%s\nContext:\n%s\n")
+                  (PP.ppterm ~subst ~metasenv ~context he)
+                  (PP.ppterm ~subst ~metasenv ~context arg)
+                  (PP.ppterm ~subst ~metasenv ~context ty_arg)
+                  (PP.ppterm ~subst ~metasenv ~context s)
+                  (PP.ppcontext ~subst ~metasenv context))))
+       | _ ->
+          raise 
+            (TypeCheckerFailure
+              (lazy (Printf.sprintf
+                "Appl: %s is not a function, it cannot be applied"
+                (PP.ppterm ~subst ~metasenv ~context
+                 (let res = List.length tl in
+                  let eaten = List.length args_with_ty - res in
+                   (C.Appl
+                    (he::List.map fst
+                     (fst (HExtlib.split_nth eaten args_with_ty)))))))))
+  in
+    aux ty_he args_with_ty
+
+
 (*
 open Printf
 
@@ -1053,7 +1465,7 @@ and type_of_aux' ?(clean_dummy_dependent_types=true) ?(localization_tbl = Cic.Ci
       check_exp_named_subst_aux metasubst metasenv [] tl ugraph
 
 
-  and sort_of_prod localization_tbl subst metasenv context (name,s) t (t1, t2)
+  and sort_of_prod localize subst metasenv context (name,s) t (t1, t2)
    ugraph
   =
     let module C = Cic in
index 25227e550577ee41a42271737627b02359436303..cdd382206331c6e8a2490f927e2628ad32a855c8 100644 (file)
 
 (* $Id$ *)
 
+exception RefineFailure of (Stdpp.location * string) Lazy.t;;
+exception Uncertain of (Stdpp.location * string) Lazy.t;;
+exception AssertFailure of string Lazy.t;;
+
+val typeof :
+ ?localize:(NCic.term -> Stdpp.location) ->
+  NCic.metasenv -> NCic.substitution -> NCic.context -> 
+  NCic.term -> NCic.term option -> (* term, expected type *)
+    NCic.metasenv * NCic.substitution * NCic.term * NCic.term
+    (*  menv, subst,refined term, type *)
+
 (*
 exception RefineFailure of string Lazy.t;;
 exception Uncertain of string Lazy.t;;