(***********************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) (* Lexer.add_token("",s)) constr_kw let _ = Options.v7 := false (* For Correctness syntax; doesn't work if in psyntax (freeze pb?) *) let _ = Lexer.add_token ("","!") let pair loc = Qualid (loc, Libnames.qualid_of_string "Coq.Init.Datatypes.pair") let mk_cast = function (c,(_,None)) -> c | (c,(_,Some ty)) -> CCast(join_loc (constr_loc c) (constr_loc ty), c, ty) let mk_lam = function ([],c) -> c | (bl,c) -> CLambdaN(constr_loc c, bl,c) let coerce_to_id c = match c with CRef(Ident(loc,id)) -> (loc,Name id) | CHole loc -> (loc,Anonymous) | _ -> error "ill-formed match type annotation" let match_bind_of_pat loc (oid,ty) = let l2 = match oid with None -> [] | Some x -> [([x],CHole loc)] in let l1 = match ty with None -> [] (* No: should lookup inductive arity! *) | Some (CApp(_,_,args)) -> (* parameters do not appear *) List.map (fun (c,_) -> ([coerce_to_id c],CHole loc)) args | _ -> error "ill-formed match type annotation" in l1@l2 let mk_match (loc,cil,rty,br) = (* let (lc,pargs) = List.split cil in let pr = match rty with None -> None | Some ty -> let idargs = (* TODO: not forget the list lengths for PP! *) List.flatten (List.map (match_bind_of_pat loc) pargs) in Some (CLambdaN(loc,idargs,ty)) in *) CCases(loc,(None,rty),cil,br) let mk_fixb (loc,id,bl,ann,body,(tloc,tyc)) = let n = match bl,ann with [([_],_)], None -> 0 | _, Some x -> let ids = List.map snd (List.flatten (List.map fst bl)) in (try list_index (snd x) ids - 1 with Not_found -> error "no such fix variable") | _ -> error "cannot find decreasing arg of fix" in let ty = match tyc with None -> CHole tloc | Some t -> CProdN(loc,bl,t) in (snd id,n,ty,CLambdaN(loc,bl,body)) let mk_cofixb (loc,id,bl,ann,body,(tloc,tyc)) = let _ = option_app (fun (aloc,_) -> Util.user_err_loc (aloc,"Constr:mk_cofixb", Pp.str"Annotation forbidden in cofix expression")) ann in let ty = match tyc with None -> CHole tloc | Some t -> CProdN(loc,bl,t) in (snd id,ty,CLambdaN(loc,bl,body)) let mk_fix(loc,kw,id,dcls) = if kw then let fb = List.map mk_fixb dcls in CFix(loc,id,fb) else let fb = List.map mk_cofixb dcls in CCoFix(loc,id,fb) let binder_constr = create_constr_entry (get_univ "constr") "binder_constr" let rec mkCProdN loc bll c = match bll with | LocalRawAssum ((loc1,_)::_ as idl,t) :: bll -> CProdN (loc,[idl,t],mkCProdN (join_loc loc1 loc) bll c) | LocalRawDef ((loc1,_) as id,b) :: bll -> CLetIn (loc,id,b,mkCProdN (join_loc loc1 loc) bll c) | [] -> c | LocalRawAssum ([],_) :: bll -> mkCProdN loc bll c let rec mkCLambdaN loc bll c = match bll with | LocalRawAssum ((loc1,_)::_ as idl,t) :: bll -> CLambdaN (loc,[idl,t],mkCLambdaN (join_loc loc1 loc) bll c) | LocalRawDef ((loc1,_) as id,b) :: bll -> CLetIn (loc,id,b,mkCLambdaN (join_loc loc1 loc) bll c) | [] -> c | LocalRawAssum ([],_) :: bll -> mkCLambdaN loc bll c GEXTEND Gram GLOBAL: binder_constr lconstr constr operconstr sort global constr_pattern Constr.ident binder_let tuple_constr; Constr.ident: [ [ id = Prim.ident -> id (* This is used in quotations and Syntax *) | id = METAIDENT -> id_of_string id ] ] ; global: [ [ r = Prim.reference -> r (* This is used in quotations *) | id = METAIDENT -> Ident (loc,id_of_string id) ] ] ; constr_pattern: [ [ c = constr -> c ] ] ; sort: [ [ "Set" -> RProp Pos | "Prop" -> RProp Null | "Type" -> RType None ] ] ; lconstr: [ [ c = operconstr -> c ] ] ; constr: [ [ c = operconstr LEVEL "9" -> c ] ] ; tuple_constr: [ [ (*c1 = tuple_constr; ","; c2 = tuple_constr -> CNotation (loc,"_ , _",[c1;c2]) | *) c = operconstr -> c ] ] ; operconstr: [ "200" RIGHTA [ c = binder_constr -> c ] | "100" RIGHTA [ c1 = operconstr; ":"; c2 = binder_constr -> CCast(loc,c1,c2) | c1 = operconstr; ":"; c2 = operconstr -> CCast(loc,c1,c2) ] | "80" RIGHTA [ c1 = operconstr; "->"; c2 = binder_constr -> CArrow(loc,c1,c2) | c1 = operconstr; "->"; c2 = operconstr -> CArrow(loc,c1,c2) ] | "40L" LEFTA [ "-"; c = operconstr -> CNotation(loc,"- _",[c]) ] | "10L" LEFTA [ f=operconstr; args=LIST1 appl_arg -> CApp(loc,(None,f),args) | "@"; f=global; args=LIST0 NEXT -> CAppExpl(loc,(None,f),args) ] | "9" [ ] | "1L" LEFTA [ c=operconstr; ".("; f=global; args=LIST0 appl_arg; ")" -> CApp(loc,(Some (List.length args+1),CRef f),args@[c,None]) | c=operconstr; ".("; "@"; f=global; args=LIST0 NEXT; ")" -> CAppExpl(loc,(Some (List.length args+1),f),args@[c]) | c=operconstr; "%"; key=IDENT -> CDelimiters (loc,key,c) ] | "0" [ c=atomic_constr -> c | c=match_constr -> c | "("; c=tuple_constr; ")" -> c ] ] ; binder_constr: [ [ "forall"; bl = binder_list; ","; c = operconstr LEVEL "200" -> mkCProdN loc bl c | "fun"; bl = binder_list; ty = type_cstr; "=>"; c = operconstr LEVEL "200" -> mkCLambdaN loc bl (mk_cast(c,ty)) | "let"; id=name; bl = LIST0 binder_let; ty = type_cstr; ":="; c1 = operconstr; "in"; c2 = operconstr LEVEL "200" -> let loc1 = match bl with | LocalRawAssum ((loc,_)::_,_)::_ -> loc | LocalRawDef ((loc,_),_)::_ -> loc | _ -> dummy_loc in CLetIn(loc,id,mkCLambdaN loc1 bl (mk_cast(c1,ty)),c2) | "let"; lb = ["("; l=LIST0 name SEP ","; ")" -> l | "()" -> []]; po = return_type; ":="; c1 = operconstr; "in"; c2 = operconstr LEVEL "200" -> CLetTuple (loc,List.map snd lb,po,c1,c2) | "if"; c=operconstr; po = return_type; "then"; b1=operconstr LEVEL "200"; "else"; b2=operconstr LEVEL "200" -> CIf (loc, c, po, b1, b2) | c=fix_constr -> c ] ] ; appl_arg: [ [ "@"; n=INT; ":="; c=constr -> (c,Some(int_of_string n)) | c=constr -> (c,None) ] ] ; atomic_constr: [ [ g=global -> CRef g | s=sort -> CSort(loc,s) | n=INT -> CNumeral (loc,Bignat.POS (Bignat.of_string n)) | IDENT "_" -> CHole loc (* | "?"; n=Prim.natural -> CPatVar(loc,(false,Pattern.patvar_of_int n)) ] ] *) | "?"; id=ident -> CPatVar(loc,(false,id)) ] ] ; fix_constr: [ [ kw=fix_kw; dcl=fix_decl -> let (_,n,_,_,_,_) = dcl in mk_fix(loc,kw,n,[dcl]) | kw=fix_kw; dcl1=fix_decl; "with"; dcls=LIST1 fix_decl SEP "with"; "for"; id=identref -> mk_fix(loc,kw,id,dcl1::dcls) ] ] ; fix_kw: [ [ "fix" -> true | "cofix" -> false ] ] ; fix_decl: [ [ id=identref; bl=LIST0 binder; ann=fixannot; ty=type_cstr; ":="; c=operconstr LEVEL "200" -> (loc,id,bl,ann,c,ty) ] ] ; fixannot: [ [ "{"; IDENT "struct"; id=name; "}" -> Some id | -> None ] ] ; match_constr: [ [ "match"; ci=LIST1 case_item SEP ","; ty=case_type; "with"; br=branches; "end" -> mk_match (loc,ci,ty,br) ] ] ; case_item: [ [ c=operconstr LEVEL "100"; p=pred_pattern -> match c,p with | CRef (Ident (_,id)), (Names.Anonymous,x) -> (c,(Name id,x)) | _ -> (c,p) ] ] ; pred_pattern: [ [ oid = ["as"; id=name -> snd id | -> Names.Anonymous]; ty = OPT ["in"; r=global; nal=LIST0 name -> (loc,r,List.map snd nal)] -> (oid,ty) ] ] ; case_type: [ [ ty = OPT [ "return"; c = operconstr LEVEL "100" -> c ] -> ty ] ] ; return_type: [ [ na = ["as"; id=name -> snd id | -> Names.Anonymous]; ty = case_type -> (na,ty) ] ] ; branches: [ [ OPT"|"; br=LIST0 eqn SEP "|" -> br ] ] ; eqn: [ [ pl = LIST1 pattern SEP ","; "=>"; rhs = lconstr -> (loc,pl,rhs) ] ] ; simple_pattern: [ [ r = Prim.reference -> CPatAtom (loc,Some r) | IDENT "_" -> CPatAtom (loc,None) | "("; p = lpattern; ")" -> p | n = bigint -> CPatNumeral (loc,n) ] ] ; pattern: [ [ p = pattern ; lp = LIST1 simple_pattern -> (match p with | CPatAtom (_, Some r) -> CPatCstr (loc, r, lp) | _ -> Util.user_err_loc (loc, "compound_pattern", Pp.str "Constructor expected")) | p = pattern; "as"; id = base_ident -> CPatAlias (loc, p, id) | c = pattern; "%"; key=IDENT -> CPatDelimiters (loc,key,c) | p = simple_pattern -> p ] ] ; lpattern: [ [ c = pattern -> c | p1=pattern; ","; p2=lpattern -> CPatCstr (loc, pair loc, [p1;p2]) ] ] ; binder_list: [ [ idl=LIST1 name; bl=LIST0 binder_let -> LocalRawAssum (idl,CHole loc)::bl | idl=LIST1 name; ":"; c=lconstr -> [LocalRawAssum (idl,c)] | "("; idl=LIST1 name; ":"; c=lconstr; ")"; bl=LIST0 binder_let -> LocalRawAssum (idl,c)::bl ] ] ; binder_let: [ [ id=name -> LocalRawAssum ([id],CHole loc) | "("; id=name; idl=LIST1 name; ":"; c=lconstr; ")" -> LocalRawAssum (id::idl,c) | "("; id=name; ":"; c=lconstr; ")" -> LocalRawAssum ([id],c) | "("; id=name; ":="; c=lconstr; ")" -> LocalRawDef (id,c) | "("; id=name; ":"; t=lconstr; ":="; c=lconstr; ")" -> LocalRawDef (id,CCast (join_loc (constr_loc t) loc,c,t)) ] ] ; binder: [ [ id=name -> ([id],CHole loc) | "("; idl=LIST1 name; ":"; c=lconstr; ")" -> (idl,c) ] ] ; type_cstr: [ [ c=OPT [":"; c=lconstr -> c] -> (loc,c) ] ] ; END;;