diff options
author | Hugo Herbelin <Hugo.Herbelin@inria.fr> | 2016-07-19 13:19:34 +0200 |
---|---|---|
committer | Hugo Herbelin <Hugo.Herbelin@inria.fr> | 2016-07-19 13:45:23 +0200 |
commit | f7ae4e6433e44a0b3a838847c58ab72ffffa3d48 (patch) | |
tree | f05d58a6f51c77ce890452ff00babd8cadf2e990 | |
parent | a67bd7f93224c61b6a59459ea1114a6670daa857 (diff) |
Some extra fixes in printing patterns in binders.
- typo in notation_ops.ml
- factorization of patterns in ppconstr.ml
- update of test-suite
- printing of cast of a binding pattern if in mode "printing all"
The question of whether or not to print the type of a binding pattern
by default seems open to me.
-rw-r--r-- | interp/constrextern.ml | 5 | ||||
-rw-r--r-- | interp/notation_ops.ml | 2 | ||||
-rw-r--r-- | printing/ppconstr.ml | 43 | ||||
-rw-r--r-- | test-suite/output/PatternsInBinders.out | 22 | ||||
-rw-r--r-- | test-suite/output/PatternsInBinders.v | 32 |
5 files changed, 52 insertions, 52 deletions
diff --git a/interp/constrextern.ml b/interp/constrextern.ml index 747687189..e71daef99 100644 --- a/interp/constrextern.ml +++ b/interp/constrextern.ml @@ -848,10 +848,11 @@ and extern_local_binder scopes vars = function | (Inr p,bk,Some bd,ty)::l -> assert false | (Inr p,bk,None,ty)::l -> - let ty = extern_typ scopes vars ty in + let ty = + if !Flags.raw_print then Some (extern_typ scopes vars ty) else None in let p = extern_cases_pattern vars p in let (assums,ids,l) = extern_local_binder scopes vars l in - (assums,ids, LocalPattern(Loc.ghost,p,Some ty) :: l) + (assums,ids, LocalPattern(Loc.ghost,p,ty) :: l) and extern_eqn inctx scopes vars (loc,ids,pl,c) = (loc,[loc,List.map (extern_cases_pattern_in_scope scopes vars) pl], diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml index 12da34462..6478ade61 100644 --- a/interp/notation_ops.ml +++ b/interp/notation_ops.ml @@ -869,7 +869,7 @@ let rec match_iterated_binders islambda decls = function match_iterated_binders islambda ((Inr cp,bk,None,t)::decls) b | GLambda (_,na,bk,t,b) when islambda -> match_iterated_binders islambda ((Inl na,bk,None,t)::decls) b - | GLambda (_,Name p,bk,t,GCases (_,LetPatternStyle,None,[(GVar(_,e),_)],[(_,_,[cp],b)])) + | GProd (_,Name p,bk,t,GCases (_,LetPatternStyle,None,[(GVar(_,e),_)],[(_,_,[cp],b)])) when not islambda && Id.equal p e -> match_iterated_binders islambda ((Inr cp,bk,None,t)::decls) b | GProd (_,(Name _ as na),bk,t,b) when not islambda -> diff --git a/printing/ppconstr.ml b/printing/ppconstr.ml index 7aa6d4858..dab97d603 100644 --- a/printing/ppconstr.ml +++ b/printing/ppconstr.ml @@ -350,7 +350,12 @@ end) = struct surround (pr_lname na ++ pr_opt_type pr_c topt ++ str":=" ++ cut() ++ pr_c c) | LocalPattern (loc,p,tyo) -> - str "'" ++ pr_patt lsimplepatt p + let p = pr_patt lsimplepatt p in + match tyo with + | None -> + str "'" ++ p + | Some ty -> + str "'" ++ surround (p ++ spc () ++ str ":" ++ ws 1 ++ pr_c ty) let pr_undelimited_binders sep pr_c = prlist_with_sep sep (pr_binder_among_many pr_c) @@ -374,6 +379,9 @@ end) = struct if bl = [] then [], x else LocalRawDef (na,b) :: bl, c*) | CProdN (loc,[],c) -> extract_prod_binders c + | CProdN (loc,[[_,Name id],bk,t],CCases (_,LetPatternStyle,None, [CRef (Ident (_,id'),None),None,None],[(_,[_,[p]],b)])) -> + let bl,c = extract_prod_binders b in + LocalPattern (loc,p,None) :: bl, c | CProdN (loc,(nal,bk,t)::bl,c) -> let bl,c = extract_prod_binders (CProdN(loc,bl,c)) in LocalRawAssum (nal,bk,t) :: bl, c @@ -385,6 +393,9 @@ end) = struct if bl = [] then [], x else LocalRawDef (na,b) :: bl, c*) | CLambdaN (loc,[],c) -> extract_lam_binders c + | CLambdaN (loc,[[_,Name id],bk,t],CCases (_,LetPatternStyle,None, [CRef (Ident (_,id'),None),None,None],[(_,[_,[p]],b)])) -> + let bl,c = extract_lam_binders b in + LocalPattern (loc,p,None) :: bl, c | CLambdaN (loc,(nal,bk,t)::bl,c) -> let bl,c = extract_lam_binders (CLambdaN(loc,bl,c)) in LocalRawAssum (nal,bk,t) :: bl, c @@ -536,21 +547,6 @@ end) = struct (pr_cofixdecl (pr mt) (pr_dangling_with_for mt pr)) (snd id) cofix), lfix ) - | CProdN - (_, - [([(_,Name n)],_,_)], - CCases - (_,LetPatternStyle,None,[(CRef(Ident(_,m),None),None,None)], - [(_,[(_,[p])],a)])) - when - Id.equal m n && - not (Id.Set.mem n (Topconstr.free_vars_of_constr_expr a)) -> - return ( - hov 0 ( - keyword "forall" ++ spc () ++ str "'" ++ pr_patt lsimplepatt p ++ - str "," ++ pr spc ltop a), - llambda - ) | CProdN _ -> let (bl,a) = extract_prod_binders a in return ( @@ -560,21 +556,6 @@ end) = struct str "," ++ pr spc ltop a), lprod ) - | CLambdaN - (_, - [([(_,Name n)],_,_)], - CCases - (_,LetPatternStyle,None,[(CRef(Ident(_,m),None),None,None)], - [(_,[(_,[p])],a)])) - when - Id.equal m n && - not (Id.Set.mem n (Topconstr.free_vars_of_constr_expr a)) -> - return ( - hov 0 ( - keyword "fun" ++ spc () ++ str "'" ++ pr_patt lsimplepatt p ++ - pr_fun_sep ++ pr spc ltop a), - llambda - ) | CLambdaN _ -> let (bl,a) = extract_lam_binders a in return ( diff --git a/test-suite/output/PatternsInBinders.out b/test-suite/output/PatternsInBinders.out index 6acaa0cce..c012a86b0 100644 --- a/test-suite/output/PatternsInBinders.out +++ b/test-suite/output/PatternsInBinders.out @@ -9,23 +9,31 @@ proj_informative = fun '(exist _ x _) => x : A foo = fun '(Bar n b tt p) => if b then n + p else n - p : Foo -> nat baz = -fun '(Bar n1 _ tt p1) => fun '(Bar _ _ tt _) => n1 + p1 +fun '(Bar n1 _ tt p1) '(Bar _ _ tt _) => n1 + p1 : Foo -> Foo -> nat -λ '(x, y), (y, x) - : A * B → B * A -∀ '(x, y), swap (x, y) = (y, x) - : Prop swap = -fun (A B : Type) (pat : A * B) => let '(x, y) := pat in (y, x) +fun (A B : Type) '(x, y) => (y, x) : forall A B : Type, A * B -> B * A Arguments A, B are implicit and maximally inserted Argument scopes are [type_scope type_scope _] -forall (A B : Type) (pat : A * B), let '(x, y) := pat in swap (x, y) = (y, x) +fun (A B : Type) '(x, y) => swap (x, y) = (y, x) + : forall A B : Type, A * B -> Prop +forall (A B : Type) '(x, y), swap (x, y) = (y, x) : Prop exists '(x, y), swap (x, y) = (y, x) : Prop +exists '(x, y) '(z, w), swap (x, y) = (z, w) + : Prop +λ '(x, y), (y, x) + : A * B → B * A +∀ '(x, y), swap (x, y) = (y, x) + : Prop both_z = fun pat : nat * nat => let '(n, p) as pat0 := pat return (F pat0) in (Z n, Z p) : F (n, p) : forall pat : nat * nat, F pat +fun '(x, y) '(z, t) => swap (x, y) = (z, t) + : A * B -> B * A -> Prop +forall '(x, y) '(z, t), swap (x, y) = (z, t) + : Prop diff --git a/test-suite/output/PatternsInBinders.v b/test-suite/output/PatternsInBinders.v index 8911909ab..b5c91e347 100644 --- a/test-suite/output/PatternsInBinders.v +++ b/test-suite/output/PatternsInBinders.v @@ -21,7 +21,20 @@ Print foo. Definition baz '(Bar n1 b1 tt p1) '(Bar n2 b2 tt p2) := n1+p1. Print baz. -(** Some test involving unicode noations. *) +Module WithParameters. + +Definition swap {A B} '((x,y) : A*B) := (y,x). +Print swap. + +Check fun (A B:Type) '((x,y) : A*B) => swap (x,y) = (y,x). +Check forall (A B:Type) '((x,y) : A*B), swap (x,y) = (y,x). + +Check exists '((x,y):A*A), swap (x,y) = (y,x). +Check exists '((x,y):A*A) '(z,w), swap (x,y) = (z,w). + +End WithParameters. + +(** Some test involving unicode notations. *) Module WithUnicode. Require Import Coq.Unicode.Utf8. @@ -31,24 +44,21 @@ Module WithUnicode. End WithUnicode. - (** * Suboptimal printing *) -(** These tests show examples which expose the [let] introduced by - the pattern notation in binders. *) - Module Suboptimal. -Definition swap {A B} '((x,y) : A*B) := (y,x). -Print swap. - -Check forall (A B:Type) '((x,y) : A*B), swap (x,y) = (y,x). - -Check exists '((x,y):A*A), swap (x,y) = (y,x). +(** This test shows an example which exposes the [let] introduced by + the pattern notation in binders. *) Inductive Fin (n:nat) := Z : Fin n. Definition F '(n,p) : Type := (Fin n * Fin p)%type. Definition both_z '(n,p) : F (n,p) := (Z _,Z _). Print both_z. +(** These tests show examples which do not factorize binders *) + +Check fun '((x,y) : A*B) '(z,t) => swap (x,y) = (z,t). +Check forall '(x,y) '((z,t) : B*A), swap (x,y) = (z,t). + End Suboptimal. |