diff options
author | herbelin <herbelin@85f007b7-540e-0410-9357-904b9bb8a0f7> | 2006-10-09 16:11:01 +0000 |
---|---|---|
committer | herbelin <herbelin@85f007b7-540e-0410-9357-904b9bb8a0f7> | 2006-10-09 16:11:01 +0000 |
commit | 366fa1bdea12b522c98984f50428ef8aa48cf8d0 (patch) | |
tree | 4d0683375ec32d681e1e6e5e4788654d8281b2b1 | |
parent | c03b138c8c5e85ca1636465582c3242f50415a73 (diff) |
Notations:
- prise en compte des variables liées non liées par la notation (bug #1186),
- test pour affichage des notations aussi sur les sous-ensembles
des lieurs multiples (cf notation "#" dans output/Notations.v),
- extension, correction et uniformisation de quelques fonctions sur
les constr_expr et cases_pattern (avec incidences sur rawterm.ml,
parsing et contrib/interface).
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@9226 85f007b7-540e-0410-9357-904b9bb8a0f7
-rw-r--r-- | contrib/interface/dad.ml | 2 | ||||
-rw-r--r-- | interp/constrextern.ml | 14 | ||||
-rw-r--r-- | interp/constrintern.ml | 41 | ||||
-rw-r--r-- | interp/topconstr.ml | 243 | ||||
-rw-r--r-- | interp/topconstr.mli | 50 | ||||
-rw-r--r-- | parsing/g_constr.ml4 | 2 | ||||
-rw-r--r-- | parsing/ppconstr.ml | 2 | ||||
-rw-r--r-- | pretyping/rawterm.ml | 2 | ||||
-rw-r--r-- | pretyping/rawterm.mli | 11 | ||||
-rw-r--r-- | test-suite/output/Notations.out | 6 | ||||
-rw-r--r-- | test-suite/output/Notations.v | 17 |
11 files changed, 256 insertions, 134 deletions
diff --git a/contrib/interface/dad.ml b/contrib/interface/dad.ml index 578abc490..8096bc311 100644 --- a/contrib/interface/dad.ml +++ b/contrib/interface/dad.ml @@ -73,7 +73,7 @@ let rec map_subst (env :env) (subst:patvar_map) = function | CPatVar (_,(_,i)) -> let constr = List.assoc i subst in extern_constr false env constr - | x -> map_constr_expr_with_binders (map_subst env) (fun _ x -> x) subst x;; + | x -> map_constr_expr_with_binders (fun _ x -> x) (map_subst env) subst x;; let map_subst_tactic env subst = function | TacExtend (loc,("Rewrite" as x),[b;cbl]) -> diff --git a/interp/constrextern.ml b/interp/constrextern.ml index 600392811..d38662294 100644 --- a/interp/constrextern.ml +++ b/interp/constrextern.ml @@ -405,7 +405,7 @@ let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat = match availability_of_prim_token sc scopes with | None -> raise No_match | Some key -> - let loc = pattern_loc pat in + let loc = cases_pattern_loc pat in insert_pat_alias loc (insert_pat_delimiters loc (CPatPrim(loc,p)) key) na with No_match -> try @@ -754,7 +754,11 @@ and extern_typ (_,scopes) = and sub_extern inctx (_,scopes) = extern inctx (None,scopes) -and factorize_prod scopes vars aty = function +and factorize_prod scopes vars aty c = + try + if !Options.raw_print or !print_no_symbol then raise No_match; + ([],extern_symbol scopes vars c (uninterp_notations c)) + with No_match -> match c with | RProd (loc,(Name id as na),ty,c) when same aty (extern_typ scopes vars (anonymize_if_reserved na ty)) & not (occur_var_constr_expr id aty) (* avoid na in ty escapes scope *) @@ -762,7 +766,11 @@ and factorize_prod scopes vars aty = function ((loc,Name id)::nal,c) | c -> ([],extern_typ scopes vars c) -and factorize_lambda inctx scopes vars aty = function +and factorize_lambda inctx scopes vars aty c = + try + if !Options.raw_print or !print_no_symbol then raise No_match; + ([],extern_symbol scopes vars c (uninterp_notations c)) + with No_match -> match c with | RLambda (loc,na,ty,c) when same aty (extern_typ scopes vars (anonymize_if_reserved na ty)) & not (occur_name na aty) (* To avoid na in ty' escapes scope *) diff --git a/interp/constrintern.ml b/interp/constrintern.ml index 386f1f4a2..64637bd09 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -156,7 +156,7 @@ let loc_of_notation f loc args ntn = else snd (unloc (f (List.hd args))) let ntn_loc = loc_of_notation constr_loc -let patntn_loc = loc_of_notation cases_pattern_loc +let patntn_loc = loc_of_notation cases_pattern_expr_loc let dump_notation_location pos ((path,df),sc) = let rec next growing = @@ -361,8 +361,8 @@ let rec has_duplicate = function | x::l -> if List.mem x l then (Some x) else has_duplicate l let loc_of_lhs lhs = - join_loc (cases_pattern_loc (List.hd (List.hd lhs))) - (cases_pattern_loc (list_last (list_last lhs))) + join_loc (cases_pattern_expr_loc (List.hd (List.hd lhs))) + (cases_pattern_expr_loc (list_last (list_last lhs))) let check_linearity lhs ids = match has_duplicate ids with @@ -693,15 +693,20 @@ let extract_explicit_arg imps args = (**********************************************************************) (* Syntax extensions *) -let traverse_binder subst id (ids,tmpsc,scopes as env) = +let traverse_binder subst (renaming,(ids,tmpsc,scopes as env)) id = try - (* Binders bound in the notation are consider first-order object *) - (* and binders not bound in the notation do not capture variables *) - (* outside the notation *) + (* Binders bound in the notation are considered first-order objects *) let _,id' = coerce_to_id (fst (List.assoc id subst)) in - id', (Idset.add id' ids,tmpsc,scopes) + (renaming,(Idset.add id' ids,tmpsc,scopes)), id' with Not_found -> - id, env + (* Binders not bound in the notation do not capture variables *) + (* outside the notation (i.e. in the substitution) *) + let fvs1 = List.map (fun (_,(c,_)) -> free_vars_of_constr_expr c) subst in + let fvs2 = List.map snd renaming in + let fvs = List.flatten (List.map Idset.elements fvs1) @ fvs2 in + let id' = next_ident_away id fvs in + let renaming' = if id=id' then renaming else (id,id')::renaming in + (renaming',env), id' let decode_constrlist_value = function | CAppExpl (_,_,l) -> l @@ -711,7 +716,7 @@ let rec subst_iterator y t = function | RVar (_,id) as x -> if id = y then t else x | x -> map_rawconstr (subst_iterator y t) x -let rec subst_aconstr_in_rawconstr loc interp subst (ids,_,scopes as _env) = +let rec subst_aconstr_in_rawconstr loc interp subst (renaming,(ids,_,scopes)) = function | AVar id -> begin @@ -721,6 +726,9 @@ let rec subst_aconstr_in_rawconstr loc interp subst (ids,_,scopes as _env) = let (a,(scopt,subscopes)) = List.assoc id subst in interp (ids,scopt,subscopes@scopes) a with Not_found -> + try + RVar (loc,List.assoc id renaming) + with Not_found -> (* Happens for local notation joint with inductive/fixpoint defs *) RVar (loc,id) end @@ -729,27 +737,28 @@ let rec subst_aconstr_in_rawconstr loc interp subst (ids,_,scopes as _env) = (* All elements of the list are in scopes (scopt,subscopes) *) let (a,(scopt,subscopes)) = List.assoc x subst in let termin = - subst_aconstr_in_rawconstr loc interp subst (ids,None,scopes) - terminator in + subst_aconstr_in_rawconstr loc interp subst + (renaming,(ids,None,scopes)) terminator in let l = decode_constrlist_value a in List.fold_right (fun a t -> subst_iterator ldots_var t (subst_aconstr_in_rawconstr loc interp ((x,(a,(scopt,subscopes)))::subst) - (ids,None,scopes) iter)) + (renaming,(ids,None,scopes)) iter)) (if lassoc then List.rev l else l) termin with Not_found -> anomaly "Inconsistent substitution of recursive notation") | t -> rawconstr_of_aconstr_with_binders loc (traverse_binder subst) - (subst_aconstr_in_rawconstr loc interp subst) (ids,None,scopes) t + (subst_aconstr_in_rawconstr loc interp subst) + (renaming,(ids,None,scopes)) t let intern_notation intern (_,tmp_scope,scopes as env) loc ntn args = let ntn,args = contract_notation ntn args in let ((ids,c),df) = Notation.interp_notation loc ntn (tmp_scope,scopes) in if !dump then dump_notation_location (ntn_loc loc args ntn) df; let subst = List.map2 (fun (id,scl) a -> (id,(a,scl))) ids args in - subst_aconstr_in_rawconstr loc intern subst env c + subst_aconstr_in_rawconstr loc intern subst ([],env) c let set_type_scope (ids,tmp_scope,scopes) = (ids,Some Notation.type_scope,scopes) @@ -962,7 +971,7 @@ let internalise sigma globalenv env allow_soapp lvar c = let tm' = intern env tm in let ids,typ = match t with | Some t -> - let tids = names_of_cases_indtype t in + let tids = ids_of_cases_indtype t in let tids = List.fold_right Idset.add tids Idset.empty in let t = intern_type (tids,None,scopes) t in let loc,ind,l = match t with diff --git a/interp/topconstr.ml b/interp/topconstr.ml index 213b9a48e..bd36002eb 100644 --- a/interp/topconstr.ml +++ b/interp/topconstr.ml @@ -39,17 +39,24 @@ type aconstr = | ALetIn of name * aconstr * aconstr | ACases of aconstr option * (aconstr * (name * (inductive * int * name list) option)) list * - (identifier list * cases_pattern list * aconstr) list + (cases_pattern list * aconstr) list | ALetTuple of name list * (name * aconstr option) * aconstr * aconstr | AIf of aconstr * (name * aconstr option) * aconstr * aconstr | ASort of rawsort | AHole of Evd.hole_kind | APatVar of patvar | ACast of aconstr * cast_type * aconstr - -let name_app f e = function - | Name id -> let (id, e) = f id e in (e, Name id) - | Anonymous -> e,Anonymous + +(**********************************************************************) +(* Re-interpret a notation as a rawconstr, taking care of binders *) + +let rec cases_pattern_fold_map loc g e = function + | PatVar (_,na) -> + let e',na' = name_fold_map g e na in e', PatVar (loc,na') + | PatCstr (_,cstr,patl,na) -> + let e',na' = name_fold_map g e na in + let e',patl' = list_fold_map (cases_pattern_fold_map loc g) e patl in + e', PatCstr (loc,cstr,patl',na') let rec subst_rawvars l = function | RVar (_,id) as r -> (try List.assoc id l with Not_found -> r) @@ -67,32 +74,33 @@ let rawconstr_of_aconstr_with_binders loc g f e = function let outerl = (ldots_var,inner)::(if swap then [x,RVar(loc,y)] else []) in subst_rawvars outerl it | ALambda (na,ty,c) -> - let e,na = name_app g e na in RLambda (loc,na,f e ty,f e c) + let e,na = name_fold_map g e na in RLambda (loc,na,f e ty,f e c) | AProd (na,ty,c) -> - let e,na = name_app g e na in RProd (loc,na,f e ty,f e c) + let e,na = name_fold_map g e na in RProd (loc,na,f e ty,f e c) | ALetIn (na,b,c) -> - let e,na = name_app g e na in RLetIn (loc,na,f e b,f e c) + let e,na = name_fold_map g e na in RLetIn (loc,na,f e b,f e c) | ACases (rtntypopt,tml,eqnl) -> let e',tml' = List.fold_right (fun (tm,(na,t)) (e',tml') -> let e',t' = match t with | None -> e',None | Some (ind,npar,nal) -> let e',nal' = List.fold_right (fun na (e',nal) -> - let e',na' = name_app g e' na in e',na'::nal) nal (e',[]) in + let e',na' = name_fold_map g e' na in e',na'::nal) nal (e',[]) in e',Some (loc,ind,npar,nal') in - let e',na' = name_app g e' na in + let e',na' = name_fold_map g e' na in (e',(f e tm,(na',t'))::tml')) tml (e,[]) in - let fold id (idl,e) = let (id,e) = g id e in (id::idl,e) in - let eqnl' = List.map (fun (idl,pat,rhs) -> - let (idl,e) = List.fold_right fold idl ([],e) in - (loc,idl,pat,f e rhs)) eqnl in + let fold (idl,e) id = let (e,id) = g e id in ((id::idl,e),id) in + let eqnl' = List.map (fun (patl,rhs) -> + let ((idl,e),patl) = + list_fold_map (cases_pattern_fold_map loc fold) ([],e) patl in + (loc,idl,patl,f e rhs)) eqnl in RCases (loc,option_map (f e') rtntypopt,tml',eqnl') | ALetTuple (nal,(na,po),b,c) -> - let e,nal = list_fold_map (name_app g) e nal in - let e,na = name_app g e na in + let e,nal = list_fold_map (name_fold_map g) e nal in + let e,na = name_fold_map g e na in RLetTuple (loc,nal,(na,option_map (f e) po),f e b,f e c) | AIf (c,(na,po),b1,b2) -> - let e,na = name_app g e na in + let e,na = name_fold_map g e na in RIf (loc,f e c,(na,option_map (f e) po),f e b1,f e b2) | ACast (c,k,t) -> RCast (loc,f e c, k,f e t) | ASort x -> RSort (loc,x) @@ -102,17 +110,11 @@ let rawconstr_of_aconstr_with_binders loc g f e = function let rec rawconstr_of_aconstr loc x = let rec aux () x = - rawconstr_of_aconstr_with_binders loc (fun id () -> (id,())) aux () x + rawconstr_of_aconstr_with_binders loc (fun () id -> ((),id)) aux () x in aux () x -let rec subst_pat subst pat = - match pat with - | PatVar _ -> pat - | PatCstr (loc,((kn,i),j),cpl,n) -> - let kn' = subst_kn subst kn - and cpl' = list_smartmap (subst_pat subst) cpl in - if kn' == kn && cpl' == cpl then pat else - PatCstr (loc,((kn',i),j),cpl',n) +(****************************************************************************) +(* Translating a rawconstr into a notation, interpreting recursive patterns *) let add_name r = function | Anonymous -> () @@ -179,9 +181,7 @@ let aconstr_and_vars_of_rawconstr a = | RProd (_,na,ty,c) -> add_name found na; AProd (na,aux ty,aux c) | RLetIn (_,na,b,c) -> add_name found na; ALetIn (na,aux b,aux c) | RCases (_,rtntypopt,tml,eqnl) -> - let f (_,idl,pat,rhs) = - found := idl@(!found); - (idl,pat,aux rhs) in + let f (_,idl,pat,rhs) = found := idl@(!found); (pat,aux rhs) in ACases (option_map aux rtntypopt, List.map (fun (tm,(na,x)) -> add_name found na; @@ -246,9 +246,20 @@ let aconstr_of_rawconstr vars a = List.iter check_type vars; a +(* Substitution of kernel names, avoiding a list of bound identifiers *) + let aconstr_of_constr avoiding t = aconstr_of_rawconstr [] (Detyping.detype false avoiding [] t) +let rec subst_pat subst pat = + match pat with + | PatVar _ -> pat + | PatCstr (loc,((kn,i),j),cpl,n) -> + let kn' = subst_kn subst kn + and cpl' = list_smartmap (subst_pat subst) cpl in + if kn' == kn && cpl' == cpl then pat else + PatCstr (loc,((kn',i),j),cpl',n) + let rec subst_aconstr subst bound raw = match raw with | ARef ref -> @@ -265,22 +276,26 @@ let rec subst_aconstr subst bound raw = AApp(r',rl') | AList (id1,id2,r1,r2,b) -> - let r1' = subst_aconstr subst bound r1 and r2' = subst_aconstr subst bound r2 in + let r1' = subst_aconstr subst bound r1 + and r2' = subst_aconstr subst bound r2 in if r1' == r1 && r2' == r2 then raw else AList (id1,id2,r1',r2',b) | ALambda (n,r1,r2) -> - let r1' = subst_aconstr subst bound r1 and r2' = subst_aconstr subst bound r2 in + let r1' = subst_aconstr subst bound r1 + and r2' = subst_aconstr subst bound r2 in if r1' == r1 && r2' == r2 then raw else ALambda (n,r1',r2') | AProd (n,r1,r2) -> - let r1' = subst_aconstr subst bound r1 and r2' = subst_aconstr subst bound r2 in + let r1' = subst_aconstr subst bound r1 + and r2' = subst_aconstr subst bound r2 in if r1' == r1 && r2' == r2 then raw else AProd (n,r1',r2') | ALetIn (n,r1,r2) -> - let r1' = subst_aconstr subst bound r1 and r2' = subst_aconstr subst bound r2 in + let r1' = subst_aconstr subst bound r1 + and r2' = subst_aconstr subst bound r2 in if r1' == r1 && r2' == r2 then raw else ALetIn (n,r1',r2') @@ -295,11 +310,11 @@ let rec subst_aconstr subst bound raw = if a' == a && signopt' == signopt then x else (a',(n,signopt'))) rl and branches' = list_smartmap - (fun (idl,cpl,r as branch) -> + (fun (cpl,r as branch) -> let cpl' = list_smartmap (subst_pat subst) cpl and r' = subst_aconstr subst bound r in if cpl' == cpl && r' == r then branch else - (idl,cpl',r')) + (cpl',r')) branches in if rtntypopt' == rtntypopt && rtntypopt == rtntypopt' & @@ -331,7 +346,8 @@ let rec subst_aconstr subst bound raw = Evd.InternalHole | Evd.TomatchTypeParameter _) -> raw | ACast (r1,k,r2) -> - let r1' = subst_aconstr subst bound r1 and r2' = subst_aconstr subst bound r2 in + let r1' = subst_aconstr subst bound r1 + and r2' = subst_aconstr subst bound r2 in if r1' == r1 && r2' == r2 then raw else ACast (r1',k,r2') @@ -471,7 +487,7 @@ and match_binders alp metas na1 na2 sigma b1 b2 = let (alp,sigma) = match_names metas (alp,sigma) na1 na2 in match_ alp metas sigma b1 b2 -and match_equations alp metas sigma (_,_,patl1,rhs1) (_,patl2,rhs2) = +and match_equations alp metas sigma (_,_,patl1,rhs1) (patl2,rhs2) = (* patl1 and patl2 have the same length because they respectively correspond to some tml1 and tml2 that have the same length *) let (alp,sigma) = @@ -606,7 +622,7 @@ let constr_loc = function | CDelimiters (loc,_,_) -> loc | CDynamic _ -> dummy_loc -let cases_pattern_loc = function +let cases_pattern_expr_loc = function | CPatAlias (loc,_,_) -> loc | CPatCstr (loc,_,_) -> loc | CPatAtom (loc,_) -> loc @@ -619,35 +635,98 @@ let occur_var_constr_ref id = function | Ident (loc,id') -> id = id' | Qualid _ -> false -let rec occur_var_constr_expr id = function - | CRef r -> occur_var_constr_ref id r - | CArrow (loc,a,b) -> occur_var_constr_expr id a or occur_var_constr_expr id b - | CAppExpl (loc,(_,r),l) -> - occur_var_constr_ref id r or List.exists (occur_var_constr_expr id) l - | CApp (loc,(_,f),l) -> - occur_var_constr_expr id f or - List.exists (fun (a,_) -> occur_var_constr_expr id a) l - | CProdN (_,l,b) -> occur_var_binders id b l - | CLambdaN (_,l,b) -> occur_var_binders id b l - | CLetIn (_,na,a,b) -> occur_var_binders id b [[na],a] - | CCast (loc,a,_,b) -> - occur_var_constr_expr id a or occur_var_constr_expr id b - | CNotation (_,_,l) -> List.exists (occur_var_constr_expr id) l - | CDelimiters (loc,_,a) -> occur_var_constr_expr id a - | CHole _ | CEvar _ | CPatVar _ | CSort _ | CPrim _ | CDynamic _ -> false - | CCases (loc,_,_,_) - | CLetTuple (loc,_,_,_,_) - | CIf (loc,_,_,_,_) - | CFix (loc,_,_) +let ids_of_cases_indtype = + let add_var ids = function CRef (Ident (_,id)) -> id::ids | _ -> ids in + let rec vars_of = function + (* We deal only with the regular cases *) + | CApp (_,_,l) -> List.fold_left add_var [] (List.map fst l) + | CNotation (_,_,l) + (* assume the ntn is applicative and does not instantiate the head !! *) + | CAppExpl (_,_,l) -> List.fold_left add_var [] l + | CDelimiters(_,_,c) -> vars_of c + | _ -> [] in + vars_of + +let ids_of_cases_tomatch tms = + List.fold_right + (fun (_,(ona,indnal)) l -> + option_fold_right (fun t -> (@) (ids_of_cases_indtype t)) + indnal (option_fold_right name_cons ona l)) + tms [] + +let is_constructor id = + try ignore (Nametab.extended_locate (make_short_qualid id)); true + with Not_found -> true + +let rec cases_pattern_fold_names f a = function + | CPatAlias (_,pat,id) -> f id a + | CPatCstr (_,_,patl) | CPatOr (_,patl) | CPatNotation (_,_,patl) -> + List.fold_left (cases_pattern_fold_names f) a patl + | CPatDelimiters (_,_,pat) -> cases_pattern_fold_names f a pat + | CPatAtom (_,Some (Ident (_,id))) when not (is_constructor id) -> f id a + | CPatPrim _ | CPatAtom _ -> a + +let ids_of_pattern_list = + List.fold_left (List.fold_left (cases_pattern_fold_names Idset.add)) + Idset.empty + +let rec fold_constr_expr_binders g f n acc b = function + | (nal,t)::l -> + let nal = snd (List.split nal) in + let n' = List.fold_right (name_fold g) nal n in + f n (fold_constr_expr_binders g f n' acc b l) t + | [] -> + f n acc b + +let rec fold_local_binders g f n acc b = function + | LocalRawAssum (nal,t)::l -> + let nal = snd (List.split nal) in + let n' = List.fold_right (name_fold g) nal n in + f n (fold_local_binders g f n' acc b l) t + | LocalRawDef ((_,na),t)::l -> + f n (fold_local_binders g f (name_fold g na n) acc b l) t + | _ -> + f n acc b + +let fold_constr_expr_with_binders g f n acc = function + | CArrow (loc,a,b) -> f n (f n acc a) b + | CAppExpl (loc,(_,_),l) -> List.fold_left (f n) acc l + | CApp (loc,(_,t),l) -> List.fold_left (f n) (f n acc t) (List.map fst l) + | CProdN (_,l,b) | CLambdaN (_,l,b) -> fold_constr_expr_binders g f n acc b l + | CLetIn (_,na,a,b) -> fold_constr_expr_binders g f n acc b [[na],a] + | CCast (loc,a,_,b) -> f n (f n acc a) b + | CNotation (_,_,l) -> List.fold_left (f n) acc l + | CDelimiters (loc,_,a) -> f n acc a + | CHole _ | CEvar _ | CPatVar _ | CSort _ | CPrim _ | CDynamic _ | CRef _ -> + acc + | CCases (loc,rtnpo,al,bl) -> + let ids = ids_of_cases_tomatch al in + let acc = option_fold_left (f (List.fold_right g ids n)) acc rtnpo in + let acc = List.fold_left (f n) acc (List.map fst al) in + List.fold_right (fun (loc,patl,rhs) acc -> + let ids = ids_of_pattern_list patl in + f (Idset.fold g ids n) acc rhs) bl acc + | CLetTuple (loc,nal,(ona,po),b,c) -> + let n' = List.fold_right (name_fold g) nal n in + f (option_fold_right (name_fold g) ona n') (f n acc b) c + | CIf (_,c,(ona,po),b1,b2) -> + let acc = f n (f n (f n acc b1) b2) c in + option_fold_left (f (option_fold_right (name_fold g) ona n)) acc po + | CFix (loc,_,l) -> + let n' = List.fold_right (fun (id,_,_,_,_) -> g id) l n in + List.fold_right (fun (_,(_,o),lb,t,c) acc -> + fold_local_binders g f n' + (fold_local_binders g f n acc t lb) c lb) l acc | CCoFix (loc,_,_) -> - Pp.warning "Capture check in multiple binders not done"; false + Pp.warning "Capture check in multiple binders not done"; acc -and occur_var_binders id b = function - | (idl,a)::l -> - occur_var_constr_expr id a or - (not (List.mem (Name id) (snd (List.split idl))) - & occur_var_binders id b l) - | [] -> occur_var_constr_expr id b +let free_vars_of_constr_expr c = + let rec aux bdvars l = function + | CRef (Ident (_,id)) -> if List.mem id bdvars then l else Idset.add id l + | c -> fold_constr_expr_with_binders (fun a l -> a::l) aux bdvars l c + in aux [] Idset.empty c + +let occur_var_constr_expr id c = Idset.mem id (free_vars_of_constr_expr c) let mkIdentC id = CRef (Ident (dummy_loc, id)) let mkRefC r = CRef r @@ -679,19 +758,6 @@ let coerce_to_id = function (* Used in correctness and interface *) - -let names_of_cases_indtype = - let add_var ids = function CRef (Ident (_,id)) -> id::ids | _ -> ids in - let rec vars_of = function - (* We deal only with the regular cases *) - | CApp (_,_,l) -> List.fold_left add_var [] (List.map fst l) - | CNotation (_,_,l) - (* assume the ntn is applicative and does not instantiate the head !! *) - | CAppExpl (_,_,l) -> List.fold_left add_var [] l - | CDelimiters(_,_,c) -> vars_of c - | _ -> [] in - vars_of - let map_binder g e nal = List.fold_right (fun (_,na) -> name_fold g na) nal e let map_binders f g e bl = @@ -710,7 +776,7 @@ let map_local_binders f g e bl = let (e,rbl) = List.fold_left h (e,[]) bl in (e, List.rev rbl) -let map_constr_expr_with_binders f g e = function +let map_constr_expr_with_binders g f e = function | CArrow (loc,a,b) -> CArrow (loc,f e a,f e b) | CAppExpl (loc,r,l) -> CAppExpl (loc,r,List.map (f e) l) | CApp (loc,(p,a),l) -> @@ -728,18 +794,9 @@ let map_constr_expr_with_binders f g e = function | CCases (loc,rtnpo,a,bl) -> (* TODO: apply g on the binding variables in pat... *) let bl = List.map (fun (loc,pat,rhs) -> (loc,pat,f e rhs)) bl in - let e' = - List.fold_right - (fun (tm,(na,indnal)) e -> - option_fold_right - (fun t -> - let ids = names_of_cases_indtype t in - List.fold_right g ids) - indnal (option_fold_right (name_fold g) na e)) - a e - in - CCases (loc,option_map (f e') rtnpo, - List.map (fun (tm,x) -> (f e tm,x)) a,bl) + let ids = ids_of_cases_tomatch a in + let po = option_map (f (List.fold_right g ids e)) rtnpo in + CCases (loc, po, List.map (fun (tm,x) -> (f e tm,x)) a,bl) | CLetTuple (loc,nal,(ona,po),b,c) -> let e' = List.fold_right (name_fold g) nal e in let e'' = option_fold_right (name_fold g) ona e in @@ -767,8 +824,8 @@ let map_constr_expr_with_binders f g e = function let rec replace_vars_constr_expr l = function | CRef (Ident (loc,id)) as x -> (try CRef (Ident (loc,List.assoc id l)) with Not_found -> x) - | c -> map_constr_expr_with_binders replace_vars_constr_expr - (fun id l -> List.remove_assoc id l) l c + | c -> map_constr_expr_with_binders List.remove_assoc + replace_vars_constr_expr l c (**********************************************************************) (* Concrete syntax for modules and modules types *) diff --git a/interp/topconstr.mli b/interp/topconstr.mli index 85514c139..5ceaea67a 100644 --- a/interp/topconstr.mli +++ b/interp/topconstr.mli @@ -35,7 +35,7 @@ type aconstr = | ALetIn of name * aconstr * aconstr | ACases of aconstr option * (aconstr * (name * (inductive * int * name list) option)) list * - (identifier list * cases_pattern list * aconstr) list + (cases_pattern list * aconstr) list | ALetTuple of name list * (name * aconstr option) * aconstr * aconstr | AIf of aconstr * (name * aconstr option) * aconstr * aconstr | ASort of rawsort @@ -43,20 +43,36 @@ type aconstr = | APatVar of patvar | ACast of aconstr * cast_type * aconstr +(**********************************************************************) +(* Translate a rawconstr into a notation given the list of variables *) +(* bound by the notation; also interpret recursive patterns *) + +val aconstr_of_rawconstr : identifier list -> rawconstr -> aconstr + +(* Name of the special identifier used to encode recursive notations *) +val ldots_var : identifier + +(* Equality of rawconstr (warning: only partially implemented) *) +val eq_rawconstr : rawconstr -> rawconstr -> bool + +(**********************************************************************) +(* Re-interpret a notation as a rawconstr, taking care of binders *) + val rawconstr_of_aconstr_with_binders : loc -> - (identifier -> 'a -> identifier * 'a) -> + ('a -> identifier -> 'a * identifier) -> ('a -> aconstr -> rawconstr) -> 'a -> aconstr -> rawconstr val rawconstr_of_aconstr : loc -> aconstr -> rawconstr -val subst_aconstr : substitution -> Names.identifier list -> aconstr -> aconstr +(**********************************************************************) +(* Substitution of kernel names, avoiding a list of bound identifiers *) -val aconstr_of_rawconstr : identifier list -> rawconstr -> aconstr +val subst_aconstr : substitution -> identifier list -> aconstr -> aconstr -val eq_rawconstr : rawconstr -> rawconstr -> bool +(**********************************************************************) +(* [match_aconstr metas] matches a rawconstr against an aconstr with *) +(* metavariables in [metas]; raise [No_match] if the matching fails *) -(* [match_aconstr metas] match a rawconstr against an aconstr with - metavariables in [metas]; it raises [No_match] if the matching fails *) exception No_match type scope_name = string @@ -69,7 +85,8 @@ type interpretation = val match_aconstr : rawconstr -> interpretation -> (rawconstr * (tmp_scope_name option * scope_name list)) list -(*s Concrete syntax for terms *) +(**********************************************************************) +(*s Concrete syntax for terms *) type notation = string @@ -131,18 +148,21 @@ and local_binder = | LocalRawDef of name located * constr_expr | LocalRawAssum of name located list * constr_expr +(**********************************************************************) +(* Utilities on constr_expr *) val constr_loc : constr_expr -> loc -val cases_pattern_loc : cases_pattern_expr -> loc +val cases_pattern_expr_loc : cases_pattern_expr -> loc val replace_vars_constr_expr : (identifier * identifier) list -> constr_expr -> constr_expr +val free_vars_of_constr_expr : constr_expr -> Idset.t val occur_var_constr_expr : identifier -> constr_expr -> bool (* Specific function for interning "in indtype" syntax of "match" *) -val names_of_cases_indtype : constr_expr -> identifier list +val ids_of_cases_indtype : constr_expr -> identifier list val mkIdentC : identifier -> constr_expr val mkRefC : reference -> constr_expr @@ -175,10 +195,11 @@ val names_of_local_binders : local_binder list -> name located list (* in pattern-matching clauses and in binders of the form [x,y:T(x)] *) val map_constr_expr_with_binders : - ('a -> constr_expr -> constr_expr) -> - (identifier -> 'a -> 'a) -> 'a -> constr_expr -> constr_expr + (identifier -> 'a -> 'a) -> ('a -> constr_expr -> constr_expr) -> + 'a -> constr_expr -> constr_expr -(* Concrete syntax for modules and modules types *) +(**********************************************************************) +(* Concrete syntax for modules and module types *) type with_declaration_ast = | CWith_Module of identifier list located * qualid located @@ -191,6 +212,3 @@ type module_type_ast = type module_ast = | CMEident of qualid located | CMEapply of module_ast * module_ast - -(* Special identifier to encode recursive notations *) -val ldots_var : identifier diff --git a/parsing/g_constr.ml4 b/parsing/g_constr.ml4 index 9c8632a6d..835d66434 100644 --- a/parsing/g_constr.ml4 +++ b/parsing/g_constr.ml4 @@ -293,7 +293,7 @@ GEXTEND Gram (match p with | CPatAtom (_, Some r) -> CPatCstr (loc, r, lp) | _ -> Util.user_err_loc - (cases_pattern_loc p, "compound_pattern", + (cases_pattern_expr_loc p, "compound_pattern", Pp.str "Constructor expected")) | p = pattern; "as"; id = ident -> CPatAlias (loc, p, id) ] diff --git a/parsing/ppconstr.ml b/parsing/ppconstr.ml index 4ac2cbe9e..a003c1405 100644 --- a/parsing/ppconstr.ml +++ b/parsing/ppconstr.ml @@ -178,7 +178,7 @@ let rec pr_patt sep inh p = | CPatPrim (_,p) -> pr_prim_token p, latom | CPatDelimiters (_,k,p) -> pr_delimiters k (pr_patt mt lsimple p), 1 in - let loc = cases_pattern_loc p in + let loc = cases_pattern_expr_loc p in pr_with_comments loc (sep() ++ if prec_less prec inh then strm else surround strm) diff --git a/pretyping/rawterm.ml b/pretyping/rawterm.ml index 78e616d98..e426e299b 100644 --- a/pretyping/rawterm.ml +++ b/pretyping/rawterm.ml @@ -26,7 +26,7 @@ type cases_pattern = | PatVar of loc * name | PatCstr of loc * constructor * cases_pattern list * name -let pattern_loc = function +let cases_pattern_loc = function PatVar(loc,_) -> loc | PatCstr(loc,_,_,_) -> loc diff --git a/pretyping/rawterm.mli b/pretyping/rawterm.mli index c61c61b0f..8d1ac2e65 100644 --- a/pretyping/rawterm.mli +++ b/pretyping/rawterm.mli @@ -17,7 +17,8 @@ open Libnames open Nametab (*i*) -(* Untyped intermediate terms, after ASTs and before constr. *) +(**********************************************************************) +(* The kind of patterns that occurs in "match ... with ... end" *) (* locs here refers to the ident's location, not whole pat *) (* the last argument of PatCstr is a possible alias ident for the pattern *) @@ -25,7 +26,13 @@ type cases_pattern = | PatVar of loc * name | PatCstr of loc * constructor * cases_pattern list * name -val pattern_loc : cases_pattern -> loc +val cases_pattern_loc : cases_pattern -> loc + +(**********************************************************************) +(* Untyped intermediate terms, after constr_expr and before constr *) +(* Resolution of names, insertion of implicit arguments placeholder, *) +(* and notations are done, but coercions, inference of implicit *) +(* arguments and pattern-matching compilation are not *) type patvar = identifier diff --git a/test-suite/output/Notations.out b/test-suite/output/Notations.out index e8e206791..975e2ffdf 100644 --- a/test-suite/output/Notations.out +++ b/test-suite/output/Notations.out @@ -14,6 +14,12 @@ forall n : nat, n = 0 : Prop !(0 = 0) : Prop +forall n : nat, #(n = n) + : Prop +forall n n0 : nat, ##(n = n0) + : Prop +forall n n0 : nat, ###(n = n0) + : Prop 3 + 3 : Z 3 + 3 diff --git a/test-suite/output/Notations.v b/test-suite/output/Notations.v index 93ccf7a8c..1d24e70a1 100644 --- a/test-suite/output/Notations.v +++ b/test-suite/output/Notations.v @@ -28,6 +28,23 @@ Check forall n:nat, 0=0. End A. (**********************************************************************) +(* Behaviour wrt to binding variables (cf bug report #1186) *) + +Section B. + +Notation "# A" := (forall n:nat, n=n->A) (at level 60). +Check forall n:nat, # (n=n). + +Notation "## A" := (forall n n0:nat, n=n0->A) (at level 60). +Check forall n n0:nat, ## (n=n0). + +Notation "### A" := + (forall n n0:nat, match n with O => True | S n => n=n0 end ->A) (at level 60). +Check forall n n0:nat, ### (n=n0). + +End B. + +(**********************************************************************) (* Conflict between notation and notation below coercions *) (* Case of a printer conflict *) |