X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fcomponents%2Fng_refiner%2FnCicRefiner.ml;h=15fcc069aaae4e3a589dbffb2e3d9abb49d788c6;hb=df0dc72bccac82b3dd69108b5996d7008d007601;hp=5727dae676ac51a9f0f545a11c834aa99ba0d384;hpb=14b8be2d1011ed155513aeda78b2e525d66db45f;p=helm.git diff --git a/matita/components/ng_refiner/nCicRefiner.ml b/matita/components/ng_refiner/nCicRefiner.ml index 5727dae67..15fcc069a 100644 --- a/matita/components/ng_refiner/nCicRefiner.ml +++ b/matita/components/ng_refiner/nCicRefiner.ml @@ -18,6 +18,12 @@ exception AssertFailure of string Lazy.t;; module C = NCic module Ref = NReference +type 'a expected_type = [ `XTNone (* unknown *) + | `XTSome of 'a (* the given term *) + | `XTSort (* any sort *) + | `XTInd (* any (co)inductive type *) + ] + let debug = ref false;; let indent = ref "";; let times = ref [];; @@ -62,15 +68,17 @@ let exp_implicit status ~localise metasenv subst context with_type t = | `Closed -> let metasenv,subst,expty = match with_type with - None -> metasenv,subst,None - | Some typ -> + `XTSort + | `XTInd + | `XTNone -> metasenv,subst,None + | `XTSome typ -> let (metasenv,subst),typ = try NCicMetaSubst.delift status ~unify:(fun m s c t1 t2 -> try Some (NCicUnification.unify status m s c t1 t2) with NCicUnification.UnificationFailure _ | NCicUnification.Uncertain _ -> None) - metasenv subst context 0 (0,NCic.Irl 0) typ + metasenv subst context (-1) (0,C.Irl 0) typ with NCicMetaSubst.MetaSubstFailure _ | NCicMetaSubst.Uncertain _ -> @@ -80,9 +88,14 @@ let exp_implicit status ~localise metasenv subst context with_type t = metasenv,subst,Some typ in NCicMetaSubst.mk_meta metasenv [] ?with_type:expty `IsTerm,subst - | `Type -> NCicMetaSubst.mk_meta metasenv context ?with_type `IsType,subst - | `Term -> NCicMetaSubst.mk_meta metasenv context ?with_type `IsTerm,subst + | `Type -> + let with_type = match with_type with `XTSome t -> Some t | _ -> None in + NCicMetaSubst.mk_meta metasenv context ?with_type `IsType,subst + | `Term -> + let with_type = match with_type with `XTSome t -> Some t | _ -> None in + NCicMetaSubst.mk_meta metasenv context ?with_type `IsTerm,subst | `Tagged s -> + let with_type = match with_type with `XTSome t -> Some t | _ -> None in NCicMetaSubst.mk_meta ~attrs:[`Name s] metasenv context ?with_type `IsTerm,subst | `Vector -> @@ -139,6 +152,42 @@ let check_allowed_sort_elimination status localise r orig = aux ;; +(* CSC: temporary thing, waiting for better times *) +let mk_fresh_name context name = +try + let rex = Str.regexp "[0-9']*$" in + let rex2 = Str.regexp "'*$" in + let basename = Str.global_replace rex "" in + let suffix name = + ignore (Str.search_forward rex name 0); + let n = Str.matched_string name in + let n = Str.global_replace rex2 "" n in + if n = "" then 0 else int_of_string n +in + let name' = basename name in + let name' = if name' = "_" then "H" else name' in + let last = + List.fold_left + (fun last (name,_) -> + if basename name = name' then + max last (suffix name) + else + last + ) (-1) context + in + name' ^ (if last = -1 then "" else string_of_int (last + 1)) +with exn -> prerr_endline ("XXX" ^ Printexc.to_string exn); assert false + +(* let eq, eq_refl = + let uri = NUri.uri_of_string "cic:/matita/ng/Plogic/equality/eq.ind" in + C.Const (Ref.reference_of_spec uri (Ref.Ind (true,0,2))), + C.Const (Ref.reference_of_spec uri Ref.Con (0,1,2)) +*) +let eq, eq_refl = + let uri = NUri.uri_of_string "cic:/matita/basics/jmeq/jmeq.ind" in + C.Const (Ref.reference_of_spec uri (Ref.Ind(true,0,2))), + C.Const (Ref.reference_of_spec uri (Ref.Con(0,1,2))) + let rec typeof (status:#NCicCoercion.status) ?(localise=fun _ -> Stdpp.dummy_loc) metasenv subst context term expty @@ -146,7 +195,7 @@ let rec typeof (status:#NCicCoercion.status) let force_ty skip_lambda skip_appl metasenv subst context orig t infty expty = (*D*)inside 'F'; try let rc = match expty with - | Some expty -> + | `XTSome expty -> (match t with | C.Implicit _ -> assert false | C.Lambda _ when skip_lambda -> metasenv, subst, t, expty @@ -166,15 +215,30 @@ let rec typeof (status:#NCicCoercion.status) | NCicUnification.Uncertain _ | NCicUnification.UnificationFailure _ as exc -> try_coercions status ~localise - metasenv subst context t orig infty expty true exc) - | None -> metasenv, subst, t, infty + metasenv subst context t orig infty (`XTSome expty) exc) + | `XTNone -> metasenv, subst, t, infty + | `XTSort -> + (match t with + | C.Implicit _ -> assert false + | _ -> + pp (lazy ("forcing infty=any sort: "^ + (status#ppterm ~metasenv ~subst ~context infty) ^ " === any sort")); + force_to_sort status metasenv subst context t orig localise infty) + | `XTInd -> + (match t with + | C.Implicit _ -> assert false + | _ -> + pp (lazy ("forcing infty=any (co)inductive type: "^ + (status#ppterm ~metasenv ~subst ~context infty) ^ " === any (co)inductive type")); + force_to_inductive status metasenv subst context t orig localise infty) (*D*)in outside true; rc with exc -> outside false; raise exc in let rec typeof_aux metasenv subst context expty = fun t as orig -> (*D*)inside 'R'; try let rc = pp (lazy (status#ppterm ~metasenv ~subst ~context t ^ " ::exp:: " ^ - match expty with None -> "None" | Some e -> + match expty with `XTSort -> "Any sort" | `XTInd -> "Any (co)inductive type" + | `XTNone -> "None" | `XTSome e -> status#ppterm ~metasenv ~subst ~context e)); let metasenv, subst, t, infty = match t with @@ -198,10 +262,9 @@ let rec typeof (status:#NCicCoercion.status) let (metasenv,_,t,ty),subst = exp_implicit status ~localise metasenv subst context expty t infos in - if expty = None then - typeof_aux metasenv subst context None t - else - metasenv, subst, t, ty + (match expty with + | `XTSome _ -> metasenv, subst, t, ty + | _ -> typeof_aux metasenv subst context expty t) | C.Meta (n,l) as t -> let metasenv, ty = try @@ -213,12 +276,12 @@ let rec typeof (status:#NCicCoercion.status) | C.Const _ -> metasenv, subst, t, NCicTypeChecker.typeof status ~subst ~metasenv context t | C.Prod (name,(s as orig_s),(t as orig_t)) -> - let metasenv, subst, s, s1 = typeof_aux metasenv subst context None s in + let metasenv, subst, s, s1 = typeof_aux metasenv subst context `XTSort s in let metasenv, subst, s, s1 = force_to_sort status metasenv subst context s orig_s localise s1 in let context1 = (name,(C.Decl s))::context in - let metasenv, subst, t, s2 = typeof_aux metasenv subst context1 None t in + let metasenv, subst, t, s2 = typeof_aux metasenv subst context1 `XTSort t in let metasenv, subst, t, s2 = force_to_sort status metasenv subst context1 t orig_t localise s2 in @@ -226,19 +289,21 @@ let rec typeof (status:#NCicCoercion.status) sort_of_prod status localise metasenv subst context orig_s orig_t (name,s) t (s1,s2) in - metasenv, subst, NCic.Prod(name,s,t), ty + metasenv, subst, C.Prod(name,s,t), ty | C.Lambda (n,(s as orig_s),t) as orig -> let exp_s, exp_ty_t, force_after = match expty with - | None -> None, None, false - | Some expty -> + | `XTSort + | `XTInd + | `XTNone -> `XTNone, `XTNone, false + | `XTSome expty -> match NCicReduction.whd status ~subst context expty with - | C.Prod (_,s,t) -> Some s, Some t, false - | _ -> None, None, true + | C.Prod (_,s,t) -> `XTSome s, `XTSome t, false + | _ -> `XTNone, `XTNone, true in let (metasenv,subst), s = match exp_s with - | Some exp_s -> + | `XTSome exp_s -> (* optimized case: implicit and implicitly typed meta * the optimization prevents proliferation of metas *) (match s with @@ -255,12 +320,14 @@ let rec typeof (status:#NCicCoercion.status) (try pp(lazy("Force source to: "^status#ppterm ~metasenv ~subst ~context exp_s)); - NCicUnification.unify ~test_eq_only:true status metasenv subst context s exp_s,s + NCicUnification.unify ~test_eq_only:true status metasenv subst context s exp_s, s with exc -> raise (wrap_exc (lazy (localise orig_s, Printf.sprintf "Source type %s was expected to be %s" (status#ppterm ~metasenv ~subst ~context s) (status#ppterm ~metasenv ~subst ~context exp_s))) exc))) - | None -> + | `XTNone + | `XTSort + | `XTInd -> let metasenv, subst, s = check_type status ~localise metasenv subst context s in (metasenv, subst), s @@ -278,18 +345,18 @@ let rec typeof (status:#NCicCoercion.status) let metasenv, subst, ty = check_type status ~localise metasenv subst context ty in let metasenv, subst, t, _ = - typeof_aux metasenv subst context (Some ty) t in + typeof_aux metasenv subst context (`XTSome ty) t in let context1 = (n, C.Def (t,ty)) :: context in let metasenv, subst, expty1 = match expty with - | None -> metasenv, subst, None - | Some x -> + | `XTSome x -> let m, s, x = NCicUnification.delift_type_wrt_terms status metasenv subst context1 (NCicSubstitution.lift status 1 x) [NCicSubstitution.lift status 1 t] in - m, s, Some x + m, s, `XTSome x + | _ -> metasenv, subst, expty in let metasenv, subst, bo, bo_ty = typeof_aux metasenv subst context1 expty1 bo @@ -303,13 +370,13 @@ let rec typeof (status:#NCicCoercion.status) in let refine_appl metasenv subst args = let metasenv, subst, he, ty_he = - typeof_aux metasenv subst context None he in + typeof_aux metasenv subst context `XTNone he in let metasenv, subst, t, ty = eat_prods status ~localise force_ty metasenv subst context expty t orig_he he ty_he args in metasenv, subst, hbr t, ty in - if args = [C.Implicit `Vector] && expty <> None then + if args = [C.Implicit `Vector] && expty <> `XTNone then (* we try here to expand the vector a 0 implicits, but we use * the expected type *) try @@ -321,10 +388,10 @@ let rec typeof (status:#NCicCoercion.status) (* CSC: whd can be useful on he too... *) (match he with C.Const (Ref.Ref (uri1,Ref.Con _)) -> ( - match - HExtlib.map_option (NCicReduction.whd status ~subst context) expty - with - Some (C.Appl(C.Const(Ref.Ref (uri2,Ref.Ind _) as ref)::expargs)) + match expty with + | `XTSome expty -> ( + match NCicReduction.whd status ~subst context expty with + C.Appl (C.Const (Ref.Ref (uri2,Ref.Ind _) as ref)::expargs) when NUri.eq uri1 uri2 -> (try let _,leftno,_,_,_ = NCicEnvironment.get_checked_indtys status ref in @@ -349,7 +416,7 @@ let rec typeof (status:#NCicCoercion.status) (C.Implicit `Term :: allargs) allexpargs) | arg::args -> let metasenv,subst,arg,_ = - typeof_aux metasenv subst context None arg in + typeof_aux metasenv subst context `XTNone arg in let metasenv,subst = NCicUnification.unify status metasenv subst context arg exparg @@ -361,7 +428,10 @@ let rec typeof (status:#NCicCoercion.status) | Sys.Break as exn -> raise exn | _ -> refine_appl metasenv subst args (* to try coercions *)) - | _ -> refine_appl metasenv subst args) + | _ -> refine_appl metasenv subst args) + | `XTNone + | `XTSort + | `XTInd -> refine_appl metasenv subst args) | _ -> refine_appl metasenv subst args) | C.Appl _ -> raise (AssertFailure (lazy "Appl of length < 2")) | C.Match (Ref.Ref (_,Ref.Ind (_,tyno,_)) as r, @@ -373,20 +443,20 @@ let rec typeof (status:#NCicCoercion.status) NCicMetaSubst.saturate status metasenv subst context arity 0 in let ind = if args = [] then C.Const r else C.Appl (C.Const r::args) in let metasenv, subst, term, _ = - typeof_aux metasenv subst context (Some ind) term in + typeof_aux metasenv subst context (`XTSome ind) term in let parameters, arguments = HExtlib.split_nth leftno args in let outtype = match outtype with | C.Implicit _ as ot -> let rec aux = function - | [] -> NCic.Lambda ("_",NCic.Implicit `Type,ot) - | _::tl -> NCic.Lambda ("_",NCic.Implicit `Type,aux tl) + | [] -> C.Lambda ("_",C.Implicit `Type,ot) + | _::tl -> C.Lambda ("_",C.Implicit `Type,aux tl) in aux arguments | _ -> outtype in let metasenv, subst, outtype, outsort = - typeof_aux metasenv subst context None outtype in + typeof_aux metasenv subst context `XTNone outtype in (* this cannot be `XTSort *) (* next lines are to do a subst-expansion of the outtype, so that when it becomes a beta-abstraction, the beta-redex is @@ -403,7 +473,7 @@ let rec typeof (status:#NCicCoercion.status) if parameters = [] then C.Const r else C.Appl ((C.Const r)::parameters) in let metasenv, subst, ind, ind_ty = - typeof_aux metasenv subst context None ind in + typeof_aux metasenv subst context `XTNone ind in (* FG: this cannot be `XTSort *) let metasenv, subst = check_allowed_sort_elimination status localise r orig_term metasenv subst context ind ind_ty outsort @@ -429,7 +499,7 @@ let rec typeof (status:#NCicCoercion.status) else C.Appl (C.Const cons::parameters) in let metasenv, subst, cons, ty_cons = - typeof_aux metasenv subst context None cons in + typeof_aux metasenv subst context `XTNone cons in (* FG: this cannot be `XTInd *) let ty_branch = NCicTypeChecker.type_of_branch status ~subst context leftno outtype cons ty_cons in @@ -437,7 +507,7 @@ let rec typeof (status:#NCicCoercion.status) status#ppterm ~metasenv ~subst ~context p ^ " ::inf:: " ^ status#ppterm ~metasenv ~subst ~context ty_branch )); let metasenv, subst, p, _ = - typeof_aux metasenv subst context (Some ty_branch) p in + typeof_aux metasenv subst context (`XTSome ty_branch) p in j-1, metasenv, subst, p :: branches) pl (List.length pl, metasenv, subst, []) in @@ -455,7 +525,7 @@ let rec typeof (status:#NCicCoercion.status) and check_type status ~localise metasenv subst context (ty as orig_ty) = let metasenv, subst, ty, sort = - typeof status ~localise metasenv subst context ty None + typeof status ~localise metasenv subst context ty `XTSort in let metasenv, subst, ty, _ = force_to_sort status metasenv subst context ty orig_ty localise sort @@ -463,26 +533,35 @@ and check_type status ~localise metasenv subst context (ty as orig_ty) = metasenv, subst, ty and try_coercions status - ~localise - metasenv subst context t orig_t infty expty perform_unification exc + ~localise metasenv subst context t orig_t infty expty exc = let cs_subst = NCicSubstitution.subst status ~avoid_beta_redexes:true in try - pp (lazy "WWW: trying coercions -- preliminary unification"); - let metasenv, subst = - NCicUnification.unify status metasenv subst context infty expty - in metasenv, subst, t, expty + (match expty with + `XTSome expty -> + pp (lazy "WWW: trying coercions -- preliminary unification"); + let metasenv, subst = + NCicUnification.unify status metasenv subst context infty expty + in metasenv, subst, t, expty + | _ -> raise (Failure "Special case XTProd, XTSort or XTInd")) with | exn -> (* we try with a coercion *) let rec first exc = function | [] -> - pp (lazy "WWW: no more coercions!"); - raise (wrap_exc (lazy (localise orig_t, Printf.sprintf + pp (lazy "WWW: no more coercions!"); + raise (wrap_exc (lazy + let expty = + match expty with + `XTSome expty -> status#ppterm ~metasenv ~subst ~context expty + | `XTSort -> "[[sort]]" + | `XTProd -> "[[prod]]" + | `XTInd -> "[[ind]]" in + (localise orig_t, Printf.sprintf "The term\n%s\nhas type\n%s\nbut is here used with type\n%s" (status#ppterm ~metasenv ~subst ~context t) (status#ppterm ~metasenv ~subst ~context infty) - (status#ppterm ~metasenv ~subst ~context expty))) exc) + expty)) exc) | (_,metasenv, newterm, newtype, meta)::tl -> try pp (lazy("K=" ^ status#ppterm ~metasenv ~subst ~context newterm)); @@ -494,21 +573,32 @@ and try_coercions status status#ppterm ~metasenv ~subst ~context t ^ " === " ^ status#ppterm ~metasenv ~subst ~context meta ^ "\n")); let metasenv, subst = - NCicUnification.unify status metasenv subst context t meta - in - pp (lazy ( "UNIFICATION in CTX:\n"^ - status#ppcontext ~metasenv ~subst context - ^ "\nMENV: " ^ - status#ppmetasenv metasenv ~subst - ^ "\nOF: " ^ - status#ppterm ~metasenv ~subst ~context newtype ^ " === " ^ - status#ppterm ~metasenv ~subst ~context expty ^ "\n")); - let metasenv, subst = - if perform_unification then - NCicUnification.unify status metasenv subst context newtype expty - else metasenv, subst - in - metasenv, subst, newterm, newtype + NCicUnification.unify status metasenv subst context t meta in + match expty with + `XTSome expty -> + pp (lazy ( "UNIFICATION in CTX:\n"^ + status#ppcontext ~metasenv ~subst context + ^ "\nMENV: " ^ + status#ppmetasenv metasenv ~subst + ^ "\nOF: " ^ + status#ppterm ~metasenv ~subst ~context newtype ^ " === " ^ + status#ppterm ~metasenv ~subst ~context expty ^ "\n")); + let metasenv,subst = + NCicUnification.unify status metasenv subst context newtype expty + in + metasenv, subst, newterm, newtype + | `XTSort -> + (match NCicReduction.whd status ~subst context newtype with + C.Sort _ -> metasenv,subst,newterm,newtype + | _ -> first exc tl) + | `XTInd -> + (match NCicReduction.whd status ~subst context newtype with + C.Const (Ref.Ref (_, Ref.Ind _)) -> metasenv,subst,newterm,newtype + | _ -> first exc tl) + | `XTProd -> + (match NCicReduction.whd status ~subst context newtype with + C.Prod _ -> metasenv,subst,newterm,newtype + | _ -> first exc tl) with | NCicUnification.UnificationFailure _ -> first exc tl | NCicUnification.Uncertain _ as exc -> first exc tl @@ -517,20 +607,21 @@ and try_coercions status try pp (lazy "WWW: trying coercions -- inner check"); match infty, expty, t with - | _,_, NCic.Match (Ref.Ref (_,Ref.Ind (_,tyno,leftno)) as r,outty,m,pl) -> + (* `XTSort|`XTProd|`XTInd + Match not implemented *) + | _,`XTSome expty, C.Match (Ref.Ref (_,Ref.Ind (_,tyno,leftno)) as r,outty,m,pl) -> (*{{{*) pp (lazy "CASE"); (* {{{ helper functions *) let get_cl_and_left_p refit outty = match refit with - | NReference.Ref (uri, NReference.Ind (_,tyno,leftno)) -> + | Ref.Ref (uri, Ref.Ind (_,tyno,leftno)) -> let _, leftno, itl, _, _ = NCicEnvironment.get_checked_indtys status r in let _, _, ty, cl = List.nth itl tyno in let constructorsno = List.length cl in let count_pis t = let rec aux ctx t = match NCicReduction.whd status ~subst ~delta:max_int ctx t with - | NCic.Prod (name,src,tgt) -> - let ctx = (name, (NCic.Decl src)) :: ctx in + | C.Prod (name,src,tgt) -> + let ctx = (name, (C.Decl src)) :: ctx in 1 + aux ctx tgt | _ -> 0 in @@ -539,17 +630,17 @@ and try_coercions status let rec skip_lambda_delifting t n = match t,n with | _,0 -> t - | NCic.Lambda (_,_,t),n -> + | C.Lambda (_,_,t),n -> pp (lazy ("WWW: failing term? "^ status#ppterm ~subst:[] ~metasenv:[] ~context:[] t)); skip_lambda_delifting (* substitute dangling indices with a dummy *) - (NCicSubstitution.subst status (NCic.Sort NCic.Prop) t) (n - 1) + (NCicSubstitution.subst status (C.Sort C.Prop) t) (n - 1) | _ -> assert false in let get_l_r_p n = function - | NCic.Lambda (_,NCic.Const (Ref.Ref (_,Ref.Ind (_,_,_))),_) -> [],[] - | NCic.Lambda (_,NCic.Appl - (NCic.Const (Ref.Ref (_,Ref.Ind (_,_,_))) :: args),_) -> + | C.Lambda (_,C.Const (Ref.Ref (_,Ref.Ind (_,_,_))),_) -> [],[] + | C.Lambda (_,C.Appl + (C.Const (Ref.Ref (_,Ref.Ind (_,_,_))) :: args),_) -> HExtlib.split_nth n args | _ -> assert false in @@ -562,7 +653,7 @@ and try_coercions status (fun ty -> List.fold_left (fun t p -> match t with - | NCic.Prod (_,_,t) -> + | C.Prod (_,_,t) -> cs_subst p t | _-> assert false) ty left_p) @@ -572,11 +663,12 @@ and try_coercions status cl, left_p, leftno, rno | _ -> raise exn in + (*{{{*) pp (lazy "CASE"); let rec keep_lambdas_and_put_expty ctx t bo right_p matched n = match t,n with | _,0 -> let rec mkr n = function - | [] -> [] | _::tl -> NCic.Rel n :: mkr (n+1) tl + | [] -> [] | _::tl -> C.Rel n :: mkr (n+1) tl in pp (lazy ("input replace: " ^ status#ppterm ~context:ctx ~metasenv:[] ~subst:[] bo)); let bo = @@ -584,37 +676,33 @@ and try_coercions status ~equality:(fun _ -> NCicRefineUtil.alpha_equivalence status) ~context:ctx ~what:(matched::right_p) - ~with_what:(NCic.Rel 1::List.rev (mkr 2 right_p)) + ~with_what:(C.Rel 1::List.rev (mkr 2 right_p)) ~where:bo in pp (lazy ("output replace: " ^ status#ppterm ~context:ctx ~metasenv:[] ~subst:[] bo)); bo - | NCic.Lambda (name, src, tgt),_ -> - NCic.Lambda (name, src, + | C.Lambda (name, src, tgt),_ -> + C.Lambda (name, src, keep_lambdas_and_put_expty - ((name, NCic.Decl src)::ctx) tgt (NCicSubstitution.lift status 1 bo) + ((name, C.Decl src)::ctx) tgt (NCicSubstitution.lift status 1 bo) (List.map (NCicSubstitution.lift status 1) right_p) (NCicSubstitution.lift status 1 matched) (n-1)) | _ -> assert false in - (*let eq = NCic.Const (NReference.reference_of_string ("cic:/matita/ng/Plogic/equality/eq.ind(1,0,2)")) in - let eq_refl = NCic.Const (NReference.reference_of_string ("cic:/matita/ng/Plogic/equality/eq.con(0,1,2)")) in*) - let eq = NCic.Const (NReference.reference_of_string ("cic:/matita/basics/jmeq/jmeq.ind(1,0,2)")) in - let eq_refl = NCic.Const (NReference.reference_of_string ("cic:/matita/basics/jmeq/jmeq.con(0,1,2)")) in let add_params metasenv subst context cty outty mty m i = let rec aux context outty par j mty m = function - | NCic.Prod (name, src, tgt) -> + | C.Prod (name, src, tgt) -> let t,k = aux - ((name, NCic.Decl src) :: context) - (NCicSubstitution.lift status 1 outty) (NCic.Rel j::par) (j+1) + ((name, C.Decl src) :: context) + (NCicSubstitution.lift status 1 outty) (C.Rel j::par) (j+1) (NCicSubstitution.lift status 1 mty) (NCicSubstitution.lift status 1 m) tgt in - NCic.Prod (name, src, t), k - | NCic.Const (Ref.Ref (_,Ref.Ind (_,_,leftno)) as r) -> + C.Prod (name, src, t), k + | C.Const (Ref.Ref (_,Ref.Ind (_,_,leftno)) as r) -> let k = - let k = NCic.Const(Ref.mk_constructor i r) in + let k = C.Const(Ref.mk_constructor i r) in NCicUntrusted.mk_appl k par in (* the type has no parameters, so kty = mty @@ -622,30 +710,30 @@ and try_coercions status try NCicTypechecker.typeof ~subst ~metasenv context k with NCicTypeChecker.TypeCheckerFailure _ -> assert false in *) - NCic.Prod ("p", - NCic.Appl [eq; mty; m; mty; k], + C.Prod ("p", + C.Appl [eq; mty; m; mty; k], (NCicSubstitution.lift status 1 (NCicReduction.head_beta_reduce status ~delta:max_int (NCicUntrusted.mk_appl outty [k])))),[mty,m,mty,k] - | NCic.Appl (NCic.Const (Ref.Ref (_,Ref.Ind (_,_,leftno)) as r)::pl) -> + | C.Appl (C.Const (Ref.Ref (_,Ref.Ind (_,_,leftno)) as r)::pl) -> let left_p,right_p = HExtlib.split_nth leftno pl in let has_rights = right_p <> [] in let k = - let k = NCic.Const(Ref.mk_constructor i r) in + let k = C.Const(Ref.mk_constructor i r) in NCicUntrusted.mk_appl k (left_p@par) in let right_p = try match NCicTypeChecker.typeof status ~subst ~metasenv context k with - | NCic.Appl (NCic.Const (Ref.Ref (_,Ref.Ind (_,_,_)))::args) -> + | C.Appl (C.Const (Ref.Ref (_,Ref.Ind (_,_,_)))::args) -> snd (HExtlib.split_nth leftno args) | _ -> assert false with NCicTypeChecker.TypeCheckerFailure _-> assert false in let orig_right_p = match mty with - | NCic.Appl (NCic.Const (Ref.Ref (_,Ref.Ind (_,_,_)))::args) -> + | C.Appl (C.Const (Ref.Ref (_,Ref.Ind (_,_,_)))::args) -> snd (HExtlib.split_nth leftno args) | _ -> assert false in @@ -659,7 +747,7 @@ and try_coercions status try NCicTypeChecker.typeof status ~subst ~metasenv context y with NCicTypeChecker.TypeCheckerFailure _ -> assert false in - NCic.Prod ("_", NCic.Appl [eq;xty;x;yty;y], + C.Prod ("_", C.Appl [eq;xty;x;yty;y], NCicSubstitution.lift status 1 tacc), (xty,x,yty,y)::pacc) (orig_right_p @ [m]) (right_p @ [k]) (NCicReduction.head_beta_reduce status ~delta:max_int @@ -667,13 +755,13 @@ and try_coercions status (* if has_rights then NCicReduction.head_beta_reduce status ~delta:max_int - (NCic.Appl (outty ::right_p @ [k])),k + (C.Appl (outty ::right_p @ [k])),k else - NCic.Prod ("p", - NCic.Appl [eq; mty; m; k], + C.Prod ("p", + C.Appl [eq; mty; m; k], (NCicSubstitution.lift status 1 (NCicReduction.head_beta_reduce status ~delta:max_int - (NCic.Appl (outty ::right_p @ [k]))))),k*) + (C.Appl (outty ::right_p @ [k]))))),k*) | _ -> assert false in aux context outty [] 1 mty m cty @@ -682,32 +770,32 @@ and try_coercions status (* k lives in the right context *) (* if rno <> 0 then pbo else *) let rec aux = function - | NCic.Lambda (name,src,bo), NCic.Prod (_,_,ty) -> - NCic.Lambda (name,src, + | C.Lambda (name,src,bo), C.Prod (_,_,ty) -> + C.Lambda (name,src, (aux (bo,ty))) | t,_ -> snd (List.fold_right - (fun (xty,x,yty,y) (n,acc) -> n+1, NCic.Lambda ("p" ^ string_of_int n, - NCic.Appl [eq; xty; x; yty; y], + (fun (xty,x,yty,y) (n,acc) -> n+1, C.Lambda ("p" ^ string_of_int n, + C.Appl [eq; xty; x; yty; y], NCicSubstitution.lift status 1 acc)) eqs (1,t)) - (*NCic.Lambda ("p", - NCic.Appl [eq; mty; m; k],NCicSubstitution.lift status 1 t)*) + (*C.Lambda ("p", + C.Appl [eq; mty; m; k],NCicSubstitution.lift status 1 t)*) in aux (pbo,cty) in let add_pi_for_refl_m context new_outty mty m lno rno = (*if rno <> 0 then new_outty else*) let rec aux context mty m = function - | NCic.Lambda (name, src, tgt) -> - let context = (name, NCic.Decl src)::context in - NCic.Lambda (name, src, aux context (NCicSubstitution.lift status 1 mty) (NCicSubstitution.lift status 1 m) tgt) + | C.Lambda (name, src, tgt) -> + let context = (name, C.Decl src)::context in + C.Lambda (name, src, aux context (NCicSubstitution.lift status 1 mty) (NCicSubstitution.lift status 1 m) tgt) | t -> let lhs = match mty with - | NCic.Appl (_::params) -> (snd (HExtlib.split_nth lno params))@[m] + | C.Appl (_::params) -> (snd (HExtlib.split_nth lno params))@[m] | _ -> [m] in let rhs = - List.map (fun x -> NCic.Rel x) + List.map (fun x -> C.Rel x) (List.rev (HExtlib.list_seq 1 (rno+2))) in List.fold_right2 (fun x y acc -> @@ -719,12 +807,12 @@ and try_coercions status try NCicTypeChecker.typeof status ~subst ~metasenv context y with NCicTypeChecker.TypeCheckerFailure _ -> assert false in - NCic.Prod - ("_", NCic.Appl [eq;xty;x;yty;y], + C.Prod + ("_", C.Appl [eq;xty;x;yty;y], (NCicSubstitution.lift status 1 acc))) lhs rhs t - (* NCic.Prod - ("_", NCic.Appl [eq;mty;m;NCic.Rel 1], + (* C.Prod + ("_", C.Appl [eq;mty;m;C.Rel 1], NCicSubstitution.lift status 1 t)*) in aux context mty m new_outty @@ -739,8 +827,8 @@ and try_coercions status match NCicTypeChecker.typeof status ~subst ~metasenv context m with - | (NCic.Const (Ref.Ref (_,Ref.Ind (_,_,_))) | NCic.Meta _) as mty -> [], mty - | NCic.Appl ((NCic.Const (Ref.Ref (_,Ref.Ind (_,_,_)))|NCic.Meta _)::args) as mty -> + | (C.Const (Ref.Ref (_,Ref.Ind (_,_,_))) | C.Meta _) as mty -> [], mty + | C.Appl ((C.Const (Ref.Ref (_,Ref.Ind (_,_,_)))|C.Meta _)::args) as mty -> snd (HExtlib.split_nth leftno args), mty | _ -> assert false with NCicTypeChecker.TypeCheckerFailure _ -> @@ -777,7 +865,7 @@ and try_coercions status ~context ~metasenv ~subst pbo)); let metasenv, subst, pbo, _ = try_coercions status ~localise menv s context pbo - orig_t (*??*) infty_pbo expty_pbo perform_unification exc + orig_t (*??*) infty_pbo (`XTSome expty_pbo) exc in pp (lazy ("CASE: pbo2: " ^ status#ppterm @@ -797,47 +885,63 @@ and try_coercions status List.map (fun t -> NCicTypeChecker.typeof status ~subst ~metasenv context t) right_p in let eqs = - List.map2 (fun x y -> NCic.Appl[eq_refl;x;y]) (right_tys@[mty]) + List.map2 (fun x y -> C.Appl[eq_refl;x;y]) (right_tys@[mty]) (right_p@[m]) in - let t = NCic.Appl (NCic.Match(r,new_outty,m,pl) :: eqs) + let t = C.Appl (C.Match(r,new_outty,m,pl) :: eqs) in metasenv, subst, t, expty (*}}}*) - | NCic.Prod (nameprod, src, ty),NCic.Prod (_, src2, ty2), _ -> - let rec mk_fresh_name ctx firstch n = - let candidate = (String.make 1 firstch) ^ (string_of_int n) in - if (List.for_all (fun (s,_) -> s <> candidate) ctx) then candidate - else mk_fresh_name ctx firstch (n+1) + | _,`XTSome expty,C.LetIn(name,ty,t,bo) -> + pp (lazy "LETIN"); + let name' = mk_fresh_name context name in + let context_bo = (name', C.Def (t,ty)) :: context in + let metasenv, subst, bo2, _ = + try_coercions status ~localise metasenv subst context_bo + bo orig_t (NCicSubstitution.lift status 1 infty) + (`XTSome (NCicSubstitution.lift status 1 expty)) exc in + let coerced = C.LetIn (name',ty,t,bo2) in + pp (lazy ("LETIN: coerced = " ^ status#ppterm ~metasenv ~subst ~context coerced)); + metasenv, subst, coerced, expty + | C.Prod (nameprod, src, ty),`XTSome (C.Prod (_, src2, ty2) as expty), _ -> (*{{{*) pp (lazy "LAM"); pp (lazy ("LAM: t = " ^ status#ppterm ~metasenv ~subst ~context t)); - let name_con = mk_fresh_name context 'c' 0 - (*FreshNamesGenerator.mk_fresh_name - ~subst metasenv context ~typ:src2 Cic.Anonymous*) - in - let context_src2 = ((name_con, NCic.Decl src2) :: context) in + let namename = match t with C.Lambda (n,_,_) -> n | _ -> nameprod in + let name_con = mk_fresh_name context namename in + let context_src2 = ((name_con, C.Decl src2) :: context) in (* contravariant part: the argument of f:src->ty *) let metasenv, subst, rel1, _ = try_coercions status ~localise metasenv subst context_src2 - (NCic.Rel 1) orig_t (NCicSubstitution.lift status 1 src2) - (NCicSubstitution.lift status 1 src) perform_unification exc + (C.Rel 1) orig_t (NCicSubstitution.lift status 1 src2) + (`XTSome (NCicSubstitution.lift status 1 src)) exc in (* covariant part: the result of f(c x); x:src2; (c x):src *) let name_t, bo = match t with - | NCic.Lambda (n,_,bo) -> n, cs_subst rel1 (NCicSubstitution.lift_from status 2 1 bo) + | C.Lambda (n,_,bo) -> n, cs_subst rel1 (NCicSubstitution.lift_from status 2 1 bo) | _ -> name_con, NCicUntrusted.mk_appl (NCicSubstitution.lift status 1 t) [rel1] in (* we fix the possible dependency problem in the source ty *) let ty = cs_subst rel1 (NCicSubstitution.lift_from status 2 1 ty) in let metasenv, subst, bo, _ = try_coercions status ~localise metasenv subst context_src2 - bo orig_t ty ty2 perform_unification exc + bo orig_t ty (`XTSome ty2) exc in - let coerced = NCic.Lambda (name_t,src2, bo) in + let coerced = C.Lambda (name_t,src2, bo) in pp (lazy ("LAM: coerced = " ^ status#ppterm ~metasenv ~subst ~context coerced)); metasenv, subst, coerced, expty (*}}}*) | _ -> raise exc with exc2 -> + let expty = + match expty with + `XTSome expty -> expty + | `XTSort -> + C.Sort (C.Type + (match NCicEnvironment.get_universes () with + | x::_ -> x + | _ -> assert false)) + | `XTProd -> C.Prod ("_",C.Implicit `Type,C.Implicit `Type) + | `XTInd -> assert false(*CSC: was not handled, OCaml 4.0 complains. ??? *) + in pp(lazy("try_coercion " ^ status#ppterm ~metasenv ~subst ~context infty ^ " |---> " ^ status#ppterm ~metasenv ~subst ~context expty)); @@ -852,15 +956,44 @@ and force_to_sort status metasenv subst context t orig_t localise ty = metasenv, subst, t, ty with Failure _ -> + let msg = + (lazy (localise orig_t, + "The type of " ^ status#ppterm ~metasenv ~subst ~context t ^ + " is not a sort: " ^ status#ppterm ~metasenv ~subst ~context ty)) in let ty = NCicReduction.whd status ~subst context ty in + let exn = + if NCicUnification.could_reduce status ~subst context ty then + Uncertain msg + else + RefineFailure msg + in try_coercions status ~localise metasenv subst context - t orig_t ty (NCic.Sort (NCic.Type - (match NCicEnvironment.get_universes () with - | x::_ -> x - | _ -> assert false))) false - (Uncertain (lazy (localise orig_t, - "The type of " ^ status#ppterm ~metasenv ~subst ~context t ^ - " is not a sort: " ^ status#ppterm ~metasenv ~subst ~context ty))) + t orig_t ty `XTSort exn + +and force_to_inductive status metasenv subst context t orig_t localise ty = + try + let metasenv, subst, ty = + NCicUnification.indfy status (Failure "indfy") metasenv subst context ty in + metasenv, subst, t, ty + with + Failure _ -> + let msg = + (lazy (localise orig_t, + "The type of " ^ status#ppterm ~metasenv ~subst ~context t ^ + " is not a (co)inductive type: " ^ status#ppterm ~metasenv ~subst ~context ty)) in + let ty = NCicReduction.whd status ~subst context ty in +(* prerr_endline ("#### " ^ status#ppterm ~metasenv ~subst ~context ty); *) + let exn = + if NCicUnification.could_reduce status ~subst context ty then + Uncertain msg + else + RefineFailure msg + in + raise exn +(* FG: this should be as follows but the case `XTInd is not imp;emented yet + try_coercions status ~localise metasenv subst context + t orig_t ty `XTInd exn +*) and sort_of_prod status localise metasenv subst context orig_s orig_t (name,s) t (t1, t2) @@ -887,8 +1020,8 @@ and sort_of_prod status localise metasenv subst context orig_s orig_t (name,s) and guess_name status subst ctx ty = let aux initial = "#" ^ String.make 1 initial in match ty with - | C.Const (NReference.Ref (u,_)) - | C.Appl (C.Const (NReference.Ref (u,_)) :: _) -> + | C.Const (Ref.Ref (u,_)) + | C.Appl (C.Const (Ref.Ref (u,_)) :: _) -> aux (String.sub (NUri.name_of_uri u) 0 1).[0] | C.Prod _ -> aux 'f' | C.Meta (n,lc) -> @@ -908,15 +1041,16 @@ and eat_prods status ~localise force_ty metasenv subst context expty orig_t orig pp(lazy("FORCE FINAL APPL: " ^ status#ppterm ~metasenv ~subst ~context res ^ " of type " ^ status#ppterm ~metasenv ~subst ~context ty_he - ^ " to type " ^ match expty with None -> "None" | Some x -> + ^ " to type " ^ match expty with `XTSort -> "Any sort" | `XTInd -> "Any (co)inductive type" + | `XTNone -> "None" | `XTSome x -> status#ppterm ~metasenv ~subst ~context x)); (* whatever the term is, we force the type. in case of ((Lambda..) ?...) * the application may also be a lambda! *) force_ty false false metasenv subst context orig_t res ty_he expty - | NCic.Implicit `Vector::tl -> + | C.Implicit `Vector::tl -> let has_some_more_pis x = match NCicReduction.whd status ~subst context x with - | NCic.Meta _ | NCic.Appl (NCic.Meta _::_) -> false + | C.Meta _ | C.Appl (C.Meta _::_) -> false | _ -> true in (try @@ -926,7 +1060,7 @@ and eat_prods status ~localise force_ty metasenv subst context expty orig_t orig | RefineFailure _ as exc when has_some_more_pis ty_he -> (try aux metasenv subst args_so_far he ty_he - (NCic.Implicit `Term :: NCic.Implicit `Vector :: tl) + (C.Implicit `Term :: C.Implicit `Vector :: tl) with Uncertain msg | RefineFailure msg -> raise (wrap_exc msg exc)) | RefineFailure msg when not (has_some_more_pis ty_he) -> @@ -936,13 +1070,13 @@ and eat_prods status ~localise force_ty metasenv subst context expty orig_t orig match NCicReduction.whd status ~subst context ty_he with | C.Prod (_,s,t) -> let metasenv, subst, arg, _ = - typeof status ~localise metasenv subst context arg (Some s) in + typeof status ~localise metasenv subst context arg (`XTSome s) in let t = NCicSubstitution.subst status ~avoid_beta_redexes:true arg t in aux metasenv subst (arg :: args_so_far) he t tl | C.Meta _ | C.Appl (C.Meta _ :: _) as t -> let metasenv, subst, arg, ty_arg = - typeof status ~localise metasenv subst context arg None in + typeof status ~localise metasenv subst context arg `XTNone in let name = guess_name status subst context ty_arg in let metasenv, _, meta, _ = NCicMetaSubst.mk_meta metasenv @@ -951,7 +1085,7 @@ and eat_prods status ~localise force_ty metasenv subst context expty orig_t orig let flex_prod = C.Prod (name, ty_arg, meta) in (* next line grants that ty_args is a type *) let metasenv,subst, flex_prod, _ = - typeof status ~localise metasenv subst context flex_prod None in + typeof status ~localise metasenv subst context flex_prod `XTSort in (* pp (lazy ( "UNIFICATION in CTX:\n"^ status#ppcontext ~metasenv ~subst context @@ -978,7 +1112,7 @@ and eat_prods status ~localise force_ty metasenv subst context expty orig_t orig aux metasenv subst (arg :: args_so_far) he meta tl | C.Match (_,_,C.Meta _,_) | C.Match (_,_,C.Appl (C.Meta _ :: _),_) - | C.Appl (C.Const (NReference.Ref (_, NReference.Fix _)) :: _) -> + | C.Appl (C.Const (Ref.Ref (_, Ref.Fix _)) :: _) -> raise (Uncertain (lazy (localise orig_he, Printf.sprintf ("The term %s is here applied to %d arguments but expects " ^^ "only %d arguments") (status#ppterm ~metasenv ~subst ~context he) @@ -987,8 +1121,7 @@ and eat_prods status ~localise force_ty metasenv subst context expty orig_t orig let metasenv, subst, newhead, newheadty = try_coercions status ~localise metasenv subst context (NCicUntrusted.mk_appl he (List.rev args_so_far)) orig_he ty - (NCic.Prod ("_",NCic.Implicit `Term,NCic.Implicit `Term)) - false + `XTProd (RefineFailure (lazy (localise orig_he, Printf.sprintf ("The term %s is here applied to %d arguments but expects " ^^ "only %d arguments") (status#ppterm ~metasenv ~subst ~context he) @@ -1043,9 +1176,9 @@ let undebruijnate status inductive ref t rev_fl = (HExtlib.list_mapi (fun (_,_,rno,_,_,_) i -> let i = len - i - 1 in - NCic.Const - (if inductive then NReference.mk_fix i rno ref - else NReference.mk_cofix i ref)) + C.Const + (if inductive then Ref.mk_fix i rno ref + else Ref.mk_cofix i ref)) rev_fl) t ;; @@ -1062,7 +1195,7 @@ let typeof_obj match bo with | Some bo -> let metasenv, subst, bo, ty = - typeof status ~localise metasenv subst [] bo (Some ty) in + typeof status ~localise metasenv subst [] bo (`XTSome ty) in let height = (* XXX recalculate *) height in metasenv, subst, Some bo, ty, height | None -> metasenv, subst, None, ty, 0 @@ -1086,7 +1219,7 @@ let typeof_obj List.fold_left (fun (metasenv,subst,fl) (relevance,name,k,ty,dbo,localise) -> let metasenv, subst, dbo, ty = - typeof status ~localise metasenv subst types dbo (Some ty) + typeof status ~localise metasenv subst types dbo (`XTSome ty) in metasenv, subst, (relevance,name,k,ty,dbo)::fl) (metasenv, subst, []) rev_fl @@ -1097,9 +1230,9 @@ let typeof_obj (fun (relevance,name,k,ty,dbo) -> let bo = undebruijnate status inductive - (NReference.reference_of_spec uri - (if inductive then NReference.Fix (0,k,0) - else NReference.CoFix 0)) dbo rev_fl + (Ref.reference_of_spec uri + (if inductive then Ref.Fix (0,k,0) + else Ref.CoFix 0)) dbo rev_fl in relevance,name,k,ty,bo) fl @@ -1113,7 +1246,7 @@ let typeof_obj (fun (metasenv,subst,res,ctx) (relevance,n,ty,cl) -> let metasenv, subst, ty = check_type status ~localise metasenv subst [] ty in - metasenv,subst,(relevance,n,ty,cl)::res,(n,NCic.Decl ty)::ctx + metasenv,subst,(relevance,n,ty,cl)::res,(n,C.Decl ty)::ctx ) (metasenv,subst,[],[]) itl in let metasenv,subst,itl,_ = List.fold_left @@ -1177,9 +1310,9 @@ let typeof_obj with Invalid_argument "List.fold_left2" -> assert false in let metasenv, subst = let rec aux context (metasenv,subst) = function - | NCic.Meta _ -> metasenv, subst - | NCic.Implicit _ -> metasenv, subst - | NCic.Appl (NCic.Rel i :: args) as t + | C.Meta _ -> metasenv, subst + | C.Implicit _ -> metasenv, subst + | C.Appl (C.Rel i :: args) as t when i > List.length context - len -> let lefts, _ = HExtlib.split_nth leftno args in let ctxlen = List.length context in @@ -1188,7 +1321,7 @@ let typeof_obj (fun ((metasenv, subst),l) arg -> NCicUnification.unify status ~test_eq_only:true metasenv subst context arg - (NCic.Rel (ctxlen - len - l)), l+1 + (C.Rel (ctxlen - len - l)), l+1 ) ((metasenv, subst), 0) lefts in @@ -1236,18 +1369,18 @@ let typeof_obj NCicSubstitution.psubst status (fun i -> if i <= leftno then - NCic.Rel i + C.Rel i else - NCic.Const (NReference.reference_of_spec uri - (NReference.Ind (ind,relsno - i,leftno)))) + C.Const (Ref.reference_of_spec uri + (Ref.Ind (ind,relsno - i,leftno)))) (HExtlib.list_seq 1 (relsno+1)) te in let te = List.fold_right (fun (name,decl) te -> match decl with - NCic.Decl ty -> NCic.Prod (name,ty,te) - | NCic.Def (bo,ty) -> NCic.LetIn (name,ty,bo,te) + C.Decl ty -> C.Prod (name,ty,te) + | C.Def (bo,ty) -> C.LetIn (name,ty,bo,te) ) sx_context_te_rev te in metasenv,subst,(k_relev,n,te)::res