From 4e70791036a1ab189579e109b428f46f45698b59 Mon Sep 17 00:00:00 2001
From: Hugo Herbelin
Date: Thu, 13 Apr 2017 12:13:04 +0200
Subject: Adding a fold_glob_constr_with_binders combinator.
Binding generalizable_vars_of_glob_constr, occur_glob_constr,
free_glob_vars, and bound_glob_vars on it.
Most of the functions of which it factorizes the code were bugged with
respect to bindings in the return clause of "match" and in either the
types or the bodies of "fix/cofix".
---
interp/implicit_quantifiers.ml | 56 +--------------
pretyping/glob_ops.ml | 153 ++++++++++++++---------------------------
pretyping/glob_ops.mli | 1 +
test-suite/.csdp.cache | Bin 89077 -> 89077 bytes
test-suite/success/boundvars.v | 5 ++
5 files changed, 57 insertions(+), 158 deletions(-)
create mode 100644 test-suite/success/boundvars.v
diff --git a/interp/implicit_quantifiers.ml b/interp/implicit_quantifiers.ml
index 7f11c0a3b6..d6749e918f 100644
--- a/interp/implicit_quantifiers.ml
+++ b/interp/implicit_quantifiers.ml
@@ -131,61 +131,7 @@ let generalizable_vars_of_glob_constr ?(bound=Id.Set.empty) ?(allowed=Id.Set.emp
if Id.List.mem_assoc id vs then vs
else (id, loc) :: vs
else vs
- | GApp (loc,f,args) -> List.fold_left (vars bound) vs (f::args)
- | GLambda (loc,na,_,ty,c) | GProd (loc,na,_,ty,c) ->
- let vs' = vars bound vs ty in
- let bound' = add_name_to_ids bound na in
- vars bound' vs' c
- | GLetIn (loc,na,b,ty,c) ->
- let vs' = vars bound vs b in
- let vs'' = Option.fold_left (vars bound) vs' ty in
- let bound' = add_name_to_ids bound na in
- vars bound' vs'' c
- | GCases (loc,sty,rtntypopt,tml,pl) ->
- let vs1 = vars_option bound vs rtntypopt in
- let vs2 = List.fold_left (fun vs (tm,_) -> vars bound vs tm) vs1 tml in
- List.fold_left (vars_pattern bound) vs2 pl
- | GLetTuple (loc,nal,rtntyp,b,c) ->
- let vs1 = vars_return_type bound vs rtntyp in
- let vs2 = vars bound vs1 b in
- let bound' = List.fold_left add_name_to_ids bound nal in
- vars bound' vs2 c
- | GIf (loc,c,rtntyp,b1,b2) ->
- let vs1 = vars_return_type bound vs rtntyp in
- let vs2 = vars bound vs1 c in
- let vs3 = vars bound vs2 b1 in
- vars bound vs3 b2
- | GRec (loc,fk,idl,bl,tyl,bv) ->
- let bound' = Array.fold_right Id.Set.add idl bound in
- let vars_fix i vs fid =
- let vs1,bound1 =
- List.fold_left
- (fun (vs,bound) (na,k,bbd,bty) ->
- let vs' = vars_option bound vs bbd in
- let vs'' = vars bound vs' bty in
- let bound' = add_name_to_ids bound na in
- (vs'',bound')
- )
- (vs,bound')
- bl.(i)
- in
- let vs2 = vars bound1 vs1 tyl.(i) in
- vars bound1 vs2 bv.(i)
- in
- Array.fold_left_i vars_fix vs idl
- | GCast (loc,c,k) -> let v = vars bound vs c in
- (match k with CastConv t | CastVM t -> vars bound v t | _ -> v)
- | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> vs
-
- and vars_pattern bound vs (loc,idl,p,c) =
- let bound' = List.fold_right Id.Set.add idl bound in
- vars bound' vs c
-
- and vars_option bound vs = function None -> vs | Some p -> vars bound vs p
-
- and vars_return_type bound vs (na,tyopt) =
- let bound' = add_name_to_ids bound na in
- vars_option bound' vs tyopt
+ | c -> Glob_ops.fold_glob_constr_with_binders Id.Set.add vars bound vs c
in fun rt ->
let vars = List.rev (vars bound [] rt) in
List.iter (fun (id, loc) ->
diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml
index ebbfa195f0..aa296aace7 100644
--- a/pretyping/glob_ops.ml
+++ b/pretyping/glob_ops.ml
@@ -214,55 +214,57 @@ let fold_glob_constr f acc = function
f acc c
| (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> acc
-let iter_glob_constr f = fold_glob_constr (fun () -> f) ()
+let fold_return_type_with_binders f g v acc (na,tyopt) =
+ Option.fold_left (f (name_fold g na v)) acc tyopt
+
+let fold_glob_constr_with_binders g f v acc = function
+ | GVar _ -> acc
+ | GApp (_,c,args) -> List.fold_left (f v) (f v acc c) args
+ | GLambda (_,na,_,b,c) | GProd (_,na,_,b,c) ->
+ f (name_fold g na v) (f v acc b) c
+ | GLetIn (_,na,b,t,c) ->
+ f (name_fold g na v) (Option.fold_left (f v) (f v acc b) t) c
+ | GCases (_,_,rtntypopt,tml,pl) ->
+ let fold_pattern acc (_,idl,p,c) = f (List.fold_right g idl v) acc c in
+ let fold_tomatch (v',acc) (tm,(na,onal)) =
+ (Option.fold_left (fun v'' (_,_,nal) -> List.fold_right (name_fold g) nal v'')
+ (name_fold g na v') onal,
+ f v acc tm) in
+ let (v',acc) = List.fold_left fold_tomatch (v,acc) tml in
+ let acc = Option.fold_left (f v') acc rtntypopt in
+ List.fold_left fold_pattern acc pl
+ | GLetTuple (_,nal,rtntyp,b,c) ->
+ f v (f v (fold_return_type_with_binders f g v acc rtntyp) b) c
+ | GIf (_,c,rtntyp,b1,b2) ->
+ f v (f v (f v (fold_return_type_with_binders f g v acc rtntyp) c) b1) b2
+ | GRec (_,_,idl,bll,tyl,bv) ->
+ let f' i acc fid =
+ let v,acc =
+ List.fold_left
+ (fun (v,acc) (na,k,bbd,bty) ->
+ (name_fold g na v, f v (Option.fold_left (f v) acc bbd) bty))
+ (v,acc)
+ bll.(i) in
+ f (Array.fold_right g idl v) (f v acc tyl.(i)) (bv.(i)) in
+ Array.fold_left_i f' acc idl
+ | GCast (_,c,k) ->
+ let acc = match k with
+ | CastConv t | CastVM t | CastNative t -> f v acc t | CastCoerce -> acc in
+ f v acc c
+ | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> acc
-let same_id na id = match na with
-| Anonymous -> false
-| Name id' -> Id.equal id id'
+let iter_glob_constr f = fold_glob_constr (fun () -> f) ()
let occur_glob_constr id =
- let rec occur = function
+ let rec occur barred acc = function
| GVar (loc,id') -> Id.equal id id'
- | GApp (loc,f,args) -> (occur f) || (List.exists occur args)
- | GLambda (loc,na,bk,ty,c) ->
- (occur ty) || (not (same_id na id) && (occur c))
- | GProd (loc,na,bk,ty,c) ->
- (occur ty) || (not (same_id na id) && (occur c))
- | GLetIn (loc,na,b,t,c) ->
- (Option.fold_left (fun b t -> occur t || b) (occur b) t) || (not (same_id na id) && (occur c))
- | GCases (loc,sty,rtntypopt,tml,pl) ->
- (occur_option rtntypopt)
- || (List.exists (fun (tm,_) -> occur tm) tml)
- || (List.exists occur_pattern pl)
- | GLetTuple (loc,nal,rtntyp,b,c) ->
- occur_return_type rtntyp id
- || (occur b) || (not (List.mem_f Name.equal (Name id) nal) && (occur c))
- | GIf (loc,c,rtntyp,b1,b2) ->
- occur_return_type rtntyp id || (occur c) || (occur b1) || (occur b2)
- | GRec (loc,fk,idl,bl,tyl,bv) ->
- not (Array.for_all4 (fun fid bl ty bd ->
- let rec occur_fix = function
- [] -> not (occur ty) && (Id.equal fid id || not(occur bd))
- | (na,k,bbd,bty)::bl ->
- not (occur bty) &&
- (match bbd with
- Some bd -> not (occur bd)
- | _ -> true) &&
- (match na with Name id' -> Id.equal id id' | _ -> not (occur_fix bl)) in
- occur_fix bl)
- idl bl tyl bv)
- | GCast (loc,c,k) -> (occur c) || (match k with CastConv t
- | CastVM t | CastNative t -> occur t | CastCoerce -> false)
- | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> false
-
- and occur_pattern (loc,idl,p,c) = not (Id.List.mem id idl) && (occur c)
-
- and occur_option = function None -> false | Some p -> occur p
-
- and occur_return_type (na,tyopt) id = not (same_id na id) && occur_option tyopt
-
- in occur
-
+ | c ->
+ (* [g] looks if [id] appears in a binding position, in which
+ case, we don't have to look in the corresponding subterm *)
+ let g id' barred = barred || Id.equal id id' in
+ let f barred acc c = acc || not barred && occur false acc c in
+ fold_glob_constr_with_binders g f barred acc c in
+ occur false false
let add_name_to_ids set na =
match na with
@@ -270,64 +272,9 @@ let add_name_to_ids set na =
| Name id -> Id.Set.add id set
let free_glob_vars =
- let rec vars bounded vs = function
- | GVar (loc,id') -> if Id.Set.mem id' bounded then vs else Id.Set.add id' vs
- | GApp (loc,f,args) -> List.fold_left (vars bounded) vs (f::args)
- | GLambda (loc,na,_,ty,c) | GProd (loc,na,_,ty,c) ->
- let vs' = vars bounded vs ty in
- let bounded' = add_name_to_ids bounded na in
- vars bounded' vs' c
- | GLetIn (loc,na,b,ty,c) ->
- let vs' = vars bounded vs b in
- let vs'' = Option.fold_left (vars bounded) vs' ty in
- let bounded' = add_name_to_ids bounded na in
- vars bounded' vs'' c
- | GCases (loc,sty,rtntypopt,tml,pl) ->
- let vs1 = vars_option bounded vs rtntypopt in
- let vs2 = List.fold_left (fun vs (tm,_) -> vars bounded vs tm) vs1 tml in
- List.fold_left (vars_pattern bounded) vs2 pl
- | GLetTuple (loc,nal,rtntyp,b,c) ->
- let vs1 = vars_return_type bounded vs rtntyp in
- let vs2 = vars bounded vs1 b in
- let bounded' = List.fold_left add_name_to_ids bounded nal in
- vars bounded' vs2 c
- | GIf (loc,c,rtntyp,b1,b2) ->
- let vs1 = vars_return_type bounded vs rtntyp in
- let vs2 = vars bounded vs1 c in
- let vs3 = vars bounded vs2 b1 in
- vars bounded vs3 b2
- | GRec (loc,fk,idl,bl,tyl,bv) ->
- let bounded' = Array.fold_right Id.Set.add idl bounded in
- let vars_fix i vs fid =
- let vs1,bounded1 =
- List.fold_left
- (fun (vs,bounded) (na,k,bbd,bty) ->
- let vs' = vars_option bounded vs bbd in
- let vs'' = vars bounded vs' bty in
- let bounded' = add_name_to_ids bounded na in
- (vs'',bounded')
- )
- (vs,bounded')
- bl.(i)
- in
- let vs2 = vars bounded1 vs1 tyl.(i) in
- vars bounded1 vs2 bv.(i)
- in
- Array.fold_left_i vars_fix vs idl
- | GCast (loc,c,k) -> let v = vars bounded vs c in
- (match k with CastConv t | CastVM t | CastNative t -> vars bounded v t | _ -> v)
- | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> vs
-
- and vars_pattern bounded vs (loc,idl,p,c) =
- let bounded' = List.fold_right Id.Set.add idl bounded in
- vars bounded' vs c
-
- and vars_option bounded vs = function None -> vs | Some p -> vars bounded vs p
-
- and vars_return_type bounded vs (na,tyopt) =
- let bounded' = add_name_to_ids bounded na in
- vars_option bounded' vs tyopt
- in
+ let rec vars bound vs = function
+ | GVar (loc,id') -> if Id.Set.mem id' bound then vs else Id.Set.add id' vs
+ | c -> fold_glob_constr_with_binders Id.Set.add vars bound vs c in
fun rt ->
let vs = vars Id.Set.empty Id.Set.empty rt in
Id.Set.elements vs
diff --git a/pretyping/glob_ops.mli b/pretyping/glob_ops.mli
index 55e6b6533f..af2834e498 100644
--- a/pretyping/glob_ops.mli
+++ b/pretyping/glob_ops.mli
@@ -37,6 +37,7 @@ val map_glob_constr_left_to_right :
val warn_variable_collision : ?loc:Loc.t -> Id.t -> unit
val fold_glob_constr : ('a -> glob_constr -> 'a) -> 'a -> glob_constr -> 'a
+val fold_glob_constr_with_binders : (Id.t -> 'a -> 'a) -> ('a -> 'b -> glob_constr -> 'b) -> 'a -> 'b -> glob_constr -> 'b
val iter_glob_constr : (glob_constr -> unit) -> glob_constr -> unit
val occur_glob_constr : Id.t -> glob_constr -> bool
val free_glob_vars : glob_constr -> Id.t list
diff --git a/test-suite/.csdp.cache b/test-suite/.csdp.cache
index ba85286dd3..b99d80e95f 100644
Binary files a/test-suite/.csdp.cache and b/test-suite/.csdp.cache differ
diff --git a/test-suite/success/boundvars.v b/test-suite/success/boundvars.v
new file mode 100644
index 0000000000..7b6696af8e
--- /dev/null
+++ b/test-suite/success/boundvars.v
@@ -0,0 +1,5 @@
+(* An example showing a bug in the detection of free variables *)
+(* "x" is not free in the common type of "x" and "y" *)
+
+Check forall (x z:unit) (x y : match z as x return x=x with tt => eq_refl end = eq_refl), x=x.
+
--
cgit v1.2.3
From b4936da085b19ad508346d8e07ce1e922ef79c2d Mon Sep 17 00:00:00 2001
From: Hugo Herbelin
Date: Thu, 13 Apr 2017 15:05:16 +0200
Subject: Using fold_glob_constr_with_binders to code bound_glob_vars.
To use the generic combinator, we introduce a side effect. I believe
that we have more to gain from a short code than from being purely
functional.
This also fixes the expected semantics since the variables binding the
return type in "match" were not taking into account.
---
pretyping/glob_ops.ml | 57 ++++++-----------------------------------
test-suite/.csdp.cache | Bin 89077 -> 89077 bytes
test-suite/success/boundvars.v | 9 +++++++
3 files changed, 17 insertions(+), 49 deletions(-)
diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml
index aa296aace7..080ec5ed12 100644
--- a/pretyping/glob_ops.ml
+++ b/pretyping/glob_ops.ml
@@ -300,57 +300,16 @@ let add_and_check_ident id set =
Id.Set.add id set
let bound_glob_vars =
- let rec vars bound = function
- | GLambda (_,na,_,_,_) | GProd (_,na,_,_,_) | GLetIn (_,na,_,_,_) as c ->
- let bound = name_fold add_and_check_ident na bound in
- fold_glob_constr vars bound c
- | GCases (loc,sty,rtntypopt,tml,pl) ->
- let bound = vars_option bound rtntypopt in
- let bound =
- List.fold_left (fun bound (tm,_) -> vars bound tm) bound tml in
- List.fold_left vars_pattern bound pl
- | GLetTuple (loc,nal,rtntyp,b,c) ->
- let bound = vars_return_type bound rtntyp in
- let bound = vars bound b in
- let bound = List.fold_right (name_fold add_and_check_ident) nal bound in
- vars bound c
- | GIf (loc,c,rtntyp,b1,b2) ->
- let bound = vars_return_type bound rtntyp in
- let bound = vars bound c in
- let bound = vars bound b1 in
- vars bound b2
- | GRec (loc,fk,idl,bl,tyl,bv) ->
- let bound = Array.fold_right Id.Set.add idl bound in
- let vars_fix i bound fid =
- let bound =
- List.fold_left
- (fun bound (na,k,bbd,bty) ->
- let bound = vars_option bound bbd in
- let bound = vars bound bty in
- name_fold add_and_check_ident na bound
- )
- bound
- bl.(i)
- in
- let bound = vars bound tyl.(i) in
- vars bound bv.(i)
- in
- Array.fold_left_i vars_fix bound idl
- | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _ | GVar _) -> bound
- | GApp _ | GCast _ as c -> fold_glob_constr vars bound c
-
- and vars_pattern bound (loc,idl,p,c) =
- let bound = List.fold_right add_and_check_ident idl bound in
- vars bound c
-
- and vars_option bound = function None -> bound | Some p -> vars bound p
-
- and vars_return_type bound (na,tyopt) =
- let bound = name_fold add_and_check_ident na bound in
- vars_option bound tyopt
+ let rec vars bound =
+ fold_glob_constr_with_binders
+ (fun id () -> bound := add_and_check_ident id !bound)
+ (fun () () -> vars bound)
+ () ()
in
fun rt ->
- vars Id.Set.empty rt
+ let bound = ref Id.Set.empty in
+ vars bound rt;
+ !bound
(** Mapping of names in binders *)
diff --git a/test-suite/.csdp.cache b/test-suite/.csdp.cache
index b99d80e95f..ba85286dd3 100644
Binary files a/test-suite/.csdp.cache and b/test-suite/.csdp.cache differ
diff --git a/test-suite/success/boundvars.v b/test-suite/success/boundvars.v
index 7b6696af8e..fafe272925 100644
--- a/test-suite/success/boundvars.v
+++ b/test-suite/success/boundvars.v
@@ -3,3 +3,12 @@
Check forall (x z:unit) (x y : match z as x return x=x with tt => eq_refl end = eq_refl), x=x.
+(* An example showing a bug in the detection of bound variables *)
+
+Goal forall x, match x return x = x with 0 => eq_refl | _ => eq_refl end = eq_refl.
+intro.
+match goal with
+|- (match x as y in nat return y = y with O => _ | S n => _ end) = _ => assert (forall y, y = 0) end.
+intro.
+Check x0. (* Check that "y" has been bound to "x0" while matching "match x as x0 return x0=x0 with ... end" *)
+Abort.
--
cgit v1.2.3
From 6795bc07f53a842bcec76ad0329d0b4444a625ab Mon Sep 17 00:00:00 2001
From: Pierre-Marie Pédrot
Date: Sun, 2 Apr 2017 21:22:55 +0200
Subject: Replacing costly merges in UGraph.
---
kernel/uGraph.ml | 16 +++++++++-------
1 file changed, 9 insertions(+), 7 deletions(-)
diff --git a/kernel/uGraph.ml b/kernel/uGraph.ml
index 4884d0deb1..6971c0a2b6 100644
--- a/kernel/uGraph.ml
+++ b/kernel/uGraph.ml
@@ -354,13 +354,15 @@ let get_new_edges g to_merge =
UMap.empty to_merge
in
let ltle =
- UMap.fold (fun _ n acc ->
- UMap.merge (fun _ strict1 strict2 ->
- match strict1, strict2 with
- | Some true, _ | _, Some true -> Some true
- | _, _ -> Some false)
- acc n.ltle)
- to_merge_lvl UMap.empty
+ let fold _ n acc =
+ let fold u strict acc =
+ if strict then UMap.add u strict acc
+ else if UMap.mem u acc then acc
+ else UMap.add u false acc
+ in
+ UMap.fold fold n.ltle acc
+ in
+ UMap.fold fold to_merge_lvl UMap.empty
in
let ltle, _ = clean_ltle g ltle in
let ltle =
--
cgit v1.2.3
From 11aaa1fd8230a347f1dca1a0f349ea7c7f2768c3 Mon Sep 17 00:00:00 2001
From: Jason Gross
Date: Fri, 10 Jun 2016 18:58:24 -0400
Subject: Add support for transparent abstract (no syntax)
This is a small change that allows a transparent version of tclABSTRACT.
Additionally, it factors the machinery of [abstract] through a
plugin-accessible function which allows alternate continuations (other
than exact_no_check.
It might be nice to factor it further, into a cache_term function that
caches a term, and a separate bit that calls cache_term with the result
of running the tactic.
---
tactics/tactics.ml | 28 ++++++++++++++++++----------
tactics/tactics.mli | 4 +++-
2 files changed, 21 insertions(+), 11 deletions(-)
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index e792585822..19627eb530 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -4907,7 +4907,7 @@ let shrink_entry sign const =
} in
(const, args)
-let abstract_subproof id gk tac =
+let cache_term_by_tactic_then id gk ?(opaque=true) tac tacK =
let open Tacticals.New in
let open Tacmach.New in
let open Proofview.Notations in
@@ -4957,8 +4957,8 @@ let abstract_subproof id gk tac =
else (const, List.rev (Context.Named.to_instance Constr.mkVar sign))
in
let args = List.map EConstr.of_constr args in
- let cd = Entries.DefinitionEntry const in
- let decl = (cd, IsProof Lemma) in
+ let cd = Entries.DefinitionEntry { const with Entries.const_entry_opaque = opaque } in
+ let decl = (cd, if opaque then IsProof Lemma else IsDefinition Definition) in
let cst () =
(** do not compute the implicit arguments, it may be costly *)
let () = Impargs.make_implicit_args false in
@@ -4976,18 +4976,21 @@ let abstract_subproof id gk tac =
Entries.(snd (Future.force const.const_entry_body)) in
let solve =
Proofview.tclEFFECTS effs <*>
- exact_no_check (applist (lem, args))
+ tacK lem args
in
let tac = if not safe then Proofview.mark_as_unsafe <*> solve else solve in
Sigma.Unsafe.of_pair (tac, evd)
end }
+let abstract_subproof id gk tac ?(opaque=true) =
+ cache_term_by_tactic_then id gk ~opaque:opaque tac (fun lem args -> exact_no_check (applist (lem, args)))
+
let anon_id = Id.of_string "anonymous"
-let tclABSTRACT name_op tac =
+let name_op_to_name name_op object_kind suffix =
let open Proof_global in
- let default_gk = (Global, false, Proof Theorem) in
- let s, gk = match name_op with
+ let default_gk = (Global, false, object_kind) in
+ match name_op with
| Some s ->
(try let _, gk, _ = current_proof_statement () in s, gk
with NoCurrentProof -> s, default_gk)
@@ -4995,9 +4998,14 @@ let tclABSTRACT name_op tac =
let name, gk =
try let name, gk, _ = current_proof_statement () in name, gk
with NoCurrentProof -> anon_id, default_gk in
- add_suffix name "_subproof", gk
- in
- abstract_subproof s gk tac
+ add_suffix name suffix, gk
+
+let tclABSTRACT ?(opaque=true) name_op tac =
+ let open Proof_global in
+ let s, gk = if opaque
+ then name_op_to_name name_op (Proof Theorem) "_subproof"
+ else name_op_to_name name_op (DefinitionBody Definition) "_subterm" in
+ abstract_subproof s gk tac ~opaque:opaque
let unify ?(state=full_transparent_state) x y =
Proofview.Goal.s_enter { s_enter = begin fun gl ->
diff --git a/tactics/tactics.mli b/tactics/tactics.mli
index ba4a9706de..d206011eed 100644
--- a/tactics/tactics.mli
+++ b/tactics/tactics.mli
@@ -401,7 +401,9 @@ val generalize_dep : ?with_let:bool (** Don't lose let bindings *) -> constr -
val unify : ?state:Names.transparent_state -> constr -> constr -> unit Proofview.tactic
-val tclABSTRACT : Id.t option -> unit Proofview.tactic -> unit Proofview.tactic
+val cache_term_by_tactic_then : Id.t -> Decl_kinds.goal_kind -> ?opaque:bool -> unit Proofview.tactic -> (constr -> constr list -> unit Proofview.tactic) -> unit Proofview.tactic
+
+val tclABSTRACT : ?opaque:bool -> Id.t option -> unit Proofview.tactic -> unit Proofview.tactic
val abstract_generalize : ?generalize_vars:bool -> ?force_dep:bool -> Id.t -> unit Proofview.tactic
val specialize_eqs : Id.t -> unit Proofview.tactic
--
cgit v1.2.3
From 5f3d20dc53ffd0537a84c93acd761c3c69081342 Mon Sep 17 00:00:00 2001
From: Jason Gross
Date: Fri, 10 Jun 2016 19:12:49 -0400
Subject: Add transparent_abstract tactic
---
doc/refman/RefMan-ltac.tex | 14 +++++++++-----
plugins/ltac/extratactics.ml4 | 13 +++++++++++++
test-suite/success/transparent_abstract.v | 21 +++++++++++++++++++++
theories/Init/Prelude.v | 2 +-
4 files changed, 44 insertions(+), 6 deletions(-)
create mode 100644 test-suite/success/transparent_abstract.v
diff --git a/doc/refman/RefMan-ltac.tex b/doc/refman/RefMan-ltac.tex
index 9378529cbe..46274e12f3 100644
--- a/doc/refman/RefMan-ltac.tex
+++ b/doc/refman/RefMan-ltac.tex
@@ -1087,8 +1087,8 @@ Fail all:let n:= numgoals in guard n=2.
Reset Initial.
\end{coq_eval}
-\subsubsection[Proving a subgoal as a separate lemma]{Proving a subgoal as a separate lemma\tacindex{abstract}\comindex{Qed exporting}
-\index{Tacticals!abstract@{\tt abstract}}}
+\subsubsection[Proving a subgoal as a separate lemma]{Proving a subgoal as a separate lemma\tacindex{abstract}\tacindex{transparent\_abstract}\comindex{Qed exporting}
+\index{Tacticals!abstract@{\tt abstract}}\index{Tacticals!transparent\_abstract@{\tt transparent\_abstract}}}
From the outside ``\texttt{abstract \tacexpr}'' is the same as
{\tt solve \tacexpr}. Internally it saves an auxiliary lemma called
@@ -1114,13 +1114,17 @@ on. This can be obtained thanks to the option below.
{\tt Set Shrink Abstract}
\end{quote}
-When set, all lemmas generated through \texttt{abstract {\tacexpr}} are
-quantified only over the variables that appear in the term constructed by
-\texttt{\tacexpr}.
+When set, all lemmas generated through \texttt{abstract {\tacexpr}}
+and \texttt{transparent\_abstract {\tacexpr}} are quantified only over the
+variables that appear in the term constructed by \texttt{\tacexpr}.
\begin{Variants}
\item \texttt{abstract {\tacexpr} using {\ident}}.\\
Give explicitly the name of the auxiliary lemma.
+\item \texttt{transparent\_abstract {\tacexpr}}.\\
+ Save the subproof in a transparent lemma rather than an opaque one.
+\item \texttt{transparent\_abstract {\tacexpr} using {\ident}}.\\
+ Give explicitly the name of the auxiliary transparent lemma.
\end{Variants}
\ErrMsg \errindex{Proof is not complete}
diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4
index 38fdfb7599..a96623a5f6 100644
--- a/plugins/ltac/extratactics.ml4
+++ b/plugins/ltac/extratactics.ml4
@@ -815,6 +815,19 @@ TACTIC EXTEND destauto
| [ "destauto" "in" hyp(id) ] -> [ destauto_in id ]
END
+(**********************************************************************)
+
+(**********************************************************************)
+(* A version of abstract constructing transparent terms *)
+(* Introduced by Jason Gross and Benjamin Delaware in June 2016 *)
+(**********************************************************************)
+
+TACTIC EXTEND transparent_abstract
+| [ "transparent_abstract" tactic3(t) ] -> [ Proofview.Goal.nf_enter { enter = fun gl ->
+ Tactics.tclABSTRACT ~opaque:false None (Tacinterp.tactic_of_value ist t) } ]
+| [ "transparent_abstract" tactic3(t) "using" ident(id) ] -> [ Proofview.Goal.nf_enter { enter = fun gl ->
+ Tactics.tclABSTRACT ~opaque:false (Some id) (Tacinterp.tactic_of_value ist t) } ]
+END
(* ********************************************************************* *)
diff --git a/test-suite/success/transparent_abstract.v b/test-suite/success/transparent_abstract.v
new file mode 100644
index 0000000000..ff4509c4a8
--- /dev/null
+++ b/test-suite/success/transparent_abstract.v
@@ -0,0 +1,21 @@
+Class by_transparent_abstract {T} (x : T) := make_by_transparent_abstract : T.
+Hint Extern 0 (@by_transparent_abstract ?T ?x) => change T; transparent_abstract exact_no_check x : typeclass_instances.
+
+Goal True /\ True.
+Proof.
+ split.
+ transparent_abstract exact I using foo.
+ let x := (eval hnf in foo) in constr_eq x I.
+ let x := constr:(ltac:(constructor) : True) in
+ let T := type of x in
+ let x := constr:(_ : by_transparent_abstract x) in
+ let x := (eval cbv delta [by_transparent_abstract] in (let y : T := x in y)) in
+ pose x as x'.
+ simpl in x'.
+ let v := eval cbv [x'] in x' in tryif constr_eq v I then fail 0 else idtac.
+ hnf in x'.
+ let v := eval cbv [x'] in x' in tryif constr_eq v I then idtac else fail 0.
+ exact x'.
+Defined.
+Check eq_refl : I = foo.
+Eval compute in foo.
diff --git a/theories/Init/Prelude.v b/theories/Init/Prelude.v
index c58d23dad0..e71a8774ed 100644
--- a/theories/Init/Prelude.v
+++ b/theories/Init/Prelude.v
@@ -23,4 +23,4 @@ Declare ML Module "cc_plugin".
Declare ML Module "ground_plugin".
Declare ML Module "recdef_plugin".
(* Default substrings not considered by queries like SearchAbout *)
-Add Search Blacklist "_subproof" "Private_".
+Add Search Blacklist "_subproof" "_subterm" "Private_".
--
cgit v1.2.3
From 84845f766d9b9d532f615352fbc8a0e78e1727e9 Mon Sep 17 00:00:00 2001
From: Jason Gross
Date: Mon, 27 Mar 2017 13:28:17 -0400
Subject: Mark transparent_abstract as risky in docs
As per Enrico's request.
---
doc/refman/RefMan-ltac.tex | 7 +++++++
1 file changed, 7 insertions(+)
diff --git a/doc/refman/RefMan-ltac.tex b/doc/refman/RefMan-ltac.tex
index 46274e12f3..c2f52e23bc 100644
--- a/doc/refman/RefMan-ltac.tex
+++ b/doc/refman/RefMan-ltac.tex
@@ -1121,10 +1121,17 @@ variables that appear in the term constructed by \texttt{\tacexpr}.
\begin{Variants}
\item \texttt{abstract {\tacexpr} using {\ident}}.\\
Give explicitly the name of the auxiliary lemma.
+ Use this feature at your own risk; explicitly named and reused subterms
+ don't play well with asynchronous proofs.
\item \texttt{transparent\_abstract {\tacexpr}}.\\
Save the subproof in a transparent lemma rather than an opaque one.
+ Use this feature at your own risk; building computationally relevant terms
+ with tactics is fragile.
\item \texttt{transparent\_abstract {\tacexpr} using {\ident}}.\\
Give explicitly the name of the auxiliary transparent lemma.
+ Use this feature at your own risk; building computationally relevant terms
+ with tactics is fragile, and explicitly named and reused subterms
+ don't play well with asynchronous proofs.
\end{Variants}
\ErrMsg \errindex{Proof is not complete}
--
cgit v1.2.3
From 12f34b2ebfcbe958ba53b49399c3fcaf01f7a18c Mon Sep 17 00:00:00 2001
From: Jason Gross
Date: Tue, 11 Apr 2017 09:50:55 -0400
Subject: Generalize cache_term_by_tactic_then
This will allow a cache_term tactic that doesn't suffer from the
Not_found anomalies of abstract in typeclass resolution.
---
tactics/tactics.ml | 7 +++++--
tactics/tactics.mli | 2 +-
2 files changed, 6 insertions(+), 3 deletions(-)
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index 19627eb530..20de56645f 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -4907,7 +4907,7 @@ let shrink_entry sign const =
} in
(const, args)
-let cache_term_by_tactic_then id gk ?(opaque=true) tac tacK =
+let cache_term_by_tactic_then id gk ?(opaque=true) ?(goal_type=None) tac tacK =
let open Tacticals.New in
let open Tacmach.New in
let open Proofview.Notations in
@@ -4927,7 +4927,10 @@ let cache_term_by_tactic_then id gk ?(opaque=true) tac tacK =
else (Context.Named.add d s1,s2))
global_sign (Context.Named.empty, empty_named_context_val) in
let id = next_global_ident_away id (pf_ids_of_hyps gl) in
- let concl = it_mkNamedProd_or_LetIn (Proofview.Goal.concl gl) sign in
+ let concl = match goal_type with
+ | None -> Proofview.Goal.concl gl
+ | Some ty -> ty in
+ let concl = it_mkNamedProd_or_LetIn concl sign in
let concl =
try flush_and_check_evars !evdref concl
with Uninstantiated_evar _ ->
diff --git a/tactics/tactics.mli b/tactics/tactics.mli
index d206011eed..082812c5a6 100644
--- a/tactics/tactics.mli
+++ b/tactics/tactics.mli
@@ -401,7 +401,7 @@ val generalize_dep : ?with_let:bool (** Don't lose let bindings *) -> constr -
val unify : ?state:Names.transparent_state -> constr -> constr -> unit Proofview.tactic
-val cache_term_by_tactic_then : Id.t -> Decl_kinds.goal_kind -> ?opaque:bool -> unit Proofview.tactic -> (constr -> constr list -> unit Proofview.tactic) -> unit Proofview.tactic
+val cache_term_by_tactic_then : Id.t -> Decl_kinds.goal_kind -> ?opaque:bool -> ?goal_type:(constr option) -> unit Proofview.tactic -> (constr -> constr list -> unit Proofview.tactic) -> unit Proofview.tactic
val tclABSTRACT : ?opaque:bool -> Id.t option -> unit Proofview.tactic -> unit Proofview.tactic
--
cgit v1.2.3
From b348a11ccc4913598b72e4ecbb58811bcccd7bfc Mon Sep 17 00:00:00 2001
From: Jason Gross
Date: Tue, 11 Apr 2017 12:34:07 -0400
Subject: Make opaque optional only for tclABSTRACT
Also move named arguments to the beginning of the functions. As per
https://github.com/coq/coq/pull/201#discussion_r110928302
---
tactics/tactics.ml | 8 ++++----
tactics/tactics.mli | 2 +-
2 files changed, 5 insertions(+), 5 deletions(-)
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index 20de56645f..8f791cdcfe 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -4907,7 +4907,7 @@ let shrink_entry sign const =
} in
(const, args)
-let cache_term_by_tactic_then id gk ?(opaque=true) ?(goal_type=None) tac tacK =
+let cache_term_by_tactic_then ~opaque ?(goal_type=None) id gk tac tacK =
let open Tacticals.New in
let open Tacmach.New in
let open Proofview.Notations in
@@ -4985,8 +4985,8 @@ let cache_term_by_tactic_then id gk ?(opaque=true) ?(goal_type=None) tac tacK =
Sigma.Unsafe.of_pair (tac, evd)
end }
-let abstract_subproof id gk tac ?(opaque=true) =
- cache_term_by_tactic_then id gk ~opaque:opaque tac (fun lem args -> exact_no_check (applist (lem, args)))
+let abstract_subproof ~opaque id gk tac =
+ cache_term_by_tactic_then ~opaque:opaque id gk tac (fun lem args -> exact_no_check (applist (lem, args)))
let anon_id = Id.of_string "anonymous"
@@ -5008,7 +5008,7 @@ let tclABSTRACT ?(opaque=true) name_op tac =
let s, gk = if opaque
then name_op_to_name name_op (Proof Theorem) "_subproof"
else name_op_to_name name_op (DefinitionBody Definition) "_subterm" in
- abstract_subproof s gk tac ~opaque:opaque
+ abstract_subproof ~opaque:opaque s gk tac
let unify ?(state=full_transparent_state) x y =
Proofview.Goal.s_enter { s_enter = begin fun gl ->
diff --git a/tactics/tactics.mli b/tactics/tactics.mli
index 082812c5a6..07a8035427 100644
--- a/tactics/tactics.mli
+++ b/tactics/tactics.mli
@@ -401,7 +401,7 @@ val generalize_dep : ?with_let:bool (** Don't lose let bindings *) -> constr -
val unify : ?state:Names.transparent_state -> constr -> constr -> unit Proofview.tactic
-val cache_term_by_tactic_then : Id.t -> Decl_kinds.goal_kind -> ?opaque:bool -> ?goal_type:(constr option) -> unit Proofview.tactic -> (constr -> constr list -> unit Proofview.tactic) -> unit Proofview.tactic
+val cache_term_by_tactic_then : opaque:bool -> ?goal_type:(constr option) -> Id.t -> Decl_kinds.goal_kind -> unit Proofview.tactic -> (constr -> constr list -> unit Proofview.tactic) -> unit Proofview.tactic
val tclABSTRACT : ?opaque:bool -> Id.t option -> unit Proofview.tactic -> unit Proofview.tactic
--
cgit v1.2.3
From 1e046726dc9352f7979ebdeba0d750e44016fea5 Mon Sep 17 00:00:00 2001
From: Jason Gross
Date: Tue, 11 Apr 2017 12:48:23 -0400
Subject: transparent abstract: Respond to review comment
https://github.com/coq/coq/pull/201#discussion_r110952601
---
tactics/tactics.ml | 2 +-
1 file changed, 1 insertion(+), 1 deletion(-)
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index 8f791cdcfe..d02fe86653 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -4986,7 +4986,7 @@ let cache_term_by_tactic_then ~opaque ?(goal_type=None) id gk tac tacK =
end }
let abstract_subproof ~opaque id gk tac =
- cache_term_by_tactic_then ~opaque:opaque id gk tac (fun lem args -> exact_no_check (applist (lem, args)))
+ cache_term_by_tactic_then ~opaque id gk tac (fun lem args -> exact_no_check (applist (lem, args)))
let anon_id = Id.of_string "anonymous"
--
cgit v1.2.3
From e4262a89d7bc3d9b985d9a4a939f34176581abcb Mon Sep 17 00:00:00 2001
From: Jason Gross
Date: Tue, 11 Apr 2017 13:05:54 -0400
Subject: transparent abstract: Respond to review comment
https://github.com/coq/coq/pull/201#discussion_r110957570
---
tactics/tactics.ml | 2 +-
1 file changed, 1 insertion(+), 1 deletion(-)
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index d02fe86653..4654817035 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -5008,7 +5008,7 @@ let tclABSTRACT ?(opaque=true) name_op tac =
let s, gk = if opaque
then name_op_to_name name_op (Proof Theorem) "_subproof"
else name_op_to_name name_op (DefinitionBody Definition) "_subterm" in
- abstract_subproof ~opaque:opaque s gk tac
+ abstract_subproof ~opaque s gk tac
let unify ?(state=full_transparent_state) x y =
Proofview.Goal.s_enter { s_enter = begin fun gl ->
--
cgit v1.2.3
From e2a8edaf595827af82be67a90c0c5b22c987abe5 Mon Sep 17 00:00:00 2001
From: Hugo Herbelin
Date: Tue, 22 Nov 2016 13:50:10 +0100
Subject: A refined solution to the beta-iota discrepancies between 8.4 and 8.5
"refine".
There is a long story of commits trying to improve the compatibility
between 8.4 and 8.5 refine, as discussed in
https://github.com/coq/coq/pull/346.
ac9c5986b77bf4a783f2bd0ad571645694c960e1 add beta-iota in hypotheses and conclusion
8afac4f87d9d7e3add1c19485f475bd2207bfde7 remove beta-iota in hypotheses
08e87eb96ab67ead60d92394eec6066d9b52e55e re-add beta-iota in hypotheses
c9c54122d1d9493a965b483939e119d52121d5a6 re-remove beta-iota in hypotheses
9194180e2da0f7f9a2b2c7574bb7261cc69ead17 revert re-remove beta-iota in hypotheses
6bb352a6743c7332b9715ac15e95c806a58d101c re-re-remove beta-iota in hypotheses if <= 8.5
d8baa76d86eaa691a5386669596a6004bb44bb7a idem if = 8.5
The current commit tries to identify (one of?) the exact points of
divergence between 8.4 and 8.5 refine, namely the types inferred for
the variables of a pattern-matching problem.
Note that for the conclusion of each new goal, there were a
nf_betaiota in 8.4 done in function Evarutil.evars_to_metas, so the
compatibility expects that such a nf_betaiota on the conclusion of
each goal remains.
---
pretyping/cases.ml | 6 ++++++
1 file changed, 6 insertions(+)
diff --git a/pretyping/cases.ml b/pretyping/cases.ml
index 6bc2a4f94b..8a49cd5488 100644
--- a/pretyping/cases.ml
+++ b/pretyping/cases.ml
@@ -1245,6 +1245,12 @@ let build_branch initial current realargs deps (realnames,curname) pb arsign eqn
let typs = List.map2 RelDecl.set_name names cs_args
in
+ (* Beta-iota-normalize types to better compatibility of refine with 8.4 behavior *)
+ (* This is a bit too strong I think, in the sense that what we would *)
+ (* really like is to have beta-iota reduction only at the positions where *)
+ (* parameters are substituted *)
+ let typs = List.map (map_type (nf_betaiota !(pb.evdref))) typs in
+
(* We build the matrix obtained by expanding the matching on *)
(* "C x1..xn as x" followed by a residual matching on eqn into *)
(* a matching on "x1 .. xn eqn" *)
--
cgit v1.2.3
From 9839375100365ea3bd28bfc2efdb701db7d83adb Mon Sep 17 00:00:00 2001
From: Hugo Herbelin
Date: Mon, 5 Dec 2016 12:35:01 +0100
Subject: Test surgical use of beta-iota in the type of variables coming from
pattern-matching for refine.
---
test-suite/bugs/closed/5219.v | 10 ++++++++++
1 file changed, 10 insertions(+)
create mode 100644 test-suite/bugs/closed/5219.v
diff --git a/test-suite/bugs/closed/5219.v b/test-suite/bugs/closed/5219.v
new file mode 100644
index 0000000000..f7cec1a0cf
--- /dev/null
+++ b/test-suite/bugs/closed/5219.v
@@ -0,0 +1,10 @@
+(* Test surgical use of beta-iota in the type of variables coming from
+ pattern-matching for refine *)
+
+Goal forall x : sigT (fun x => x = 1), True.
+ intro x; refine match x with
+ | existT _ x' e' => _
+ end.
+ lazymatch goal with
+ | [ H : _ = _ |- _ ] => idtac
+ end.
--
cgit v1.2.3
From cc1212c3cfbd9c39cbe981210758c67cf9095be2 Mon Sep 17 00:00:00 2001
From: Hugo Herbelin
Date: Thu, 27 Apr 2017 17:25:14 +0200
Subject: Tentative note in CHANGES about now applying βι while typing "match"
branches.
In practice, this is almost invisible except when using "refine". So,
in some sense, it is aligning the behavior of pretyping on the one of
logic.ml's "refine" so that the more natural behavior of 8.4's refine
on this issue is restored.
---
CHANGES | 5 +++++
1 file changed, 5 insertions(+)
diff --git a/CHANGES b/CHANGES
index 60b88ea8db..8cb5573b21 100644
--- a/CHANGES
+++ b/CHANGES
@@ -18,6 +18,11 @@ Tactics
missed before because of a missing normalization step. Hopefully this should
be fairly uncommon.
- "auto with real" can now discharge comparisons of literals
+- The types of variables in patterns of "match" are now
+ beta-iota-reduced after type-checking. This has an impact on the
+ type of the variables that the tactic "refine" introduces in the
+ context, producing types a priori closer to the expectations.
+
Standard Library
--
cgit v1.2.3
From 746066172b8ed508886feb20cee239920ca7a4c7 Mon Sep 17 00:00:00 2001
From: Pierre-Marie Pédrot
Date: Thu, 27 Apr 2017 17:33:51 +0200
Subject: Test for bug #5193: Uncaught exception
Class_tactics.Search.ReachedLimitEx.
---
test-suite/bugs/closed/5193.v | 14 ++++++++++++++
1 file changed, 14 insertions(+)
create mode 100644 test-suite/bugs/closed/5193.v
diff --git a/test-suite/bugs/closed/5193.v b/test-suite/bugs/closed/5193.v
new file mode 100644
index 0000000000..cc8739afe6
--- /dev/null
+++ b/test-suite/bugs/closed/5193.v
@@ -0,0 +1,14 @@
+Class Eqdec A := eqdec : forall a b : A, {a=b}+{a<>b}.
+
+Typeclasses eauto := debug.
+Set Typeclasses Debug Verbosity 2.
+
+Inductive Finx(n : nat) : Set :=
+| Fx1(i : nat)(e : n = S i)
+| FxS(i : nat)(f : Finx i)(e : n = S i).
+
+Context `{Finx_eqdec : forall n, Eqdec (Finx n)}.
+
+Goal {x : Type & Eqdec x}.
+ eexists.
+ try typeclasses eauto 1 with typeclass_instances.
--
cgit v1.2.3
From e574b4bdd974daa7d2ceecf799762be92fadff44 Mon Sep 17 00:00:00 2001
From: Paul Steckler
Date: Thu, 27 Apr 2017 12:05:51 -0400
Subject: fix order of command-line arguments mentioned in Add LoadPath
---
doc/refman/RefMan-oth.tex | 8 ++++----
1 file changed, 4 insertions(+), 4 deletions(-)
diff --git a/doc/refman/RefMan-oth.tex b/doc/refman/RefMan-oth.tex
index 56ce753cd6..a25da7d926 100644
--- a/doc/refman/RefMan-oth.tex
+++ b/doc/refman/RefMan-oth.tex
@@ -697,8 +697,8 @@ which can be any valid path.
\subsection[\tt Add LoadPath {\str} as {\dirpath}.]{\tt Add LoadPath {\str} as {\dirpath}.\comindex{Add LoadPath}\label{AddLoadPath}}
-This command is equivalent to the command line option {\tt -Q {\dirpath}
- {\str}}. It adds the physical directory {\str} to the current {\Coq}
+This command is equivalent to the command line option {\tt -Q {\str}
+ {\dirpath}}. It adds the physical directory {\str} to the current {\Coq}
loadpath and maps it to the logical directory {\dirpath}.
\begin{Variants}
@@ -707,8 +707,8 @@ Performs as {\tt Add LoadPath {\str} as {\dirpath}} but for the empty directory
\end{Variants}
\subsection[\tt Add Rec LoadPath {\str} as {\dirpath}.]{\tt Add Rec LoadPath {\str} as {\dirpath}.\comindex{Add Rec LoadPath}\label{AddRecLoadPath}}
-This command is equivalent to the command line option {\tt -R {\dirpath}
- {\str}}. It adds the physical directory {\str} and all its
+This command is equivalent to the command line option {\tt -R {\str}
+ {\dirpath}}. It adds the physical directory {\str} and all its
subdirectories to the current {\Coq} loadpath.
\begin{Variants}
--
cgit v1.2.3
From 68fb8e13c44c5ee95dbc9256b1d74c7c83303d2d Mon Sep 17 00:00:00 2001
From: Hugo Herbelin
Date: Fri, 28 Apr 2017 16:30:45 +0200
Subject: Fixing #5487 (v8.5 regression on ltac-matching expressions with
evars).
The fix follows an invariant enforced in proofview.ml on the kind of
evars that are goals or that occur in goals.
One day, evar kinds will need a little cleaning...
---
pretyping/patternops.ml | 4 +++-
test-suite/bugs/closed/5487.v | 9 +++++++++
2 files changed, 12 insertions(+), 1 deletion(-)
create mode 100644 test-suite/bugs/closed/5487.v
diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml
index 2090aad8a0..75d3ed30ba 100644
--- a/pretyping/patternops.ml
+++ b/pretyping/patternops.ml
@@ -160,7 +160,9 @@ let pattern_of_constr env sigma t =
let ty = Evarutil.nf_evar sigma (existential_type sigma ev) in
let () = ignore (pattern_of_constr env ty) in
assert (not b); PMeta (Some id)
- | Evar_kinds.GoalEvar ->
+ | Evar_kinds.GoalEvar | Evar_kinds.VarInstance _ ->
+ (* These are the two evar kinds used for existing goals *)
+ (* see Proofview.mark_in_evm *)
PEvar (evk,Array.map (pattern_of_constr env) ctxt)
| _ ->
let ty = Evarutil.nf_evar sigma (existential_type sigma ev) in
diff --git a/test-suite/bugs/closed/5487.v b/test-suite/bugs/closed/5487.v
new file mode 100644
index 0000000000..9b995f4503
--- /dev/null
+++ b/test-suite/bugs/closed/5487.v
@@ -0,0 +1,9 @@
+(* Was a collision between an ltac pattern variable and an evar *)
+
+Goal forall n, exists m, n = m :> nat.
+Proof.
+ eexists.
+ Fail match goal with
+ | [ |- ?x = ?y ]
+ => match x with y => idtac end
+ end.
--
cgit v1.2.3
From db28e827d21658797418c320d566fb99570b44b6 Mon Sep 17 00:00:00 2001
From: Maxime Dénès
Date: Fri, 28 Apr 2017 22:20:35 +0200
Subject: Revert "Fixing #5487 (v8.5 regression on ltac-matching expressions
with evars)."
One day I'll get bored of spending my nights fixing commits that were
pushed without being tested, and I'll ask for removal of push rights.
But for now let's pretend I haven't insisted enough:
~~~~ PLEASE TEST YOUR COMMITS BEFORE PUSHING ~~~~
Thank you!
---
pretyping/patternops.ml | 4 +---
test-suite/bugs/closed/5487.v | 9 ---------
2 files changed, 1 insertion(+), 12 deletions(-)
delete mode 100644 test-suite/bugs/closed/5487.v
diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml
index 75d3ed30ba..2090aad8a0 100644
--- a/pretyping/patternops.ml
+++ b/pretyping/patternops.ml
@@ -160,9 +160,7 @@ let pattern_of_constr env sigma t =
let ty = Evarutil.nf_evar sigma (existential_type sigma ev) in
let () = ignore (pattern_of_constr env ty) in
assert (not b); PMeta (Some id)
- | Evar_kinds.GoalEvar | Evar_kinds.VarInstance _ ->
- (* These are the two evar kinds used for existing goals *)
- (* see Proofview.mark_in_evm *)
+ | Evar_kinds.GoalEvar ->
PEvar (evk,Array.map (pattern_of_constr env) ctxt)
| _ ->
let ty = Evarutil.nf_evar sigma (existential_type sigma ev) in
diff --git a/test-suite/bugs/closed/5487.v b/test-suite/bugs/closed/5487.v
deleted file mode 100644
index 9b995f4503..0000000000
--- a/test-suite/bugs/closed/5487.v
+++ /dev/null
@@ -1,9 +0,0 @@
-(* Was a collision between an ltac pattern variable and an evar *)
-
-Goal forall n, exists m, n = m :> nat.
-Proof.
- eexists.
- Fail match goal with
- | [ |- ?x = ?y ]
- => match x with y => idtac end
- end.
--
cgit v1.2.3
From fdd5a8452bd2da22ffd1cab3b1888f2261f193b9 Mon Sep 17 00:00:00 2001
From: Gaetan Gilbert
Date: Sun, 30 Apr 2017 13:10:48 +0200
Subject: Functional choice <-> [inhabited] and [forall] commute
---
theories/Init/Logic.v | 5 +++++
theories/Init/Specif.v | 11 +++++++++++
theories/Logic/ChoiceFacts.v | 33 +++++++++++++++++++++++++++++++++
3 files changed, 49 insertions(+)
diff --git a/theories/Init/Logic.v b/theories/Init/Logic.v
index 9ae9dde28d..3eefe9a849 100644
--- a/theories/Init/Logic.v
+++ b/theories/Init/Logic.v
@@ -609,6 +609,11 @@ Proof.
destruct 1; auto.
Qed.
+Lemma inhabited_covariant (A B : Type) : (A -> B) -> inhabited A -> inhabited B.
+Proof.
+ intros f [x];exact (inhabits (f x)).
+Qed.
+
(** Declaration of stepl and stepr for eq and iff *)
Lemma eq_stepl : forall (A : Type) (x y z : A), x = y -> x = z -> z = y.
diff --git a/theories/Init/Specif.v b/theories/Init/Specif.v
index 2cc2ecbc20..43a441fc51 100644
--- a/theories/Init/Specif.v
+++ b/theories/Init/Specif.v
@@ -207,6 +207,17 @@ Definition sig2_eta {A P Q} (p : { a : A | P a & Q a })
: p = exist2 _ _ (proj1_sig (sig_of_sig2 p)) (proj2_sig (sig_of_sig2 p)) (proj3_sig p).
Proof. destruct p; reflexivity. Defined.
+(** [exists x : A, B] is equivalent to [inhabited {x : A | B}] *)
+Lemma exists_to_inhabited_sig {A P} : (exists x : A, P x) -> inhabited {x : A | P x}.
+Proof.
+ intros [x y]. exact (inhabits (exist _ x y)).
+Qed.
+
+Lemma inhabited_sig_to_exists {A P} : inhabited {x : A | P x} -> exists x : A, P x.
+Proof.
+ intros [[x y]];exists x;exact y.
+Qed.
+
(** [sumbool] is a boolean type equipped with the justification of
their value *)
diff --git a/theories/Logic/ChoiceFacts.v b/theories/Logic/ChoiceFacts.v
index 07e8b6ef8d..f1f20606b1 100644
--- a/theories/Logic/ChoiceFacts.v
+++ b/theories/Logic/ChoiceFacts.v
@@ -760,6 +760,39 @@ Proof.
destruct Heq using eq_indd; trivial.
Qed.
+(** Functional choice can be reformulated as a property on [inhabited] *)
+
+Definition InhabitedForallCommute_on (A : Type) (B : A -> Type) :=
+ (forall x, inhabited (B x)) -> inhabited (forall x, B x).
+
+Notation InhabitedForallCommute :=
+ (forall A (B : A -> Type), InhabitedForallCommute_on B).
+
+Theorem functional_choice_to_inhabited_forall_commute :
+ FunctionalChoice -> InhabitedForallCommute.
+Proof.
+ intros choose0 A B Hinhab.
+ pose proof (non_dep_dep_functional_choice choose0) as choose;clear choose0.
+ assert (Hexists : forall x, exists _ : B x, True).
+ { intros x;apply inhabited_sig_to_exists.
+ refine (inhabited_covariant _ (Hinhab x)).
+ intros y;exists y;exact I. }
+ apply choose in Hexists.
+ destruct Hexists as [f _].
+ exact (inhabits f).
+Qed.
+
+Theorem inhabited_forall_commute_to_functional_choice :
+ InhabitedForallCommute -> FunctionalChoice.
+Proof.
+ intros choose A B R Hexists.
+ assert (Hinhab : forall x, inhabited {y : B | R x y}).
+ { intros x;apply exists_to_inhabited_sig;trivial. }
+ apply choose in Hinhab. destruct Hinhab as [f].
+ exists (fun x => proj1_sig (f x)).
+ exact (fun x => proj2_sig (f x)).
+Qed.
+
(** ** Reification of dependent and non dependent functional relation are equivalent *)
Definition DependentFunctionalRelReification_on (A:Type) (B:A -> Type) :=
--
cgit v1.2.3
--
cgit v1.2.3
From 12f1c409daf2cdbd7d0323f0d61723819532b362 Mon Sep 17 00:00:00 2001
From: Hugo Herbelin
Date: Mon, 1 May 2017 16:56:25 +0200
Subject: Really fixing #2602 which was wrongly working because of #5487 hiding
the cause.
The cause was a missing evar/evar clause in ltac pattern-matching
function (constr_matching.ml).
---
pretyping/constr_matching.ml | 3 +++
test-suite/success/ltac.v | 13 +++++++++++++
2 files changed, 16 insertions(+)
diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml
index 3c47cfdc4b..afdf601c21 100644
--- a/pretyping/constr_matching.ml
+++ b/pretyping/constr_matching.ml
@@ -347,6 +347,9 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels
| PFix c1, Fix _ when eq_constr (mkFix c1) cT -> subst
| PCoFix c1, CoFix _ when eq_constr (mkCoFix c1) cT -> subst
+ | PEvar (c1,args1), Evar (c2,args2) when c1 = c2 ->
+ Array.fold_left2 (sorec ctx env) subst args1 args2
+
| _ -> raise PatternMatchingFailure
in
diff --git a/test-suite/success/ltac.v b/test-suite/success/ltac.v
index ce90990594..d7ec092d76 100644
--- a/test-suite/success/ltac.v
+++ b/test-suite/success/ltac.v
@@ -317,3 +317,16 @@ let T := constr:(fun a b : nat => a) in
end.
exact (eq_refl n).
Qed.
+
+(* A variant of #2602 which was wrongly succeeding because "a", bound to
+ "?m", was then internally turned into a "_" in the second matching *)
+
+Goal exists m, S m > 0.
+eexists.
+Fail match goal with
+ | |- context [ S ?a ] =>
+ match goal with
+ | |- S a > a => idtac
+ end
+end.
+Abort.
--
cgit v1.2.3
From c3aec655a8a33fff676c79e12888d193cc2e237b Mon Sep 17 00:00:00 2001
From: Hugo Herbelin
Date: Mon, 1 May 2017 16:58:38 +0200
Subject: Fixing #5487 (v8.5 regression on ltac-matching expressions with
evars).
The fix follows an invariant enforced in proofview.ml on the kind of
evars that are goals or that occur in goals.
One day, evar kinds will need a little cleaning...
PS: This is a second attempt, completing db28e82 which was missing the
case PEvar in constr_matching.ml. Indeed the attached fix to #5487
alone made #2602 failing, revealing that the real cause for #2602 was
actually not fixed and that if the test for #2602 was working it was
because of #5487 hiding the real problem in #2602.
---
pretyping/patternops.ml | 4 +++-
test-suite/bugs/closed/5487.v | 9 +++++++++
2 files changed, 12 insertions(+), 1 deletion(-)
create mode 100644 test-suite/bugs/closed/5487.v
diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml
index 2090aad8a0..75d3ed30ba 100644
--- a/pretyping/patternops.ml
+++ b/pretyping/patternops.ml
@@ -160,7 +160,9 @@ let pattern_of_constr env sigma t =
let ty = Evarutil.nf_evar sigma (existential_type sigma ev) in
let () = ignore (pattern_of_constr env ty) in
assert (not b); PMeta (Some id)
- | Evar_kinds.GoalEvar ->
+ | Evar_kinds.GoalEvar | Evar_kinds.VarInstance _ ->
+ (* These are the two evar kinds used for existing goals *)
+ (* see Proofview.mark_in_evm *)
PEvar (evk,Array.map (pattern_of_constr env) ctxt)
| _ ->
let ty = Evarutil.nf_evar sigma (existential_type sigma ev) in
diff --git a/test-suite/bugs/closed/5487.v b/test-suite/bugs/closed/5487.v
new file mode 100644
index 0000000000..9b995f4503
--- /dev/null
+++ b/test-suite/bugs/closed/5487.v
@@ -0,0 +1,9 @@
+(* Was a collision between an ltac pattern variable and an evar *)
+
+Goal forall n, exists m, n = m :> nat.
+Proof.
+ eexists.
+ Fail match goal with
+ | [ |- ?x = ?y ]
+ => match x with y => idtac end
+ end.
--
cgit v1.2.3
From 2553e4bf5735a2bd127832e2d26609c6a8096fb7 Mon Sep 17 00:00:00 2001
From: Pierre-Marie Pédrot
Date: Mon, 1 May 2017 17:24:29 +0200
Subject: Removing dead code in Autorewrite.
Since 260965d, an imperative code was semantically the identity because the
closure allocation was not performed at the right moment. Because of it
intricacy, I cannot really tell the original motivations of this piece of
code, although it looks like it was for there for pretty-printing of errors.
Anyway, both because the code was dubious and its effect not observed, it
cannot hurt to remove it.
---
tactics/autorewrite.ml | 44 ++++++--------------------------------------
1 file changed, 6 insertions(+), 38 deletions(-)
diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml
index e58ec5a31f..3c430cb174 100644
--- a/tactics/autorewrite.ml
+++ b/tactics/autorewrite.ml
@@ -127,45 +127,13 @@ let autorewrite ?(conds=Naive) tac_main lbas =
(Proofview.tclUNIT()) lbas))
let autorewrite_multi_in ?(conds=Naive) idl tac_main lbas =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
(* let's check at once if id exists (to raise the appropriate error) *)
- let _ = List.map (fun id -> Tacmach.New.pf_get_hyp id gl) idl in
- let general_rewrite_in id =
- let id = ref id in
- let to_be_cleared = ref false in
- fun dir cstr tac gl ->
- let last_hyp_id =
- match Tacmach.pf_hyps gl with
- d :: _ -> Context.Named.Declaration.get_id d
- | _ -> (* even the hypothesis id is missing *)
- raise (Logic.RefinerError (Logic.NoSuchHyp !id))
- in
- let gl' = Proofview.V82.of_tactic (general_rewrite_in dir AllOccurrences true ~tac:(tac, conds) false !id cstr false) gl in
- let gls = gl'.Evd.it in
- match gls with
- g::_ ->
- (match Environ.named_context_of_val (Goal.V82.hyps gl'.Evd.sigma g) with
- d ::_ ->
- let lastid = Context.Named.Declaration.get_id d in
- if not (Id.equal last_hyp_id lastid) then
- begin
- let gl'' =
- if !to_be_cleared then
- tclTHEN (fun _ -> gl') (tclTRY (Proofview.V82.of_tactic (clear [!id]))) gl
- else gl' in
- id := lastid ;
- to_be_cleared := true ;
- gl''
- end
- else
- begin
- to_be_cleared := false ;
- gl'
- end
- | _ -> assert false) (* there must be at least an hypothesis *)
- | _ -> assert false (* rewriting cannot complete a proof *)
- in
- let general_rewrite_in x y z w = Proofview.V82.tactic (general_rewrite_in x y (EConstr.of_constr z) w) in
+ let _ = List.map (fun id -> Tacmach.New.pf_get_hyp id gl) idl in
+ let general_rewrite_in id dir cstr tac =
+ let cstr = EConstr.of_constr cstr in
+ general_rewrite_in dir AllOccurrences true ~tac:(tac, conds) false id cstr false
+ in
Tacticals.New.tclMAP (fun id ->
Tacticals.New.tclREPEAT_MAIN (Proofview.tclPROGRESS
(List.fold_left (fun tac bas ->
--
cgit v1.2.3
From f6856c5022ef27cdc492daadd1301cfcad025b01 Mon Sep 17 00:00:00 2001
From: Paul Steckler
Date: Mon, 1 May 2017 11:34:00 -0400
Subject: remove unneeded -emacs flag to coq-prog-args
---
test-suite/bugs/closed/3080.v | 2 +-
test-suite/bugs/closed/3323.v | 2 +-
test-suite/bugs/closed/3332.v | 2 +-
test-suite/bugs/closed/3346.v | 2 +-
test-suite/bugs/closed/3348.v | 2 +-
test-suite/bugs/closed/3352.v | 2 +-
test-suite/bugs/closed/3375.v | 2 +-
test-suite/bugs/closed/3393.v | 2 +-
test-suite/bugs/closed/3427.v | 2 +-
test-suite/bugs/closed/3539.v | 2 +-
test-suite/bugs/closed/3612.v | 2 +-
test-suite/bugs/closed/3649.v | 2 +-
test-suite/bugs/closed/3881.v | 2 +-
test-suite/bugs/closed/3956.v | 2 +-
test-suite/bugs/closed/4089.v | 2 +-
test-suite/bugs/closed/4121.v | 2 +-
test-suite/bugs/closed/4394.v | 2 +-
test-suite/bugs/closed/4400.v | 2 +-
test-suite/bugs/closed/4456.v | 2 +-
test-suite/bugs/closed/4527.v | 2 +-
test-suite/bugs/closed/4533.v | 2 +-
test-suite/bugs/closed/4544.v | 2 +-
test-suite/bugs/closed/4656.v | 2 +-
test-suite/bugs/closed/4673.v | 2 +-
test-suite/bugs/closed/4722.v | 2 +-
test-suite/bugs/closed/4727.v | 2 +-
test-suite/bugs/closed/4733.v | 2 +-
test-suite/bugs/closed/4769.v | 2 +-
test-suite/bugs/closed/4780.v | 2 +-
test-suite/bugs/closed/4785_compat_85.v | 2 +-
test-suite/bugs/closed/4811.v | 2 +-
test-suite/bugs/closed/4818.v | 2 +-
test-suite/bugs/closed/5198.v | 2 +-
test-suite/bugs/closed/HoTT_coq_012.v | 2 +-
test-suite/bugs/closed/HoTT_coq_032.v | 2 +-
test-suite/bugs/closed/HoTT_coq_053.v | 2 +-
test-suite/bugs/closed/HoTT_coq_054.v | 1 -
test-suite/bugs/closed/HoTT_coq_055.v | 2 +-
test-suite/bugs/closed/HoTT_coq_059.v | 2 +-
test-suite/bugs/closed/HoTT_coq_062.v | 2 +-
test-suite/bugs/closed/HoTT_coq_097.v | 2 +-
test-suite/bugs/closed/HoTT_coq_107.v | 2 +-
test-suite/bugs/closed/HoTT_coq_108.v | 2 +-
test-suite/bugs/closed/HoTT_coq_123.v | 2 +-
test-suite/bugs/opened/4803.v | 2 +-
test-suite/output-modulo-time/ltacprof.v | 2 +-
test-suite/output-modulo-time/ltacprof_cutoff.v | 2 +-
test-suite/output/ErrorInModule.v | 2 +-
test-suite/output/ErrorInSection.v | 2 +-
test-suite/success/Compat84.v | 2 +-
50 files changed, 49 insertions(+), 50 deletions(-)
diff --git a/test-suite/bugs/closed/3080.v b/test-suite/bugs/closed/3080.v
index 7d0dc090e1..36ab7ff599 100644
--- a/test-suite/bugs/closed/3080.v
+++ b/test-suite/bugs/closed/3080.v
@@ -1,4 +1,4 @@
-(* -*- coq-prog-args: ("-emacs" "-nois") -*- *)
+(* -*- coq-prog-args: ("-nois") -*- *)
Delimit Scope type_scope with type.
Delimit Scope function_scope with function.
diff --git a/test-suite/bugs/closed/3323.v b/test-suite/bugs/closed/3323.v
index 22b1603b60..4622634eaa 100644
--- a/test-suite/bugs/closed/3323.v
+++ b/test-suite/bugs/closed/3323.v
@@ -1,5 +1,5 @@
Require Import TestSuite.admit.
-(* -*- coq-prog-args: ("-emacs" "-indices-matter") -*- *)
+(* -*- coq-prog-args: ("-indices-matter") -*- *)
(* File reduced by coq-bug-finder from original input, then from 5302 lines to 4649 lines, then from 4660 lines to 355 lines, then from 360 lines to 269 lines, then from 269 lines to 175 lines, then from 144 lines to 119 lines, then from 297 lines to 117 lines, then from 95 lines to 79 lines, then from 82 lines to 68 lines *)
Set Universe Polymorphism.
diff --git a/test-suite/bugs/closed/3332.v b/test-suite/bugs/closed/3332.v
index d86470cdee..a3564bfcce 100644
--- a/test-suite/bugs/closed/3332.v
+++ b/test-suite/bugs/closed/3332.v
@@ -1,4 +1,4 @@
-(* -*- coq-prog-args: ("-emacs" "-time") -*- *)
+(* -*- coq-prog-args: ("-time") -*- *)
Definition foo : True.
Proof.
Abort. (* Toplevel input, characters 15-21:
diff --git a/test-suite/bugs/closed/3346.v b/test-suite/bugs/closed/3346.v
index 638404f2cb..09bd789345 100644
--- a/test-suite/bugs/closed/3346.v
+++ b/test-suite/bugs/closed/3346.v
@@ -1,4 +1,4 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *)
Monomorphic Inductive paths (A : Type) (a : A) : A -> Type := idpath : paths A a a.
(* This should fail with -indices-matter *)
Fail Check paths nat O O : Prop.
diff --git a/test-suite/bugs/closed/3348.v b/test-suite/bugs/closed/3348.v
index d9ac09d8d3..904de68964 100644
--- a/test-suite/bugs/closed/3348.v
+++ b/test-suite/bugs/closed/3348.v
@@ -1,4 +1,4 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *)
Set Universe Polymorphism.
Set Printing Universes.
Inductive Empty : Set := .
diff --git a/test-suite/bugs/closed/3352.v b/test-suite/bugs/closed/3352.v
index f8113e4c78..555accfd51 100644
--- a/test-suite/bugs/closed/3352.v
+++ b/test-suite/bugs/closed/3352.v
@@ -20,7 +20,7 @@ while it is expected to have type "IsHProp (* Top.17 *) Empty"
Defined.
Module B.
-(* -*- coq-prog-args: ("-emacs" "-indices-matter") -*- *)
+(* -*- coq-prog-args: ("-indices-matter") -*- *)
(* File reduced by coq-bug-finder from original input, then from 7725 lines to 78 lines, then from 51 lines to 13 lines *)
Set Universe Polymorphism.
Inductive paths {A} (a : A) : A -> Type := idpath : paths a a where "x = y" := (@paths _ x y) : type_scope.
diff --git a/test-suite/bugs/closed/3375.v b/test-suite/bugs/closed/3375.v
index d7ce02ea87..1e0c8e61f4 100644
--- a/test-suite/bugs/closed/3375.v
+++ b/test-suite/bugs/closed/3375.v
@@ -1,5 +1,5 @@
Require Import TestSuite.admit.
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-impredicative-set") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-impredicative-set") -*- *)
(* File reduced by coq-bug-finder from original input, then from 5968 lines to 11933 lines, then from 11239 lines to 11231 lines, then from 10365 lines to 446 lines, then from 456 lines to 379 lines, then from 391 lines to 373 lines, then from 369 lines to 351 lines, then from 350 lines to 340 lines, then from 348 lines to 320 lines, then from 328 lines to 302 lines, then from 330 lines to 45 lines *)
Set Universe Polymorphism.
diff --git a/test-suite/bugs/closed/3393.v b/test-suite/bugs/closed/3393.v
index f7ab5f76a5..ae8e41e29e 100644
--- a/test-suite/bugs/closed/3393.v
+++ b/test-suite/bugs/closed/3393.v
@@ -1,5 +1,5 @@
Require Import TestSuite.admit.
-(* -*- coq-prog-args: ("-emacs" "-indices-matter") -*- *)
+(* -*- coq-prog-args: ("-indices-matter") -*- *)
(* File reduced by coq-bug-finder from original input, then from 8760 lines to 7519 lines, then from 7050 lines to 909 lines, then from 846 lines to 578 lines, then from 497 lines to 392 lines, then from 365 lines to 322 lines, then from 252 lines to 205 lines, then from 215 lines to 204 lines, then from 210 lines to 182 lines, then from 175 lines to 157 lines *)
Set Universe Polymorphism.
Axiom admit : forall {T}, T.
diff --git a/test-suite/bugs/closed/3427.v b/test-suite/bugs/closed/3427.v
index 374a53928d..9a57ca7703 100644
--- a/test-suite/bugs/closed/3427.v
+++ b/test-suite/bugs/closed/3427.v
@@ -1,5 +1,5 @@
Require Import TestSuite.admit.
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *)
(* File reduced by coq-bug-finder from original input, then from 0 lines to 7171 lines, then from 7184 lines to 558 lines, then from 556 lines to 209 lines *)
Generalizable All Variables.
Set Universe Polymorphism.
diff --git a/test-suite/bugs/closed/3539.v b/test-suite/bugs/closed/3539.v
index d258bb31f8..b0c4b23702 100644
--- a/test-suite/bugs/closed/3539.v
+++ b/test-suite/bugs/closed/3539.v
@@ -1,4 +1,4 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *)
(* File reduced by coq-bug-finder from original input, then from 11678 lines to 11330 lines, then from 10721 lines to 9544 lines, then from 9549 lines to 794 lines, then from 810 lines to 785 lines, then from 628 lines to 246 lines, then from 220 lines to 89 lines, then from 80 lines to 47 lines *)
(* coqc version trunk (August 2014) compiled on Aug 22 2014 4:17:28 with OCaml 4.01.0
coqtop version cagnode17:/afs/csail.mit.edu/u/j/jgross/coq-trunk,trunk (a67cc6941434124465f20b14a1256f2ede31a60e) *)
diff --git a/test-suite/bugs/closed/3612.v b/test-suite/bugs/closed/3612.v
index a547685070..90182a4881 100644
--- a/test-suite/bugs/closed/3612.v
+++ b/test-suite/bugs/closed/3612.v
@@ -1,4 +1,4 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter" "-nois") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-indices-matter" "-nois") -*- *)
(* File reduced by coq-bug-finder from original input, then from 3595 lines to 3518 lines, then from 3133 lines to 2950 lines, then from 2911 lines to 415 lines, then from 431 lines to 407 \
lines, then from 421 lines to 428 lines, then from 444 lines to 429 lines, then from 434 lines to 66 lines, then from 163 lines to 48 lines *)
(* coqc version trunk (September 2014) compiled on Sep 11 2014 14:48:8 with OCaml 4.01.0
diff --git a/test-suite/bugs/closed/3649.v b/test-suite/bugs/closed/3649.v
index fc4c171e2c..367d380ec3 100644
--- a/test-suite/bugs/closed/3649.v
+++ b/test-suite/bugs/closed/3649.v
@@ -1,4 +1,4 @@
-(* -*- coq-prog-args: ("-emacs" "-nois") -*- *)
+(* -*- coq-prog-args: ("-nois") -*- *)
(* File reduced by coq-bug-finder from original input, then from 9518 lines to 404 lines, then from 410 lines to 208 lines, then from 162 lines to 77 lines *)
(* coqc version trunk (September 2014) compiled on Sep 18 2014 21:0:5 with OCaml 4.01.0
coqtop version cagnode16:/afs/csail.mit.edu/u/j/jgross/coq-trunk,trunk (07e4438bd758c2ced8caf09a6961ccd77d84e42b) *)
diff --git a/test-suite/bugs/closed/3881.v b/test-suite/bugs/closed/3881.v
index a327bbf2a9..bb6af6a66c 100644
--- a/test-suite/bugs/closed/3881.v
+++ b/test-suite/bugs/closed/3881.v
@@ -1,4 +1,4 @@
-(* -*- coq-prog-args: ("-emacs" "-nois" "-R" "../theories" "Coq") -*- *)
+(* -*- coq-prog-args: ("-nois" "-R" "../theories" "Coq") -*- *)
(* File reduced by coq-bug-finder from original input, then from 2236 lines to 1877 lines, then from 1652 lines to 160 lines, then from 102 lines to 34 lines *)
(* coqc version trunk (December 2014) compiled on Dec 23 2014 22:6:43 with OCaml 4.01.0
coqtop version cagnode15:/afs/csail.mit.edu/u/j/jgross/coq-trunk,trunk (90ed6636dea41486ddf2cc0daead83f9f0788163) *)
diff --git a/test-suite/bugs/closed/3956.v b/test-suite/bugs/closed/3956.v
index c19a2d4a06..66dee702aa 100644
--- a/test-suite/bugs/closed/3956.v
+++ b/test-suite/bugs/closed/3956.v
@@ -1,4 +1,4 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter"); mode: visual-line -*- *)
+(* -*- mode: coq; coq-prog-args: ("-indices-matter"); mode: visual-line -*- *)
Set Universe Polymorphism.
Set Primitive Projections.
Close Scope nat_scope.
diff --git a/test-suite/bugs/closed/4089.v b/test-suite/bugs/closed/4089.v
index e4d76732a3..fc1c504f14 100644
--- a/test-suite/bugs/closed/4089.v
+++ b/test-suite/bugs/closed/4089.v
@@ -1,6 +1,6 @@
Unset Strict Universe Declaration.
Require Import TestSuite.admit.
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *)
(* File reduced by coq-bug-finder from original input, then from 6522 lines to 318 lines, then from 1139 lines to 361 lines *)
(* coqc version 8.5beta1 (February 2015) compiled on Feb 23 2015 18:32:3 with OCaml 4.01.0
coqtop version cagnode15:/afs/csail.mit.edu/u/j/jgross/coq-8.5,v8.5 (ebfc19d792492417b129063fb511aa423e9d9e08) *)
diff --git a/test-suite/bugs/closed/4121.v b/test-suite/bugs/closed/4121.v
index d34a2b8b1b..a8bf950ff2 100644
--- a/test-suite/bugs/closed/4121.v
+++ b/test-suite/bugs/closed/4121.v
@@ -1,5 +1,5 @@
Unset Strict Universe Declaration.
-(* -*- coq-prog-args: ("-emacs" "-nois") -*- *)
+(* -*- coq-prog-args: ("-nois") -*- *)
(* File reduced by coq-bug-finder from original input, then from 830 lines to 47 lines, then from 25 lines to 11 lines *)
(* coqc version 8.5beta1 (March 2015) compiled on Mar 11 2015 18:51:36 with OCaml 4.01.0
coqtop version cagnode15:/afs/csail.mit.edu/u/j/jgross/coq-8.5,v8.5 (8dbfee5c5f897af8186cb1bdfb04fd4f88eca677) *)
diff --git a/test-suite/bugs/closed/4394.v b/test-suite/bugs/closed/4394.v
index 60c9354597..1ad81345db 100644
--- a/test-suite/bugs/closed/4394.v
+++ b/test-suite/bugs/closed/4394.v
@@ -1,4 +1,4 @@
-(* -*- coq-prog-args: ("-emacs" "-compat" "8.4") -*- *)
+(* -*- coq-prog-args: ("-compat" "8.4") -*- *)
Require Import Equality List.
Inductive Foo (I : Type -> Type) (A : Type) : Type :=
diff --git a/test-suite/bugs/closed/4400.v b/test-suite/bugs/closed/4400.v
index 5c23f8404b..a89cf0cbc3 100644
--- a/test-suite/bugs/closed/4400.v
+++ b/test-suite/bugs/closed/4400.v
@@ -1,4 +1,4 @@
-(* -*- coq-prog-args: ("-emacs" "-require" "Coq.Compat.Coq84" "-compat" "8.4") -*- *)
+(* -*- coq-prog-args: ("-require" "Coq.Compat.Coq84" "-compat" "8.4") -*- *)
Require Import Coq.Lists.List Coq.Logic.JMeq Program.Equality.
Set Printing Universes.
Inductive Foo (I : Type -> Type) (A : Type) : Type :=
diff --git a/test-suite/bugs/closed/4456.v b/test-suite/bugs/closed/4456.v
index a32acf789c..56a7b4f6e9 100644
--- a/test-suite/bugs/closed/4456.v
+++ b/test-suite/bugs/closed/4456.v
@@ -1,4 +1,4 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-R" "." "Fiat" "-top" "BooleanRecognizerMin" "-R" "." "Top") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-R" "." "Fiat" "-top" "BooleanRecognizerMin" "-R" "." "Top") -*- *)
(* File reduced by coq-bug-finder from original input, then from 2475 lines to 729 lines, then from 746 lines to 658 lines, then from 675 lines to 658 lines *)
(* coqc version 8.5beta3 (November 2015) compiled on Nov 11 2015 18:23:0 with OCaml 4.01.0
coqtop version 8.5beta3 (November 2015) *)
diff --git a/test-suite/bugs/closed/4527.v b/test-suite/bugs/closed/4527.v
index 08628377f0..4fab9d44f3 100644
--- a/test-suite/bugs/closed/4527.v
+++ b/test-suite/bugs/closed/4527.v
@@ -1,4 +1,4 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-nois" "-indices-matter" "-R" "." "Top" "-top" "bug_bad_univ_length_01") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-nois" "-indices-matter" "-R" "." "Top" "-top" "bug_bad_univ_length_01") -*- *)
(* File reduced by coq-bug-finder from original input, then from 1199 lines to
430 lines, then from 444 lines to 430 lines, then from 964 lines to 255 lines,
then from 269 lines to 255 lines *)
diff --git a/test-suite/bugs/closed/4533.v b/test-suite/bugs/closed/4533.v
index ae17fb145d..ccef1c3040 100644
--- a/test-suite/bugs/closed/4533.v
+++ b/test-suite/bugs/closed/4533.v
@@ -1,4 +1,4 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-nois" "-indices-matter" "-R" "." "Top" "-top" "bug_lex_wrong_rewrite_02") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-nois" "-indices-matter" "-R" "." "Top" "-top" "bug_lex_wrong_rewrite_02") -*- *)
(* File reduced by coq-bug-finder from original input, then from 1125 lines to
346 lines, then from 360 lines to 346 lines, then from 822 lines to 271 lines,
then from 285 lines to 271 lines *)
diff --git a/test-suite/bugs/closed/4544.v b/test-suite/bugs/closed/4544.v
index da140c9318..82f1e3fe73 100644
--- a/test-suite/bugs/closed/4544.v
+++ b/test-suite/bugs/closed/4544.v
@@ -1,4 +1,4 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-nois" "-indices-matter" "-R" "." "Top" "-top" "bug_oog_looping_rewrite_01") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-nois" "-indices-matter" "-R" "." "Top" "-top" "bug_oog_looping_rewrite_01") -*- *)
(* File reduced by coq-bug-finder from original input, then from 2553 lines to 1932 lines, then from 1946 lines to 1932 lines, then from 2467 lines to 1002 lines, then from 1016 lines to 1002 lines *)
(* coqc version 8.5 (January 2016) compiled on Jan 23 2016 16:15:22 with OCaml 4.01.0
coqtop version 8.5 (January 2016) *)
diff --git a/test-suite/bugs/closed/4656.v b/test-suite/bugs/closed/4656.v
index c89a86d634..a59eed2c86 100644
--- a/test-suite/bugs/closed/4656.v
+++ b/test-suite/bugs/closed/4656.v
@@ -1,4 +1,4 @@
-(* -*- coq-prog-args: ("-emacs" "-compat" "8.4") -*- *)
+(* -*- coq-prog-args: ("-compat" "8.4") -*- *)
Goal True.
constructor 1.
Qed.
diff --git a/test-suite/bugs/closed/4673.v b/test-suite/bugs/closed/4673.v
index 1ae5081851..10e48db6dd 100644
--- a/test-suite/bugs/closed/4673.v
+++ b/test-suite/bugs/closed/4673.v
@@ -1,4 +1,4 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-R" "." "Fiat" "-top" "BooleanRecognizerOptimized" "-R" "." "Top") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-R" "." "Fiat" "-top" "BooleanRecognizerOptimized" "-R" "." "Top") -*- *)
(* File reduced by coq-bug-finder from original input, then from 2407 lines to 22 lines, then from 528 lines to 35 lines, then from 331 lines to 42 lines, then from 56 lines to 42 lines, then from 63 lines to 46 lines, then from 60 lines to 46 lines *) (* coqc version 8.5 (February 2016) compiled on Feb 21 2016 15:26:16 with OCaml 4.02.3
coqtop version 8.5 (February 2016) *)
Axiom proof_admitted : False.
diff --git a/test-suite/bugs/closed/4722.v b/test-suite/bugs/closed/4722.v
index f047624c84..2d41828f19 100644
--- a/test-suite/bugs/closed/4722.v
+++ b/test-suite/bugs/closed/4722.v
@@ -1 +1 @@
-(* -*- coq-prog-args: ("-emacs" "-R" "4722" "Foo") -*- *)
+(* -*- coq-prog-args: ("-R" "4722" "Foo") -*- *)
diff --git a/test-suite/bugs/closed/4727.v b/test-suite/bugs/closed/4727.v
index 3854bbffdd..cfb4548d2c 100644
--- a/test-suite/bugs/closed/4727.v
+++ b/test-suite/bugs/closed/4727.v
@@ -1,4 +1,4 @@
-(* -*- coq-prog-args: ("-emacs" "-compat" "8.4") -*- *)
+(* -*- coq-prog-args: ("-compat" "8.4") -*- *)
Goal forall (P : Set) (l : P) (P0 : Set) (w w0 : P0) (T : Type) (a : P * T) (o : P -> option P0),
(forall (l1 l2 : P) (w1 : P0), o l1 = Some w1 -> o l2 = Some w1 -> l1 = l2) ->
o l = Some w -> o (fst a) = Some w0 -> {w = w0} + {w <> w0} -> False.
diff --git a/test-suite/bugs/closed/4733.v b/test-suite/bugs/closed/4733.v
index a6ebda61cf..a90abd71cf 100644
--- a/test-suite/bugs/closed/4733.v
+++ b/test-suite/bugs/closed/4733.v
@@ -1,4 +1,4 @@
-(* -*- coq-prog-args: ("-emacs" "-compat" "8.4") -*- *)
+(* -*- coq-prog-args: ("-compat" "8.4") -*- *)
(*Suppose a user wants to declare a new list-like notation with support for singletons in both 8.4 and 8.5. If they use*)
Require Import Coq.Lists.List.
Require Import Coq.Vectors.Vector.
diff --git a/test-suite/bugs/closed/4769.v b/test-suite/bugs/closed/4769.v
index d87906f313..33a1d1a50b 100644
--- a/test-suite/bugs/closed/4769.v
+++ b/test-suite/bugs/closed/4769.v
@@ -1,5 +1,5 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-nois" "-R" "." "Top" "-top" "bug_hom_anom_10") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-nois" "-R" "." "Top" "-top" "bug_hom_anom_10") -*- *)
(* File reduced by coq-bug-finder from original input, then from 156 lines to 41 lines, then from 237 lines to 45 lines, then from 163 lines to 66 lines, then from 342 lines to 121 lines, then from 353 lines to 184 lines, then from 343 lines to 255 lines, then from 435 lines to 322 lines, then from 475 lines to 351 lines, then from 442 lines to 377 lines, then from 505 lines to 410 lines, then from 591 lines to 481 lines, then from 596 lines to 535 lines, then from 647 lines to 570 lines, then from 669 lines to 596 lines, then from 687 lines to 620 lines, then from 728 lines to 652 lines, then from 1384 lines to 683 lines, then from 984 lines to 707 lines, then from 1124 lines to 734 lines, then from 775 lines to 738 lines, then from 950 lines to 763 lines, then from 857 lines to 798 lines, then from 983 lines to 752 lines, then from 1598 lines to 859 lines, then from 873 lines to 859 lines, then from 875 lines to 862 lines, then from 901 lines to 863 lines, then from 1047 lines to 865 lines, then from 929 lines to 871 lines, then from 989 lines to 884 lines, then from 900 lines to 884 lines, then from 884 lines to 751 lines, then from 763 lines to 593 lines, then from 482 lines to 232 lines, then from 416 lines to 227 lines, then from 290 lines to 231 lines, then from 348 lines to 235 lines, then from 249 lines to 235 lines, then from 249 lines to 172 lines, then from 186 lines to 172 lines, then from 140 lines to 113 lines, then from 127 lines to 113 lines *) (* coqc version trunk (June 2016) compiled on Jun 2 2016 10:16:20 with OCaml 4.02.3
coqtop version trunk (June 2016) *)
diff --git a/test-suite/bugs/closed/4780.v b/test-suite/bugs/closed/4780.v
index 4cec43184c..71a51c6312 100644
--- a/test-suite/bugs/closed/4780.v
+++ b/test-suite/bugs/closed/4780.v
@@ -1,4 +1,4 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-R" "." "Top" "-top" "bug_bad_induction_01") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-R" "." "Top" "-top" "bug_bad_induction_01") -*- *)
(* File reduced by coq-bug-finder from original input, then from 1889 lines to 144 lines, then from 158 lines to 144 lines *)
(* coqc version 8.5pl1 (April 2016) compiled on Apr 18 2016 14:48:5 with OCaml 4.02.3
coqtop version 8.5pl1 (April 2016) *)
diff --git a/test-suite/bugs/closed/4785_compat_85.v b/test-suite/bugs/closed/4785_compat_85.v
index 9d65840acd..bbb34f465c 100644
--- a/test-suite/bugs/closed/4785_compat_85.v
+++ b/test-suite/bugs/closed/4785_compat_85.v
@@ -1,4 +1,4 @@
-(* -*- coq-prog-args: ("-emacs" "-compat" "8.5") -*- *)
+(* -*- coq-prog-args: ("-compat" "8.5") -*- *)
Require Coq.Lists.List Coq.Vectors.Vector.
Require Coq.Compat.Coq85.
diff --git a/test-suite/bugs/closed/4811.v b/test-suite/bugs/closed/4811.v
index a914962629..fe6e65a0f0 100644
--- a/test-suite/bugs/closed/4811.v
+++ b/test-suite/bugs/closed/4811.v
@@ -2,7 +2,7 @@
(* Submitted by Jason Gross *)
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-R" "src" "Crypto" "-R" "Bedrock" "Bedrock" "-R" "coqprime-8.5/Coqprime" "Coqprime" "-top" "GF255192") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-R" "src" "Crypto" "-R" "Bedrock" "Bedrock" "-R" "coqprime-8.5/Coqprime" "Coqprime" "-top" "GF255192") -*- *)
(* File reduced by coq-bug-finder from original input, then from 162 lines to 23 lines, then from 245 lines to 95 lines, then from 198 lines to 101 lines, then from 654 lines to 452 lines, then from 591 lines to 505 lines, then from 1770 lines to 580 lines, then from 2238 lines to 1715 lines, then from 1776 lines to 1738 lines, then from 1750 lines to 1679 lines, then from 1693 lines to 1679 lines *)
(* coqc version 8.5pl1 (April 2016) compiled on Apr 18 2016 14:48:5 with OCaml 4.02.3
coqtop version 8.5pl1 (April 2016) *)
diff --git a/test-suite/bugs/closed/4818.v b/test-suite/bugs/closed/4818.v
index 904abb2287..e411ce62f0 100644
--- a/test-suite/bugs/closed/4818.v
+++ b/test-suite/bugs/closed/4818.v
@@ -1,4 +1,4 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-R" "." "Prob" "-top" "Product") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-R" "." "Prob" "-top" "Product") -*- *)
(* File reduced by coq-bug-finder from original input, then from 391 lines to 77 lines, then from 857 lines to 119 lines, then from 1584 lines to 126 lines, then from 362 lines to 135 lines, then from 149 lines to 135 lines *)
(* coqc version 8.5pl1 (June 2016) compiled on Jun 9 2016 17:27:17 with OCaml 4.02.3
coqtop version 8.5pl1 (June 2016) *)
diff --git a/test-suite/bugs/closed/5198.v b/test-suite/bugs/closed/5198.v
index 7254afb429..72722f5f6d 100644
--- a/test-suite/bugs/closed/5198.v
+++ b/test-suite/bugs/closed/5198.v
@@ -1,4 +1,4 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-boot" "-nois") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-boot" "-nois") -*- *)
(* File reduced by coq-bug-finder from original input, then from 286 lines to
27 lines, then from 224 lines to 53 lines, then from 218 lines to 56 lines,
then from 269 lines to 180 lines, then from 132 lines to 48 lines, then from
diff --git a/test-suite/bugs/closed/HoTT_coq_012.v b/test-suite/bugs/closed/HoTT_coq_012.v
index a3c697f8ca..420aaf9745 100644
--- a/test-suite/bugs/closed/HoTT_coq_012.v
+++ b/test-suite/bugs/closed/HoTT_coq_012.v
@@ -1,4 +1,4 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *)
Definition UU := Type.
Inductive toto (B : UU) : UU := c (x : B).
diff --git a/test-suite/bugs/closed/HoTT_coq_032.v b/test-suite/bugs/closed/HoTT_coq_032.v
index 3f5d7b0215..39a7103d1b 100644
--- a/test-suite/bugs/closed/HoTT_coq_032.v
+++ b/test-suite/bugs/closed/HoTT_coq_032.v
@@ -1,4 +1,4 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-xml") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-xml") -*- *)
Set Implicit Arguments.
Generalizable All Variables.
Set Asymmetric Patterns.
diff --git a/test-suite/bugs/closed/HoTT_coq_053.v b/test-suite/bugs/closed/HoTT_coq_053.v
index e2bf1dbedb..263dec4857 100644
--- a/test-suite/bugs/closed/HoTT_coq_053.v
+++ b/test-suite/bugs/closed/HoTT_coq_053.v
@@ -1,4 +1,4 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *)
Set Printing Universes.
Set Implicit Arguments.
Generalizable All Variables.
diff --git a/test-suite/bugs/closed/HoTT_coq_054.v b/test-suite/bugs/closed/HoTT_coq_054.v
index c687965937..635024e983 100644
--- a/test-suite/bugs/closed/HoTT_coq_054.v
+++ b/test-suite/bugs/closed/HoTT_coq_054.v
@@ -1,4 +1,3 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs") -*- *)
Inductive Empty : Prop := .
Inductive paths {A : Type} (a : A) : A -> Type :=
diff --git a/test-suite/bugs/closed/HoTT_coq_055.v b/test-suite/bugs/closed/HoTT_coq_055.v
index a250987714..561b7e557d 100644
--- a/test-suite/bugs/closed/HoTT_coq_055.v
+++ b/test-suite/bugs/closed/HoTT_coq_055.v
@@ -1,5 +1,5 @@
Require Import TestSuite.admit.
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *)
Set Universe Polymorphism.
Inductive Empty : Prop := .
diff --git a/test-suite/bugs/closed/HoTT_coq_059.v b/test-suite/bugs/closed/HoTT_coq_059.v
index 9c7e394dc3..2e6c735cf5 100644
--- a/test-suite/bugs/closed/HoTT_coq_059.v
+++ b/test-suite/bugs/closed/HoTT_coq_059.v
@@ -1,4 +1,4 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *)
Set Universe Polymorphism.
Inductive eq {A} (x : A) : A -> Type := eq_refl : eq x x.
diff --git a/test-suite/bugs/closed/HoTT_coq_062.v b/test-suite/bugs/closed/HoTT_coq_062.v
index 90d1d18306..e01f73f1b3 100644
--- a/test-suite/bugs/closed/HoTT_coq_062.v
+++ b/test-suite/bugs/closed/HoTT_coq_062.v
@@ -1,6 +1,6 @@
Unset Strict Universe Declaration.
Require Import TestSuite.admit.
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *)
(* File reduced by coq-bug-finder from 5012 lines to 4659 lines, then from 4220 lines to 501 lines, then from 513 lines to 495 lines, then from 513 lines to 495 lines, then from 412 lines to 79 lines, then from 412 lines to 79 lines. *)
Set Universe Polymorphism.
Definition admit {T} : T.
diff --git a/test-suite/bugs/closed/HoTT_coq_097.v b/test-suite/bugs/closed/HoTT_coq_097.v
index 38e8007b6c..f2b2e57b9c 100644
--- a/test-suite/bugs/closed/HoTT_coq_097.v
+++ b/test-suite/bugs/closed/HoTT_coq_097.v
@@ -1,4 +1,4 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *)
Set Universe Polymorphism.
Set Printing Universes.
Inductive Empty : Set := .
diff --git a/test-suite/bugs/closed/HoTT_coq_107.v b/test-suite/bugs/closed/HoTT_coq_107.v
index 7c1ab8dc2c..fa4072a8f6 100644
--- a/test-suite/bugs/closed/HoTT_coq_107.v
+++ b/test-suite/bugs/closed/HoTT_coq_107.v
@@ -1,4 +1,4 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-nois" "-R" "../theories" "Coq") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-nois" "-R" "../theories" "Coq") -*- *)
(* File reduced by coq-bug-finder from 4897 lines to 2605 lines, then from 2297 lines to 236 lines, then from 239 lines to 137 lines, then from 118 lines to 67 lines, then from 520 lines to 76 lines. *)
(** Note: The bug here is the same as the #113, that is, HoTT_coq_113.v *)
Require Import Coq.Init.Logic.
diff --git a/test-suite/bugs/closed/HoTT_coq_108.v b/test-suite/bugs/closed/HoTT_coq_108.v
index b6c0da76ba..ea4bcd8b45 100644
--- a/test-suite/bugs/closed/HoTT_coq_108.v
+++ b/test-suite/bugs/closed/HoTT_coq_108.v
@@ -1,5 +1,5 @@
Require Import TestSuite.admit.
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *)
(* NOTE: This bug is only triggered with -load-vernac-source / in interactive mode. *)
(* File reduced by coq-bug-finder from 139 lines to 124 lines. *)
Set Universe Polymorphism.
diff --git a/test-suite/bugs/closed/HoTT_coq_123.v b/test-suite/bugs/closed/HoTT_coq_123.v
index 6ee6e65323..cd9cad4064 100644
--- a/test-suite/bugs/closed/HoTT_coq_123.v
+++ b/test-suite/bugs/closed/HoTT_coq_123.v
@@ -1,5 +1,5 @@
Require Import TestSuite.admit.
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") *)
+(* -*- mode: coq; coq-prog-args: ("-indices-matter") *)
(* File reduced by coq-bug-finder from original input, then from 4988 lines to 856 lines, then from 648 lines to 398 lines, then from 401 lines to 332 lines, then from 287 lines to 250 lines, then from 257 lines to 241 lines, then from 223 lines to 175 lines *)
Set Universe Polymorphism.
Set Asymmetric Patterns.
diff --git a/test-suite/bugs/opened/4803.v b/test-suite/bugs/opened/4803.v
index 4530548b8f..4541f13d01 100644
--- a/test-suite/bugs/opened/4803.v
+++ b/test-suite/bugs/opened/4803.v
@@ -1,4 +1,4 @@
-(* -*- coq-prog-args: ("-emacs" "-compat" "8.4") -*- *)
+(* -*- coq-prog-args: ("-compat" "8.4") -*- *)
(*Suppose a user wants to declare a new list-like notation with support for singletons in both 8.4 and 8.5. If they use*)
Require Import Coq.Lists.List.
Require Import Coq.Vectors.Vector.
diff --git a/test-suite/output-modulo-time/ltacprof.v b/test-suite/output-modulo-time/ltacprof.v
index 6611db70e2..1e9e919797 100644
--- a/test-suite/output-modulo-time/ltacprof.v
+++ b/test-suite/output-modulo-time/ltacprof.v
@@ -1,4 +1,4 @@
-(* -*- coq-prog-args: ("-emacs" "-profile-ltac-cutoff" "0.0") -*- *)
+(* -*- coq-prog-args: ("-profile-ltac-cutoff" "0.0") -*- *)
Ltac sleep' := do 100 (do 100 (do 100 idtac)).
Ltac sleep := sleep'.
diff --git a/test-suite/output-modulo-time/ltacprof_cutoff.v b/test-suite/output-modulo-time/ltacprof_cutoff.v
index 50131470eb..3dad6271af 100644
--- a/test-suite/output-modulo-time/ltacprof_cutoff.v
+++ b/test-suite/output-modulo-time/ltacprof_cutoff.v
@@ -1,4 +1,4 @@
-(* -*- coq-prog-args: ("-emacs" "-profile-ltac") -*- *)
+(* -*- coq-prog-args: ("-profile-ltac") -*- *)
Require Coq.ZArith.BinInt.
Ltac sleep := do 50 (idtac; let sleep := (eval vm_compute in Coq.ZArith.BinInt.Z.div_eucl) in idtac).
diff --git a/test-suite/output/ErrorInModule.v b/test-suite/output/ErrorInModule.v
index e69e23276b..b2e3c3e923 100644
--- a/test-suite/output/ErrorInModule.v
+++ b/test-suite/output/ErrorInModule.v
@@ -1,4 +1,4 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-quick") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-quick") -*- *)
Module M.
Definition foo := nonexistent.
End M.
diff --git a/test-suite/output/ErrorInSection.v b/test-suite/output/ErrorInSection.v
index 3036f8f05b..505c5ce378 100644
--- a/test-suite/output/ErrorInSection.v
+++ b/test-suite/output/ErrorInSection.v
@@ -1,4 +1,4 @@
-(* -*- mode: coq; coq-prog-args: ("-emacs" "-quick") -*- *)
+(* -*- mode: coq; coq-prog-args: ("-quick") -*- *)
Section S.
Definition foo := nonexistent.
End S.
diff --git a/test-suite/success/Compat84.v b/test-suite/success/Compat84.v
index db6348fa17..732a024fc1 100644
--- a/test-suite/success/Compat84.v
+++ b/test-suite/success/Compat84.v
@@ -1,4 +1,4 @@
-(* -*- coq-prog-args: ("-emacs" "-compat" "8.4") -*- *)
+(* -*- coq-prog-args: ("-compat" "8.4") -*- *)
Goal True.
solve [ constructor 1 ]. Undo.
--
cgit v1.2.3
From cea40f37ab638031b9d5c6434ee5651a16ea1f3e Mon Sep 17 00:00:00 2001
From: Hugo Herbelin
Date: Mon, 1 May 2017 09:04:17 +0200
Subject: Fixing Set Rewriting Schemes bugs introduced in v8.5.
- Fixing a typo introduced in 31dbba5f.
- Adapting to computation of universe constraints in pretyping.
- Adding a regression test.
---
tactics/eqschemes.ml | 8 +++++++-
test-suite/success/Scheme.v | 23 +++++++++++++++++++++++
2 files changed, 30 insertions(+), 1 deletion(-)
diff --git a/tactics/eqschemes.ml b/tactics/eqschemes.ml
index 1a45217a4a..e39159fb82 100644
--- a/tactics/eqschemes.ml
+++ b/tactics/eqschemes.ml
@@ -89,6 +89,11 @@ let get_coq_eq ctx =
with Not_found ->
error "eq not found."
+let univ_of_eq env eq =
+ match kind_of_term (Retyping.get_type_of env Evd.empty eq) with
+ | Prod (_,t,_) -> (match kind_of_term t with Sort (Type u) -> u | _ -> assert false)
+ | _ -> assert false
+
(**********************************************************************)
(* Check if an inductive type [ind] has the form *)
(* *)
@@ -744,7 +749,7 @@ let build_congr env (eq,refl,ctx) ind =
let ty = get_type (lookup_rel (mip.mind_nrealargs - i + 1) env_with_arity) in
let constrsign,ccl = decompose_prod_assum mip.mind_nf_lc.(0) in
let _,constrargs = decompose_app ccl in
- if Int.equal (Context.Rel.length constrsign) (Context.Rel.length mib.mind_params_ctxt) then
+ if not (Int.equal (Context.Rel.length constrsign) (Context.Rel.length mib.mind_params_ctxt)) then
error "Constructor must have no arguments";
let b = List.nth constrargs (i + mib.mind_nparams - 1) in
let varB = fresh env (Id.of_string "B") in
@@ -752,6 +757,7 @@ let build_congr env (eq,refl,ctx) ind =
let varf = fresh env (Id.of_string "f") in
let ci = make_case_info (Global.env()) ind RegularStyle in
let uni, ctx = Universes.extend_context (Universes.new_global_univ ()) ctx in
+ let ctx = (fst ctx, Univ.enforce_leq uni (univ_of_eq env eq) (snd ctx)) in
let c =
my_it_mkLambda_or_LetIn paramsctxt
(mkNamedLambda varB (mkSort (Type uni))
diff --git a/test-suite/success/Scheme.v b/test-suite/success/Scheme.v
index dd5aa81d1d..855f26698c 100644
--- a/test-suite/success/Scheme.v
+++ b/test-suite/success/Scheme.v
@@ -2,3 +2,26 @@
Scheme Induction for eq Sort Prop.
Check eq_ind_dep.
+
+(* This was broken in v8.5 *)
+
+Set Rewriting Schemes.
+Inductive myeq A (a:A) : A -> Prop := myrefl : myeq A a a.
+Unset Rewriting Schemes.
+
+Check myeq_rect.
+Check myeq_ind.
+Check myeq_rec.
+Check myeq_congr.
+Check myeq_sym_internal.
+Check myeq_rew.
+Check myeq_rew_dep.
+Check myeq_rew_fwd_dep.
+Check myeq_rew_r.
+Check internal_myeq_sym_involutive.
+Check myeq_rew_r_dep.
+Check myeq_rew_fwd_r_dep.
+
+Set Rewriting Schemes.
+Inductive myeq_true : bool -> Prop := myrefl_true : myeq_true true.
+Unset Rewriting Schemes.
--
cgit v1.2.3
From 8dc2adfd102185c6c3c4b61709f2b55aefab2641 Mon Sep 17 00:00:00 2001
From: Jason Gross
Date: Sat, 29 Apr 2017 13:07:45 -0400
Subject: Add bmsherman/topology to the ci
This development of @bmsherman tests universe polymorphism and setoid
rewriting in type, and should build with v8.6 and trunk.
---
.travis.yml | 1 +
Makefile.ci | 2 +-
dev/ci/ci-basic-overlay.sh | 6 ++++++
dev/ci/ci-formal-topology.sh | 28 ++++++++++++++++++++++++++++
4 files changed, 36 insertions(+), 1 deletion(-)
create mode 100755 dev/ci/ci-formal-topology.sh
diff --git a/.travis.yml b/.travis.yml
index 77aac23b78..3ed547bb15 100644
--- a/.travis.yml
+++ b/.travis.yml
@@ -41,6 +41,7 @@ env:
- TEST_TARGET="ci-math-classes"
- TEST_TARGET="ci-math-comp"
- TEST_TARGET="ci-sf"
+ - TEST_TARGET="ci-formal-topology"
- TEST_TARGET="ci-unimath"
- TEST_TARGET="ci-vst"
# Not ready yet for 8.7
diff --git a/Makefile.ci b/Makefile.ci
index 4c4606aff6..0136852180 100644
--- a/Makefile.ci
+++ b/Makefile.ci
@@ -1,7 +1,7 @@
CI_TARGETS=ci-all ci-hott ci-math-comp ci-compcert ci-sf ci-cpdt \
ci-color ci-math-classes ci-tlc ci-fiat-crypto ci-fiat-parsers \
ci-coquelicot ci-flocq ci-iris-coq ci-metacoq ci-geocoq \
- ci-unimath ci-vst ci-bedrock-src ci-bedrock-facade
+ ci-unimath ci-vst ci-bedrock-src ci-bedrock-facade ci-formal-topology
.PHONY: $(CI_TARGETS)
diff --git a/dev/ci/ci-basic-overlay.sh b/dev/ci/ci-basic-overlay.sh
index e0851816ce..5da13c1008 100644
--- a/dev/ci/ci-basic-overlay.sh
+++ b/dev/ci/ci-basic-overlay.sh
@@ -106,6 +106,12 @@
: ${bedrock_facade_CI_BRANCH:=master}
: ${bedrock_facade_CI_GITURL:=https://github.com/JasonGross/bedrock.git}
+########################################################################
+# formal-topology
+########################################################################
+: ${formal_topology_CI_BRANCH:=ci}
+: ${formal_topology_CI_GITURL:=https://github.com/bmsherman/topology.git}
+
########################################################################
# CoLoR
########################################################################
diff --git a/dev/ci/ci-formal-topology.sh b/dev/ci/ci-formal-topology.sh
new file mode 100755
index 0000000000..ecb36349fb
--- /dev/null
+++ b/dev/ci/ci-formal-topology.sh
@@ -0,0 +1,28 @@
+#!/usr/bin/env bash
+
+ci_dir="$(dirname "$0")"
+source ${ci_dir}/ci-common.sh
+
+math_classes_CI_DIR=${CI_BUILD_DIR}/math-classes
+
+Corn_CI_DIR=${CI_BUILD_DIR}/corn
+
+formal_topology_CI_DIR=${CI_BUILD_DIR}/formal-topology
+
+# Setup Math-Classes
+
+git_checkout ${math_classes_CI_BRANCH} ${math_classes_CI_GITURL} ${math_classes_CI_DIR}
+
+( cd ${math_classes_CI_DIR} && make -j ${NJOBS} && make install )
+
+# Setup Corn
+
+git_checkout ${Corn_CI_BRANCH} ${Corn_CI_GITURL} ${Corn_CI_DIR}
+
+( cd ${Corn_CI_DIR} && make -j ${NJOBS} && make install )
+
+# Setup formal-topology
+
+git_checkout ${formal_topology_CI_BRANCH} ${formal_topology_CI_GITURL} ${formal_topology_CI_DIR}
+
+( cd ${formal_topology_CI_DIR} && make -j ${NJOBS} )
--
cgit v1.2.3
From 824caa1f93563ab9437fb238459d757447a0aa12 Mon Sep 17 00:00:00 2001
From: Hugo Herbelin
Date: Mon, 1 May 2017 14:15:55 +0200
Subject: Avoiding registering files from _build_ci when not calling
Makefile.ci.
---
Makefile | 3 ++-
1 file changed, 2 insertions(+), 1 deletion(-)
diff --git a/Makefile b/Makefile
index e84d5e3775..e50a1b18f7 100644
--- a/Makefile
+++ b/Makefile
@@ -52,7 +52,8 @@ FIND_VCS_CLAUSE:='(' \
-name '.bzr' -o \
-name 'debian' -o \
-name "$${GIT_DIR}" -o \
- -name '_build' \
+ -name '_build' -o \
+ -name '_build_ci' \
')' -prune -o
define find
--
cgit v1.2.3
From 4a84961049f4f00897ae72a13954edbcc9aaba5e Mon Sep 17 00:00:00 2001
From: Théo Zimmermann
Date: Wed, 3 May 2017 10:23:46 +0200
Subject: Fix outdated description in RefMan.
---
doc/refman/RefMan-pro.tex | 2 +-
1 file changed, 1 insertion(+), 1 deletion(-)
diff --git a/doc/refman/RefMan-pro.tex b/doc/refman/RefMan-pro.tex
index c37367de5b..4c333379bd 100644
--- a/doc/refman/RefMan-pro.tex
+++ b/doc/refman/RefMan-pro.tex
@@ -118,7 +118,7 @@ the current proof and declare the initial goal as an axiom.
\subsection[\tt Proof {\term}.]{\tt Proof {\term}.\comindex{Proof}
\label{BeginProof}}
This command applies in proof editing mode. It is equivalent to {\tt
- exact {\term}; Save.} That is, you have to give the full proof in
+ exact {\term}. Qed.} That is, you have to give the full proof in
one gulp, as a proof term (see Section~\ref{exact}).
\variant {\tt Proof.}
--
cgit v1.2.3
From e9b745af47ba3386724b874e3fd74b6dad33b015 Mon Sep 17 00:00:00 2001
From: Gaetan Gilbert
Date: Thu, 6 Apr 2017 22:48:32 +0200
Subject: Allow flexible anonymous universes in instances and sorts.
The addition to the test suite showcases the usage.
---
intf/misctypes.mli | 4 +--
library/declare.ml | 5 ++--
parsing/g_constr.ml4 | 8 ++---
pretyping/detyping.ml | 7 +++--
pretyping/miscops.ml | 2 +-
pretyping/pretyping.ml | 61 +++++++++++++++++++++------------------
printing/ppconstr.ml | 8 ++---
test-suite/success/polymorphism.v | 32 +++++++++++++++++++-
8 files changed, 82 insertions(+), 45 deletions(-)
diff --git a/intf/misctypes.mli b/intf/misctypes.mli
index 33dc2a335c..7c2dc5177c 100644
--- a/intf/misctypes.mli
+++ b/intf/misctypes.mli
@@ -48,8 +48,8 @@ type 'a glob_sort_gen =
| GProp (** representation of [Prop] literal *)
| GSet (** representation of [Set] literal *)
| GType of 'a (** representation of [Type] literal *)
-type sort_info = string Loc.located list
-type level_info = string Loc.located option
+type sort_info = Name.t Loc.located list
+type level_info = Name.t Loc.located option
type glob_sort = sort_info glob_sort_gen
type glob_level = level_info glob_sort_gen
diff --git a/library/declare.ml b/library/declare.ml
index 31c9c24bc3..91e0cb44b3 100644
--- a/library/declare.ml
+++ b/library/declare.ml
@@ -514,11 +514,10 @@ let do_constraint poly l =
match x with
| GProp -> Loc.dummy_loc, (false, Univ.Level.prop)
| GSet -> Loc.dummy_loc, (false, Univ.Level.set)
- | GType None ->
+ | GType None | GType (Some (_, Anonymous)) ->
user_err ~hdr:"Constraint"
(str "Cannot declare constraints on anonymous universes")
- | GType (Some (loc, id)) ->
- let id = Id.of_string id in
+ | GType (Some (loc, Name id)) ->
let names, _ = Global.global_universe_names () in
try loc, Idmap.find id names
with Not_found ->
diff --git a/parsing/g_constr.ml4 b/parsing/g_constr.ml4
index 0f2ed88fea..15f100c3b0 100644
--- a/parsing/g_constr.ml4
+++ b/parsing/g_constr.ml4
@@ -146,12 +146,12 @@ GEXTEND Gram
[ [ "Set" -> GSet
| "Prop" -> GProp
| "Type" -> GType []
- | "Type"; "@{"; u = universe; "}" -> GType (List.map (fun (loc,x) -> (loc, Id.to_string x)) u)
+ | "Type"; "@{"; u = universe; "}" -> GType u
] ]
;
universe:
- [ [ IDENT "max"; "("; ids = LIST1 identref SEP ","; ")" -> ids
- | id = identref -> [id]
+ [ [ IDENT "max"; "("; ids = LIST1 name SEP ","; ")" -> ids
+ | id = name -> [id]
] ]
;
lconstr:
@@ -298,7 +298,7 @@ GEXTEND Gram
[ [ "Set" -> GSet
| "Prop" -> GProp
| "Type" -> GType None
- | id = identref -> GType (Some (fst id, Id.to_string (snd id)))
+ | id = name -> GType (Some id)
] ]
;
fix_constr:
diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml
index 483e2b4320..8a90a3f9bc 100644
--- a/pretyping/detyping.ml
+++ b/pretyping/detyping.ml
@@ -422,7 +422,9 @@ let detype_sort sigma = function
| Type u ->
GType
(if !print_universes
- then [dl, Pp.string_of_ppcmds (Univ.Universe.pr_with (Termops.pr_evd_level sigma) u)]
+ then
+ let u = Pp.string_of_ppcmds (Univ.Universe.pr_with (Termops.pr_evd_level sigma) u) in
+ [dl, Name.mk_name (Id.of_string_soft u)]
else [])
type binder_kind = BProd | BLambda | BLetIn
@@ -434,7 +436,8 @@ let detype_anonymous = ref (fun loc n -> anomaly ~label:"detype" (Pp.str "index
let set_detype_anonymous f = detype_anonymous := f
let detype_level sigma l =
- GType (Some (dl, Pp.string_of_ppcmds (Termops.pr_evd_level sigma l)))
+ let l = Pp.string_of_ppcmds (Termops.pr_evd_level sigma l) in
+ GType (Some (dl, Name.mk_name (Id.of_string_soft l)))
let detype_instance sigma l =
let l = EInstance.kind sigma l in
diff --git a/pretyping/miscops.ml b/pretyping/miscops.ml
index 7fe81c9a43..1669f8334b 100644
--- a/pretyping/miscops.ml
+++ b/pretyping/miscops.ml
@@ -30,7 +30,7 @@ let smartmap_cast_type f c =
let glob_sort_eq g1 g2 = match g1, g2 with
| GProp, GProp -> true
| GSet, GSet -> true
-| GType l1, GType l2 -> List.equal (fun x y -> CString.equal (snd x) (snd y)) l1 l2
+| GType l1, GType l2 -> List.equal (fun x y -> Names.Name.equal (snd x) (snd y)) l1 l2
| _ -> false
let intro_pattern_naming_eq nam1 nam2 = match nam1, nam2 with
diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml
index 68ef976592..767e4be35b 100644
--- a/pretyping/pretyping.ml
+++ b/pretyping/pretyping.ml
@@ -192,45 +192,50 @@ let _ =
optwrite = (:=) Universes.set_minimization })
(** Miscellaneous interpretation functions *)
-let interp_universe_level_name evd (loc,s) =
- let names, _ = Global.global_universe_names () in
- if CString.string_contains ~where:s ~what:"." then
- match List.rev (CString.split '.' s) with
- | [] -> anomaly (str"Invalid universe name " ++ str s)
- | n :: dp ->
- let num = int_of_string n in
- let dp = DirPath.make (List.map Id.of_string dp) in
- let level = Univ.Level.make dp num in
- let evd =
- try Evd.add_global_univ evd level
- with UGraph.AlreadyDeclared -> evd
- in evd, level
- else
- try
- let level = Evd.universe_of_name evd s in
- evd, level
- with Not_found ->
- try
- let id = try Id.of_string s with _ -> raise Not_found in
- evd, snd (Idmap.find id names)
- with Not_found ->
- if not (is_strict_universe_declarations ()) then
- new_univ_level_variable ~loc ~name:s univ_rigid evd
- else user_err ~loc ~hdr:"interp_universe_level_name"
- (Pp.(str "Undeclared universe: " ++ str s))
+let interp_universe_level_name ~anon_rigidity evd (loc,s) =
+ match s with
+ | Anonymous ->
+ new_univ_level_variable ~loc anon_rigidity evd
+ | Name s ->
+ let s = Id.to_string s in
+ let names, _ = Global.global_universe_names () in
+ if CString.string_contains ~where:s ~what:"." then
+ match List.rev (CString.split '.' s) with
+ | [] -> anomaly (str"Invalid universe name " ++ str s)
+ | n :: dp ->
+ let num = int_of_string n in
+ let dp = DirPath.make (List.map Id.of_string dp) in
+ let level = Univ.Level.make dp num in
+ let evd =
+ try Evd.add_global_univ evd level
+ with UGraph.AlreadyDeclared -> evd
+ in evd, level
+ else
+ try
+ let level = Evd.universe_of_name evd s in
+ evd, level
+ with Not_found ->
+ try
+ let id = try Id.of_string s with _ -> raise Not_found in
+ evd, snd (Idmap.find id names)
+ with Not_found ->
+ if not (is_strict_universe_declarations ()) then
+ new_univ_level_variable ~loc ~name:s univ_rigid evd
+ else user_err ~loc ~hdr:"interp_universe_level_name"
+ (Pp.(str "Undeclared universe: " ++ str s))
let interp_universe ?loc evd = function
| [] -> let evd, l = new_univ_level_variable ?loc univ_rigid evd in
evd, Univ.Universe.make l
| l ->
List.fold_left (fun (evd, u) l ->
- let evd', l = interp_universe_level_name evd l in
+ let evd', l = interp_universe_level_name ~anon_rigidity:univ_flexible_alg evd l in
(evd', Univ.sup u (Univ.Universe.make l)))
(evd, Univ.Universe.type0m) l
let interp_level_info loc evd : Misctypes.level_info -> _ = function
| None -> new_univ_level_variable ~loc univ_rigid evd
- | Some (loc,s) -> interp_universe_level_name evd (loc,s)
+ | Some (loc,s) -> interp_universe_level_name ~anon_rigidity:univ_flexible evd (loc,s)
let interp_sort ?loc evd = function
| GProp -> evd, Prop Null
diff --git a/printing/ppconstr.ml b/printing/ppconstr.ml
index 3041ac259e..b546c39aec 100644
--- a/printing/ppconstr.ml
+++ b/printing/ppconstr.ml
@@ -151,8 +151,8 @@ let tag_var = tag Tag.variable
let pr_univ l =
match l with
- | [_,x] -> str x
- | l -> str"max(" ++ prlist_with_sep (fun () -> str",") (fun x -> str (snd x)) l ++ str")"
+ | [_,x] -> pr_name x
+ | l -> str"max(" ++ prlist_with_sep (fun () -> str",") (fun x -> pr_name (snd x)) l ++ str")"
let pr_univ_annot pr x = str "@{" ++ pr x ++ str "}"
@@ -166,7 +166,7 @@ let tag_var = tag Tag.variable
| GProp -> tag_type (str "Prop")
| GSet -> tag_type (str "Set")
| GType None -> tag_type (str "Type")
- | GType (Some (_, u)) -> tag_type (str u)
+ | GType (Some (_, u)) -> tag_type (pr_name u)
let pr_qualid sp =
let (sl, id) = repr_qualid sp in
@@ -191,7 +191,7 @@ let tag_var = tag Tag.variable
tag_type (str "Set")
| GType u ->
(match u with
- | Some (_,u) -> str u
+ | Some (_,u) -> pr_name u
| None -> tag_type (str "Type"))
let pr_universe_instance l =
diff --git a/test-suite/success/polymorphism.v b/test-suite/success/polymorphism.v
index 878875bd92..0a58fe89a1 100644
--- a/test-suite/success/polymorphism.v
+++ b/test-suite/success/polymorphism.v
@@ -321,4 +321,34 @@ Definition unwrap' := fun (X : Type) (b : box X) => let (unw) := b in unw.
Fail Definition bad : False := TypeNeqSmallType.paradox (unwrap' Type (wrap _
Type)) eq_refl.
-End Hurkens'.
\ No newline at end of file
+End Hurkens'.
+
+Module Anonymous.
+ Set Universe Polymorphism.
+
+ Definition defaultid := (fun x => x) : Type -> Type.
+ Definition collapseid := defaultid@{_ _}.
+ Check collapseid@{_}.
+
+ Definition anonid := (fun x => x) : Type -> Type@{_}.
+ Check anonid@{_}.
+
+ Definition defaultalg := Type : Type.
+ Definition usedefaultalg := defaultalg@{_ _}.
+ Check usedefaultalg@{_ _}.
+
+ Definition anonalg := (fun x => x) (Type : Type@{_}).
+ Check anonalg@{_}.
+
+ Definition unrelated@{i j} := nat.
+ Definition useunrelated := unrelated@{_ _}.
+ Check useunrelated@{_ _}.
+
+ Definition inthemiddle@{i j k} :=
+ let _ := defaultid@{i j} in
+ defaultalg@{k j}.
+ (* i <= j < k *)
+ Definition collapsethemiddle := inthemiddle@{i _ j}.
+ Check collapsethemiddle@{_ _}.
+
+End Anonymous.
--
cgit v1.2.3
From 4361c1ed9ac5646055f9f0eecc4a003d720c1994 Mon Sep 17 00:00:00 2001
From: Gaetan Gilbert
Date: Wed, 12 Apr 2017 13:29:16 +0200
Subject: Type@{_} should not produce a flexible algebraic universe.
Otherwise [(fun x => x) (Type : Type@{_})] becomes
[(fun x : Type@{i+1} => x) (Type@{i} : Type@{i+1})]
breaking the invariant that terms do not contain algebraic universes
(at the lambda abstraction).
---
pretyping/pretyping.ml | 3 ++-
test-suite/success/polymorphism.v | 10 +++++-----
2 files changed, 7 insertions(+), 6 deletions(-)
diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml
index 767e4be35b..4886423bd0 100644
--- a/pretyping/pretyping.ml
+++ b/pretyping/pretyping.ml
@@ -229,7 +229,8 @@ let interp_universe ?loc evd = function
evd, Univ.Universe.make l
| l ->
List.fold_left (fun (evd, u) l ->
- let evd', l = interp_universe_level_name ~anon_rigidity:univ_flexible_alg evd l in
+ (* [univ_flexible_alg] can produce algebraic universes in terms *)
+ let evd', l = interp_universe_level_name ~anon_rigidity:univ_flexible evd l in
(evd', Univ.sup u (Univ.Universe.make l)))
(evd, Univ.Universe.type0m) l
diff --git a/test-suite/success/polymorphism.v b/test-suite/success/polymorphism.v
index 0a58fe89a1..66ff55edcb 100644
--- a/test-suite/success/polymorphism.v
+++ b/test-suite/success/polymorphism.v
@@ -333,12 +333,12 @@ Module Anonymous.
Definition anonid := (fun x => x) : Type -> Type@{_}.
Check anonid@{_}.
- Definition defaultalg := Type : Type.
- Definition usedefaultalg := defaultalg@{_ _}.
+ Definition defaultalg := (fun x : Type => x) (Type : Type).
+ Definition usedefaultalg := defaultalg@{_ _ _}.
Check usedefaultalg@{_ _}.
- Definition anonalg := (fun x => x) (Type : Type@{_}).
- Check anonalg@{_}.
+ Definition anonalg := (fun x : Type@{_} => x) (Type : Type).
+ Check anonalg@{_ _}.
Definition unrelated@{i j} := nat.
Definition useunrelated := unrelated@{_ _}.
@@ -346,7 +346,7 @@ Module Anonymous.
Definition inthemiddle@{i j k} :=
let _ := defaultid@{i j} in
- defaultalg@{k j}.
+ anonalg@{k j}.
(* i <= j < k *)
Definition collapsethemiddle := inthemiddle@{i _ j}.
Check collapsethemiddle@{_ _}.
--
cgit v1.2.3
From 8adfa0e5290056b7683a3a8b778ca16182a1eb3d Mon Sep 17 00:00:00 2001
From: Gaetan Gilbert
Date: Tue, 2 May 2017 14:43:32 +0200
Subject: Reorganize comment documentation of ChoiceFacts.v
Shortnames and natural language descriptions of principles are moved
next to each principle.
The table of contents is moved to after the principle definitions.
Extra definitions are moved to the definition section (eg DependentFunctionalChoice)
Compatibility notations have been moved to the end of the file.
Details:
The following used to be announced but were neither defined or used,
and have been removed:
- OAC!
- Ext_pred = extensionality of predicates
- Ext_fun_prop_repr = choice of a representative among extensional functions to Prop
GuardedFunctionalRelReification was announced with shortname GAC! but
shortname GFR_fun was used next to it. Only the former has been retained.
Shortnames and descriptions have been invented for
InhabitedForallCommute DependentFunctionalRelReification
ExtensionalPropositionRepresentative ExtensionalFunctionRepresentative
Some modification of headlines
---
theories/Logic/ChoiceFacts.v | 283 ++++++++++++++++++++-----------------------
1 file changed, 131 insertions(+), 152 deletions(-)
diff --git a/theories/Logic/ChoiceFacts.v b/theories/Logic/ChoiceFacts.v
index f1f20606b1..116897f4ce 100644
--- a/theories/Logic/ChoiceFacts.v
+++ b/theories/Logic/ChoiceFacts.v
@@ -8,94 +8,9 @@
(************************************************************************)
(** Some facts and definitions concerning choice and description in
- intuitionistic logic.
-
-We investigate the relations between the following choice and
-description principles
-
-- AC_rel = relational form of the (non extensional) axiom of choice
- (a "set-theoretic" axiom of choice)
-- AC_fun = functional form of the (non extensional) axiom of choice
- (a "type-theoretic" axiom of choice)
-- DC_fun = functional form of the dependent axiom of choice
-- ACw_fun = functional form of the countable axiom of choice
-- AC! = functional relation reification
- (known as axiom of unique choice in topos theory,
- sometimes called principle of definite description in
- the context of constructive type theory, sometimes
- called axiom of no choice)
-
-- AC_fun_repr = functional choice of a representative in an equivalence class
-- AC_fun_setoid_gen = functional form of the general form of the (so-called
- extensional) axiom of choice over setoids
-- AC_fun_setoid = functional form of the (so-called extensional) axiom of
- choice from setoids
-- AC_fun_setoid_simple = functional form of the (so-called extensional) axiom of
- choice from setoids on locally compatible relations
-
-- GAC_rel = guarded relational form of the (non extensional) axiom of choice
-- GAC_fun = guarded functional form of the (non extensional) axiom of choice
-- GAC! = guarded functional relation reification
-
-- OAC_rel = "omniscient" relational form of the (non extensional) axiom of choice
-- OAC_fun = "omniscient" functional form of the (non extensional) axiom of choice
- (called AC* in Bell [[Bell]])
-- OAC!
-
-- ID_iota = intuitionistic definite description
-- ID_epsilon = intuitionistic indefinite description
-
-- D_iota = (weakly classical) definite description principle
-- D_epsilon = (weakly classical) indefinite description principle
-
-- PI = proof irrelevance
-- IGP = independence of general premises
- (an unconstrained generalisation of the constructive principle of
- independence of premises)
-- Drinker = drinker's paradox (small form)
- (called Ex in Bell [[Bell]])
-- EM = excluded-middle
-
-- Ext_pred_repr = choice of a representative among extensional predicates
-- Ext_pred = extensionality of predicates
-- Ext_fun_prop_repr = choice of a representative among extensional functions to Prop
-
-We let also
-
-- IPL_2 = 2nd-order impredicative minimal predicate logic (with ex. quant.)
-- IPL^2 = 2nd-order functional minimal predicate logic (with ex. quant.)
-- IPL_2^2 = 2nd-order impredicative, 2nd-order functional minimal pred. logic (with ex. quant.)
-
-with no prerequisite on the non-emptiness of domains
-
-Table of contents
-
-1. Definitions
-
-2. IPL_2^2 |- AC_rel + AC! = AC_fun
-
-3.1. typed IPL_2 + Sigma-types + PI |- AC_rel = GAC_rel and IPL_2 |- AC_rel + IGP -> GAC_rel and IPL_2 |- GAC_rel = OAC_rel
-
-3.2. IPL^2 |- AC_fun + IGP = GAC_fun = OAC_fun = AC_fun + Drinker
-
-3.3. D_iota -> ID_iota and D_epsilon <-> ID_epsilon + Drinker
-
-4. Derivability of choice for decidable relations with well-ordered codomain
-
-5. Equivalence of choices on dependent or non dependent functional types
-
-6. Non contradiction of constructive descriptions wrt functional choices
-
-7. Definite description transports classical logic to the computational world
-
-8. Choice -> Dependent choice -> Countable choice
-
-9.1. AC_fun_ext = AC_fun + Ext_fun_repr + EM
-
-9.2. AC_fun_ext = AC_fun + Ext_prop_fun_repr + PI
-
-References:
-
+ intuitionistic logic. *)
+(** * References: *)
+(**
[[Bell]] John L. Bell, Choice principles in intuitionistic set theory,
unpublished.
@@ -133,47 +48,75 @@ Variable P:A->Prop.
(** ** Constructive choice and description *)
-(** AC_rel *)
+(** AC_rel = relational form of the (non extensional) axiom of choice
+ (a "set-theoretic" axiom of choice) *)
Definition RelationalChoice_on :=
forall R:A->B->Prop,
(forall x : A, exists y : B, R x y) ->
(exists R' : A->B->Prop, subrelation R' R /\ forall x, exists! y, R' x y).
-(** AC_fun *)
+(** AC_fun = functional form of the (non extensional) axiom of choice
+ (a "type-theoretic" axiom of choice) *)
(* Note: This is called Type-Theoretic Description Axiom (TTDA) in
[[Werner97]] (using a non-standard meaning of "description"). This
is called intensional axiom of choice (AC_int) in [[Carlström04]] *)
+Definition FunctionalChoice_on_rel (R:A->B->Prop) :=
+ (forall x:A, exists y : B, R x y) ->
+ exists f : A -> B, (forall x:A, R x (f x)).
+
Definition FunctionalChoice_on :=
forall R:A->B->Prop,
(forall x : A, exists y : B, R x y) ->
(exists f : A->B, forall x : A, R x (f x)).
-(** DC_fun *)
+(** AC_fun_dep = functional form of the (non extensional) axiom of
+ choice, with dependent functions *)
+Definition DependentFunctionalChoice_on (A:Type) (B:A -> Type) :=
+ forall R:forall x:A, B x -> Prop,
+ (forall x:A, exists y : B x, R x y) ->
+ (exists f : (forall x:A, B x), forall x:A, R x (f x)).
+
+(** AC_trunc = axiom of choice for propositional truncations
+ (truncation and quantification commute) *)
+Definition InhabitedForallCommute_on (A : Type) (B : A -> Type) :=
+ (forall x, inhabited (B x)) -> inhabited (forall x, B x).
+
+(** DC_fun = functional form of the dependent axiom of choice *)
Definition FunctionalDependentChoice_on :=
forall (R:A->A->Prop),
(forall x, exists y, R x y) -> forall x0,
(exists f : nat -> A, f 0 = x0 /\ forall n, R (f n) (f (S n))).
-(** ACw_fun *)
+(** ACw_fun = functional form of the countable axiom of choice *)
Definition FunctionalCountableChoice_on :=
forall (R:nat->A->Prop),
(forall n, exists y, R n y) ->
(exists f : nat -> A, forall n, R n (f n)).
-(** AC! or Functional Relation Reification (known as Axiom of Unique Choice
- in topos theory; also called principle of definite description *)
+(** AC! = functional relation reification
+ (known as axiom of unique choice in topos theory,
+ sometimes called principle of definite description in
+ the context of constructive type theory, sometimes
+ called axiom of no choice) *)
Definition FunctionalRelReification_on :=
forall R:A->B->Prop,
(forall x : A, exists! y : B, R x y) ->
(exists f : A->B, forall x : A, R x (f x)).
-(** AC_fun_repr *)
+(** AC_dep! = functional relation reification, with dependent functions
+ see AC! *)
+Definition DependentFunctionalRelReification_on (A:Type) (B:A -> Type) :=
+ forall (R:forall x:A, B x -> Prop),
+ (forall x:A, exists! y : B x, R x y) ->
+ (exists f : (forall x:A, B x), forall x:A, R x (f x)).
+
+(** AC_fun_repr = functional choice of a representative in an equivalence class *)
(* Note: This is called Type-Theoretic Choice Axiom (TTCA) in
[[Werner97]] (by reference to the extensional set-theoretic
@@ -187,7 +130,8 @@ Definition RepresentativeFunctionalChoice_on :=
(Equivalence R) ->
(exists f : A->A, forall x : A, (R x (f x)) /\ forall x', R x x' -> f x = f x').
-(** AC_fun_setoid *)
+(** AC_fun_setoid = functional form of the (so-called extensional) axiom of
+ choice from setoids *)
Definition SetoidFunctionalChoice_on :=
forall R : A -> A -> Prop,
@@ -197,7 +141,8 @@ Definition SetoidFunctionalChoice_on :=
(forall x, exists y, T x y) ->
exists f : A -> B, forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x').
-(** AC_fun_setoid_gen *)
+(** AC_fun_setoid_gen = functional form of the general form of the (so-called
+ extensional) axiom of choice over setoids *)
(* Note: This is called extensional axiom of choice (AC_ext) in
[[Carlström04]]. *)
@@ -213,7 +158,8 @@ Definition GeneralizedSetoidFunctionalChoice_on :=
exists f : A -> B,
forall x : A, T x (f x) /\ (forall x' : A, R x x' -> S (f x) (f x')).
-(** AC_fun_setoid_simple *)
+(** AC_fun_setoid_simple = functional form of the (so-called extensional) axiom of
+ choice from setoids on locally compatible relations *)
Definition SimpleSetoidFunctionalChoice_on A B :=
forall R : A -> A -> Prop,
@@ -222,19 +168,19 @@ Definition SimpleSetoidFunctionalChoice_on A B :=
(forall x, exists y, forall x', R x x' -> T x' y) ->
exists f : A -> B, forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x').
-(** ID_epsilon (constructive version of indefinite description;
- combined with proof-irrelevance, it may be connected to
- Carlström's type theory with a constructive indefinite description
- operator) *)
+(** ID_epsilon = constructive version of indefinite description;
+ combined with proof-irrelevance, it may be connected to
+ Carlström's type theory with a constructive indefinite description
+ operator *)
Definition ConstructiveIndefiniteDescription_on :=
forall P:A->Prop,
(exists x, P x) -> { x:A | P x }.
-(** ID_iota (constructive version of definite description; combined
- with proof-irrelevance, it may be connected to Carlström's and
- Stenlund's type theory with a constructive definite description
- operator) *)
+(** ID_iota = constructive version of definite description;
+ combined with proof-irrelevance, it may be connected to
+ Carlström's and Stenlund's type theory with a
+ constructive definite description operator) *)
Definition ConstructiveDefiniteDescription_on :=
forall P:A->Prop,
@@ -242,7 +188,7 @@ Definition ConstructiveDefiniteDescription_on :=
(** ** Weakly classical choice and description *)
-(** GAC_rel *)
+(** GAC_rel = guarded relational form of the (non extensional) axiom of choice *)
Definition GuardedRelationalChoice_on :=
forall P : A->Prop, forall R : A->B->Prop,
@@ -250,7 +196,7 @@ Definition GuardedRelationalChoice_on :=
(exists R' : A->B->Prop,
subrelation R' R /\ forall x, P x -> exists! y, R' x y).
-(** GAC_fun *)
+(** GAC_fun = guarded functional form of the (non extensional) axiom of choice *)
Definition GuardedFunctionalChoice_on :=
forall P : A->Prop, forall R : A->B->Prop,
@@ -258,7 +204,7 @@ Definition GuardedFunctionalChoice_on :=
(forall x : A, P x -> exists y : B, R x y) ->
(exists f : A->B, forall x, P x -> R x (f x)).
-(** GFR_fun *)
+(** GAC! = guarded functional relation reification *)
Definition GuardedFunctionalRelReification_on :=
forall P : A->Prop, forall R : A->B->Prop,
@@ -266,27 +212,28 @@ Definition GuardedFunctionalRelReification_on :=
(forall x : A, P x -> exists! y : B, R x y) ->
(exists f : A->B, forall x : A, P x -> R x (f x)).
-(** OAC_rel *)
+(** OAC_rel = "omniscient" relational form of the (non extensional) axiom of choice *)
Definition OmniscientRelationalChoice_on :=
forall R : A->B->Prop,
exists R' : A->B->Prop,
subrelation R' R /\ forall x : A, (exists y : B, R x y) -> exists! y, R' x y.
-(** OAC_fun *)
+(** OAC_fun = "omniscient" functional form of the (non extensional) axiom of choice
+ (called AC* in Bell [[Bell]]) *)
Definition OmniscientFunctionalChoice_on :=
forall R : A->B->Prop,
inhabited B ->
exists f : A->B, forall x : A, (exists y : B, R x y) -> R x (f x).
-(** D_epsilon *)
+(** D_epsilon = (weakly classical) indefinite description principle *)
Definition EpsilonStatement_on :=
forall P:A->Prop,
inhabited A -> { x:A | (exists x, P x) -> P x }.
-(** D_iota *)
+(** D_iota = (weakly classical) definite description principle *)
Definition IotaStatement_on :=
forall P:A->Prop,
@@ -300,14 +247,20 @@ Notation RelationalChoice :=
(forall A B : Type, RelationalChoice_on A B).
Notation FunctionalChoice :=
(forall A B : Type, FunctionalChoice_on A B).
-Definition FunctionalDependentChoice :=
+Notation DependentFunctionalChoice :=
+ (forall A (B:A->Type), DependentFunctionalChoice_on B).
+Notation InhabitedForallCommute :=
+ (forall A (B : A -> Type), InhabitedForallCommute_on B).
+Notation FunctionalDependentChoice :=
(forall A : Type, FunctionalDependentChoice_on A).
-Definition FunctionalCountableChoice :=
+Notation FunctionalCountableChoice :=
(forall A : Type, FunctionalCountableChoice_on A).
Notation FunctionalChoiceOnInhabitedSet :=
(forall A B : Type, inhabited B -> FunctionalChoice_on A B).
Notation FunctionalRelReification :=
(forall A B : Type, FunctionalRelReification_on A B).
+Notation DependentFunctionalRelReification :=
+ (forall A (B:A->Type), DependentFunctionalRelReification_on B).
Notation RepresentativeFunctionalChoice :=
(forall A : Type, RepresentativeFunctionalChoice_on A).
Notation SetoidFunctionalChoice :=
@@ -341,38 +294,87 @@ Notation EpsilonStatement :=
(** Subclassical schemes *)
+(** PI = proof irrelevance *)
Definition ProofIrrelevance :=
forall (A:Prop) (a1 a2:A), a1 = a2.
+(** IGP = independence of general premises
+ (an unconstrained generalisation of the constructive principle of
+ independence of premises) *)
Definition IndependenceOfGeneralPremises :=
forall (A:Type) (P:A -> Prop) (Q:Prop),
inhabited A ->
(Q -> exists x, P x) -> exists x, Q -> P x.
+(** Drinker = drinker's paradox (small form)
+ (called Ex in Bell [[Bell]]) *)
Definition SmallDrinker'sParadox :=
forall (A:Type) (P:A -> Prop), inhabited A ->
exists x, (exists x, P x) -> P x.
+(** EM = excluded-middle *)
Definition ExcludedMiddle :=
forall P:Prop, P \/ ~ P.
(** Extensional schemes *)
+(** Ext_prop_repr = choice of a representative among extensional propositions *)
Local Notation ExtensionalPropositionRepresentative :=
(forall (A:Type),
exists h : Prop -> Prop,
forall P : Prop, (P <-> h P) /\ forall Q, (P <-> Q) -> h P = h Q).
+(** Ext_pred_repr = choice of a representative among extensional predicates *)
Local Notation ExtensionalPredicateRepresentative :=
(forall (A:Type),
exists h : (A->Prop) -> (A->Prop),
forall (P : A -> Prop), (forall x, P x <-> h P x) /\ forall Q, (forall x, P x <-> Q x) -> h P = h Q).
+(** Ext_fun_repr = choice of a representative among extensional functions *)
Local Notation ExtensionalFunctionRepresentative :=
(forall (A B:Type),
exists h : (A->B) -> (A->B),
forall (f : A -> B), (forall x, f x = h f x) /\ forall g, (forall x, f x = g x) -> h f = h g).
+(** We let also
+
+- IPL_2 = 2nd-order impredicative minimal predicate logic (with ex. quant.)
+- IPL^2 = 2nd-order functional minimal predicate logic (with ex. quant.)
+- IPL_2^2 = 2nd-order impredicative, 2nd-order functional minimal pred. logic (with ex. quant.)
+
+with no prerequisite on the non-emptiness of domains
+*)
+
+(**********************************************************************)
+(** * Table of contents *)
+
+(* This is very fragile. *)
+(**
+1. Definitions
+
+2. IPL_2^2 |- AC_rel + AC! = AC_fun
+
+3.1. typed IPL_2 + Sigma-types + PI |- AC_rel = GAC_rel and IPL_2 |- AC_rel + IGP -> GAC_rel and IPL_2 |- GAC_rel = OAC_rel
+
+3.2. IPL^2 |- AC_fun + IGP = GAC_fun = OAC_fun = AC_fun + Drinker
+
+3.3. D_iota -> ID_iota and D_epsilon <-> ID_epsilon + Drinker
+
+4. Derivability of choice for decidable relations with well-ordered codomain
+
+5. AC_fun = AC_fun_dep = AC_trunc
+
+6. Non contradiction of constructive descriptions wrt functional choices
+
+7. Definite description transports classical logic to the computational world
+
+8. Choice -> Dependent choice -> Countable choice
+
+9.1. AC_fun_setoid = AC_fun + Ext_fun_repr + EM
+
+9.2. AC_fun_setoid = AC_fun + Ext_pred_repr + PI
+ *)
+
(**********************************************************************)
(** * AC_rel + AC! = AC_fun
@@ -400,9 +402,6 @@ Proof.
apply HR'R; assumption.
Qed.
-Notation description_rel_choice_imp_funct_choice :=
- functional_rel_reification_and_rel_choice_imp_fun_choice (compat "8.6").
-
Lemma fun_choice_imp_rel_choice :
forall A B : Type, FunctionalChoice_on A B -> RelationalChoice_on A B.
Proof.
@@ -416,8 +415,6 @@ Proof.
trivial.
Qed.
-Notation funct_choice_imp_rel_choice := fun_choice_imp_rel_choice (compat "8.6").
-
Lemma fun_choice_imp_functional_rel_reification :
forall A B : Type, FunctionalChoice_on A B -> FunctionalRelReification_on A B.
Proof.
@@ -431,8 +428,6 @@ Proof.
exists f; exact H0.
Qed.
-Notation funct_choice_imp_description := fun_choice_imp_functional_rel_reification (compat "8.6").
-
Corollary fun_choice_iff_rel_choice_and_functional_rel_reification :
forall A B : Type, FunctionalChoice_on A B <->
RelationalChoice_on A B /\ FunctionalRelReification_on A B.
@@ -444,8 +439,6 @@ Proof.
intros [H H0]; exact (functional_rel_reification_and_rel_choice_imp_fun_choice H0 H).
Qed.
-Notation FunChoice_Equiv_RelChoice_and_ParamDefinDescr :=
- fun_choice_iff_rel_choice_and_functional_rel_reification (compat "8.6").
(**********************************************************************)
(** * Connection between the guarded, non guarded and omniscient choices *)
@@ -687,10 +680,6 @@ Qed.
Require Import Wf_nat.
Require Import Decidable.
-Definition FunctionalChoice_on_rel (A B:Type) (R:A->B->Prop) :=
- (forall x:A, exists y : B, R x y) ->
- exists f : A -> B, (forall x:A, R x (f x)).
-
Lemma classical_denumerable_description_imp_fun_choice :
forall A:Type,
FunctionalRelReification_on A nat ->
@@ -712,18 +701,10 @@ Proof.
Qed.
(**********************************************************************)
-(** * Choice on dependent and non dependent function types are equivalent *)
+(** * AC_fun = AC_fun_dep = AC_trunc *)
(** ** Choice on dependent and non dependent function types are equivalent *)
-Definition DependentFunctionalChoice_on (A:Type) (B:A -> Type) :=
- forall R:forall x:A, B x -> Prop,
- (forall x:A, exists y : B x, R x y) ->
- (exists f : (forall x:A, B x), forall x:A, R x (f x)).
-
-Notation DependentFunctionalChoice :=
- (forall A (B:A->Type), DependentFunctionalChoice_on B).
-
(** The easy part *)
Theorem dep_non_dep_functional_choice :
@@ -760,13 +741,7 @@ Proof.
destruct Heq using eq_indd; trivial.
Qed.
-(** Functional choice can be reformulated as a property on [inhabited] *)
-
-Definition InhabitedForallCommute_on (A : Type) (B : A -> Type) :=
- (forall x, inhabited (B x)) -> inhabited (forall x, B x).
-
-Notation InhabitedForallCommute :=
- (forall A (B : A -> Type), InhabitedForallCommute_on B).
+(** ** Functional choice and truncation choice are equivalent *)
Theorem functional_choice_to_inhabited_forall_commute :
FunctionalChoice -> InhabitedForallCommute.
@@ -795,14 +770,6 @@ Qed.
(** ** Reification of dependent and non dependent functional relation are equivalent *)
-Definition DependentFunctionalRelReification_on (A:Type) (B:A -> Type) :=
- forall (R:forall x:A, B x -> Prop),
- (forall x:A, exists! y : B x, R x y) ->
- (exists f : (forall x:A, B x), forall x:A, R x (f x)).
-
-Notation DependentFunctionalRelReification :=
- (forall A (B:A->Type), DependentFunctionalRelReification_on B).
-
(** The easy part *)
Theorem dep_non_dep_functional_rel_reification :
@@ -1337,3 +1304,15 @@ Proof.
apply repr_fun_choice_imp_excluded_middle.
now apply setoid_fun_choice_imp_repr_fun_choice.
Qed.
+
+(**********************************************************************)
+(** * Compatibility notations *)
+Notation description_rel_choice_imp_funct_choice :=
+ functional_rel_reification_and_rel_choice_imp_fun_choice (compat "8.6").
+
+Notation funct_choice_imp_rel_choice := fun_choice_imp_rel_choice (compat "8.6").
+
+Notation FunChoice_Equiv_RelChoice_and_ParamDefinDescr :=
+ fun_choice_iff_rel_choice_and_functional_rel_reification (compat "8.6").
+
+Notation funct_choice_imp_description := fun_choice_imp_functional_rel_reification (compat "8.6").
--
cgit v1.2.3
From ca4aee0fcf1b54363a6a1eb837cd05cd7ffcc0d9 Mon Sep 17 00:00:00 2001
From: Tej Chajed
Date: Wed, 3 May 2017 07:47:51 -0400
Subject: Report a useful error for dependent induction
The dependent induction tactic notation is in the standard library but
not loaded by default, leading to a parser error message that is
confusing and unhelpful. This commit adds a notation for dependent
induction to Init that fails and reports [Require Import
Coq.Program.Equality.] is required to use [dependent induction].
---
theories/Init/Tactics.v | 7 +++++++
1 file changed, 7 insertions(+)
diff --git a/theories/Init/Tactics.v b/theories/Init/Tactics.v
index 5d1e87ae0c..7a846cd1b3 100644
--- a/theories/Init/Tactics.v
+++ b/theories/Init/Tactics.v
@@ -236,3 +236,10 @@ Tactic Notation "clear" "dependent" hyp(h) :=
Tactic Notation "revert" "dependent" hyp(h) :=
generalize dependent h.
+
+(** Provide an error message for dependent induction that reports an import is
+required to use it. Importing Coq.Program.Equality will shadow this notation
+with the actual [dependent induction] tactic. *)
+
+Tactic Notation "dependent" "induction" ident(H) :=
+ fail "To use dependent induction, first [Require Import Coq.Program.Equality.]".
--
cgit v1.2.3
From 844bffb7d6c84a02dcef300dda9099487b23c09a Mon Sep 17 00:00:00 2001
From: Hugo Herbelin
Date: Tue, 11 Apr 2017 21:17:03 +0200
Subject: Added an option Set Debug Cbv.
---
pretyping/cbv.ml | 27 ++++++++++++++++++++++++---
1 file changed, 24 insertions(+), 3 deletions(-)
diff --git a/pretyping/cbv.ml b/pretyping/cbv.ml
index e18625c427..bd7350dc4e 100644
--- a/pretyping/cbv.ml
+++ b/pretyping/cbv.ml
@@ -175,6 +175,19 @@ let cofixp_reducible flgs _ stk =
else
false
+let debug_cbv = ref false
+let _ = Goptions.declare_bool_option {
+ Goptions.optsync = true; Goptions.optdepr = false;
+ Goptions.optname = "cbv visited constants display";
+ Goptions.optkey = ["Debug";"Cbv"];
+ Goptions.optread = (fun () -> !debug_cbv);
+ Goptions.optwrite = (fun a -> debug_cbv:=a);
+}
+
+let pr_key = function
+ | ConstKey (sp,_) -> Names.Constant.print sp
+ | VarKey id -> Names.Id.print id
+ | RelKey n -> Pp.(str "REL_" ++ int n)
(* The main recursive functions
*
@@ -254,9 +267,17 @@ let rec norm_head info env t stack =
and norm_head_ref k info env stack normt =
if red_set_ref (info_flags info) normt then
match ref_value_cache info normt with
- | Some body -> strip_appl (shift_value k body) stack
- | None -> (VAL(0,make_constr_ref k normt),stack)
- else (VAL(0,make_constr_ref k normt),stack)
+ | Some body ->
+ if !debug_cbv then Feedback.msg_debug Pp.(str "Unfolding " ++ pr_key normt);
+ strip_appl (shift_value k body) stack
+ | None ->
+ if !debug_cbv then Feedback.msg_debug Pp.(str "Not unfolding " ++ pr_key normt);
+ (VAL(0,make_constr_ref k normt),stack)
+ else
+ begin
+ if !debug_cbv then Feedback.msg_debug Pp.(str "Not unfolding " ++ pr_key normt);
+ (VAL(0,make_constr_ref k normt),stack)
+ end
(* cbv_stack_term performs weak reduction on constr t under the subs
* env, with context stack, i.e. ([env]t stack). First computes weak
--
cgit v1.2.3
From 952e9aea47d3fad2d0f488d968ff0e90fa403ebc Mon Sep 17 00:00:00 2001
From: Pierre Courtieu
Date: Wed, 3 May 2017 11:37:08 +0200
Subject: Adding an even more compact mode for goal display.
We want to print variables in vertical boxes as much as possible.
The criterion to distinguish "variable" from "hypothesis" is not
obvious. We chose this one but may change it in the future:
X:T is a variable if T is not of type Prop AND T is "simple" which
means T is either id or (t T1 .. Tn) Ti's are sort-typed (typicall
Ti:Type, but not Ti:nat).
---
printing/printer.ml | 60 +++++++++++++++++++++++++++++++++++++++++++------
printing/printer.mli | 5 ++++-
vernac/vernacentries.ml | 18 +++++++++++++++
3 files changed, 75 insertions(+), 8 deletions(-)
diff --git a/printing/printer.ml b/printing/printer.ml
index 93d221f03f..0e31a4a042 100644
--- a/printing/printer.ml
+++ b/printing/printer.ml
@@ -262,6 +262,13 @@ let pr_pattern t = pr_pattern_env (Global.env()) empty_names_context t*)
(**********************************************************************)
(* Contexts and declarations *)
+
+(* Flag for compact display of goals *)
+
+let get_compact_context,set_compact_context =
+ let compact_context = ref false in
+ (fun () -> !compact_context),(fun b -> compact_context := b)
+
let pr_compacted_decl env sigma decl =
let ids, pbody, typ = match decl with
| CompactedDecl.LocalAssum (ids, typ) ->
@@ -342,15 +349,55 @@ let pr_ne_context_of header env sigma =
List.is_empty (Environ.named_context env) then (mt ())
else let penv = pr_context_unlimited env sigma in (header ++ penv ++ fnl ())
+(* Heuristic for horizontalizing hypothesis:
+ Detecting variable which type is a simple id or of the form (t x y ...)
+ where t is a product or only sorts (typically [Type -> Type -> ...]
+ and not [nat -> nat -> ...] ).
+ + Special case for non-Prop dependent terms. *)
+let rec should_compact env sigma typ =
+ get_compact_context() &&
+ match kind_of_term typ with
+ | Rel _ | Var _ | Sort _ | Const _ | Ind _ -> true
+ | App (c,args) ->
+ let _,type_of_c = Typing.type_of env sigma c in
+ let _,type_of_typ = Typing.type_of env sigma typ in
+ not (is_Prop type_of_typ)
+ && (* These two more tests detect rare cases of non-Prop-sorted
+ dependent hypothesis: *)
+ let lnamedtyp , _ = decompose_prod type_of_c in
+ (* c has a non dependent type *)
+ List.for_all (fun (_,typarg) -> isSort typarg) lnamedtyp
+ && (* and real arguments are recursively elligible to compaction. *)
+ Array.for_all (should_compact env sigma) args
+ | _ -> false
+
+
+(* If option Compact Contexts is set, we pack "simple" hypothesis in a
+ hov box (with three sapaces as a separator), the global box being a
+ v box *)
let rec bld_sign_env env sigma ctxt pps =
match ctxt with
| [] -> pps
+ | CompactedDecl.LocalAssum (ids,typ)::ctxt' when should_compact env sigma typ ->
+ let pps',ctxt' = bld_sign_env_id env sigma ctxt (mt ()) true in
+ (* putting simple hyps in a more horizontal flavor *)
+ bld_sign_env env sigma ctxt' (pps ++ brk (0,0) ++ hov 0 pps')
| d:: ctxt' ->
- let pidt = pr_var_list_decl env sigma d in
- let pps' = pps ++ brk (0,0) ++ pidt in
- bld_sign_env env sigma ctxt' pps'
+ let pidt = pr_var_list_decl env sigma d in
+ let pps' = pps ++ brk (0,0) ++ pidt in
+ bld_sign_env env sigma ctxt' pps'
+and bld_sign_env_id env sigma ctxt pps is_start =
+ match ctxt with
+ | [] -> pps,ctxt
+ | CompactedDecl.LocalAssum(ids,typ) as d :: ctxt' when should_compact env sigma typ ->
+ let pidt = pr_var_list_decl env sigma d in
+ let pps' = pps ++ (if not is_start then brk (3,0) else (mt ())) ++ pidt in
+ bld_sign_env_id env sigma ctxt' pps' false
+ | _ -> pps,ctxt
+(* compact printing an env (variables and de Bruijn). Separator: three
+ spaces between simple hyps, and newline otherwise *)
let pr_context_limit_compact ?n env sigma =
let ctxt = Termops.compact_named_context (named_context env) in
let lgth = List.length ctxt in
@@ -360,15 +407,14 @@ let pr_context_limit_compact ?n env sigma =
| Some n when n > lgth -> lgth
| Some n -> n in
let ctxt_chopped,ctxt_hidden = Util.List.chop n_capped ctxt in
- (* a dot line hinting the number of hiden hyps. *)
+ (* a dot line hinting the number of hidden hyps. *)
let hidden_dots = String.make (List.length ctxt_hidden) '.' in
let sign_env = v 0 (str hidden_dots ++ (mt ())
++ bld_sign_env env sigma (List.rev ctxt_chopped) (mt ())) in
let db_env =
- fold_rel_context
- (fun env d pps -> pps ++ fnl () ++ pr_rel_decl env sigma d)
+ fold_rel_context (fun env d pps -> pps ++ fnl () ++ pr_rel_decl env sigma d)
env ~init:(mt ()) in
- (sign_env ++ db_env)
+ sign_env ++ db_env
(* The number of printed hypothesis in a goal *)
(* If [None], no limit *)
diff --git a/printing/printer.mli b/printing/printer.mli
index 504392e35f..9c9f6e766a 100644
--- a/printing/printer.mli
+++ b/printing/printer.mli
@@ -111,6 +111,9 @@ val pr_pconstructor : env -> pconstructor -> std_ppcmds
(** Contexts *)
+(** Display compact contexts of goals (simple hyps on the same line) *)
+val set_compact_context : bool -> unit
+val get_compact_context : unit -> bool
val pr_context_unlimited : env -> evar_map -> std_ppcmds
val pr_ne_context_of : std_ppcmds -> env -> evar_map -> std_ppcmds
@@ -132,7 +135,7 @@ val pr_cpred : Cpred.t -> std_ppcmds
val pr_idpred : Id.Pred.t -> std_ppcmds
val pr_transparent_state : transparent_state -> std_ppcmds
-(** Proofs *)
+(** Proofs, these functions obey [Hyps Limit] and [Compact contexts]. *)
val pr_goal : goal sigma -> std_ppcmds
val pr_subgoals : ?pr_first:bool -> std_ppcmds option -> evar_map -> evar list -> Goal.goal list -> int list -> goal list -> std_ppcmds
diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml
index 2cb6f3918f..40cd1a29c9 100644
--- a/vernac/vernacentries.ml
+++ b/vernac/vernacentries.ml
@@ -1411,6 +1411,24 @@ let _ =
optread = (fun _ -> None);
optwrite = (fun _ -> ()) }
+let _ =
+ declare_int_option
+ { optsync = false;
+ optdepr = false;
+ optname = "the hypotheses limit";
+ optkey = ["Hyps";"Limit"];
+ optread = Flags.print_hyps_limit;
+ optwrite = Flags.set_print_hyps_limit }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "display compact goal contexts";
+ optkey = ["Printing";"Compact";"Contexts"];
+ optread = (fun () -> Printer.get_compact_context());
+ optwrite = (fun b -> Printer.set_compact_context b) }
+
let _ =
declare_int_option
{ optsync = true;
--
cgit v1.2.3
From b12af1535c0ba8cab23e7f9ff18f75688c0e523c Mon Sep 17 00:00:00 2001
From: Gaetan Gilbert
Date: Wed, 3 May 2017 19:16:48 +0200
Subject: Make congruence reuse discriminate instead of rolling its own.
This changes the produced terms a bit, eg
Axiom T : Type.
Lemma foo : true = false -> T.
Proof. congruence. Qed.
used to produce
fun H : true = false =>
let Heq := H : true = false in
@eq_rect Type True (fun X : Type => X) I T
(@f_equal bool Type (fun t : bool => if t then True else T) true false Heq)
now produces
fun H : true = false =>
let Heq : true = false := H in
let H0 : False := @eq_ind bool true (fun e : bool => if e then True else False) I false Heq in
False_rect T H0
i.e. instead of proving [True = T] by [f_equal] then transporting [I]
across this identity, it now proves [False] by [eq_ind] then uses exfalso.
---
plugins/cc/cctac.ml | 27 ++++++---------------------
1 file changed, 6 insertions(+), 21 deletions(-)
diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml
index 7c5efaea3a..1cb417bf47 100644
--- a/plugins/cc/cctac.ml
+++ b/plugins/cc/cctac.ml
@@ -397,33 +397,18 @@ let convert_to_hyp_tac c1 t1 c2 t2 p =
simplest_elim false_t]
end }
-let discriminate_tac (cstr,u as cstru) p =
+(* Essentially [assert (Heq : lhs = rhs) by proof_tac p; discriminate Heq] *)
+let discriminate_tac cstru p =
Proofview.Goal.enter { enter = begin fun gl ->
- let t1=constr_of_term p.p_lhs and t2=constr_of_term p.p_rhs in
+ let lhs=constr_of_term p.p_lhs and rhs=constr_of_term p.p_rhs in
let env = Proofview.Goal.env gl in
- let concl = Proofview.Goal.concl gl in
- let xid = Tacmach.New.pf_get_new_id (Id.of_string "X") gl in
- let identity = Universes.constr_of_global (Lazy.force _I) in
- let identity = EConstr.of_constr identity in
- let trivial = Universes.constr_of_global (Lazy.force _True) in
- let trivial = EConstr.of_constr trivial in
let evm = Tacmach.New.project gl in
- let evm, intype = refresh_type env evm (Tacmach.New.pf_unsafe_type_of gl t1) in
- let evm, outtype = Evd.new_sort_variable Evd.univ_flexible evm in
- let outtype = mkSort outtype in
- let pred = mkLambda(Name xid,outtype,mkRel 1) in
+ let evm, intype = refresh_type env evm (Tacmach.New.pf_unsafe_type_of gl lhs) in
let hid = Tacmach.New.pf_get_new_id (Id.of_string "Heq") gl in
- let proj = build_projection intype cstru trivial concl gl in
- let injt = app_global _f_equal
- [|intype;outtype;proj;t1;t2;mkVar hid|] in
- let endt k =
- injt (fun injt ->
- app_global _eq_rect
- [|outtype;trivial;pred;identity;concl;injt|] k) in
- let neweq = app_global _eq [|intype;t1;t2|] in
+ let neweq=app_global _eq [|intype;lhs;rhs|] in
Tacticals.New.tclTHEN (Proofview.Unsafe.tclEVARS evm)
(Tacticals.New.tclTHENS (neweq (assert_before (Name hid)))
- [proof_tac p; endt refine_exact_check])
+ [proof_tac p; Equality.discrHyp hid])
end }
(* wrap everything *)
--
cgit v1.2.3
From 9b54bda96f51cc5387f195614620fae53dec5bd1 Mon Sep 17 00:00:00 2001
From: Cyprien Mangin
Date: Wed, 3 May 2017 10:47:53 +0200
Subject: FIx bug #5300: uncaught exception in "Print Assumptions".
---
test-suite/bugs/closed/5300.v | 39 +++++++++++++++++++++++++++++++++++++++
toplevel/assumptions.ml | 21 +++++++++++++++++++--
2 files changed, 58 insertions(+), 2 deletions(-)
create mode 100644 test-suite/bugs/closed/5300.v
diff --git a/test-suite/bugs/closed/5300.v b/test-suite/bugs/closed/5300.v
new file mode 100644
index 0000000000..18202df508
--- /dev/null
+++ b/test-suite/bugs/closed/5300.v
@@ -0,0 +1,39 @@
+Module Test1.
+
+ Module Type Foo.
+ Parameter t : unit.
+ End Foo.
+
+ Module Bar : Foo.
+ Module Type Rnd. Definition t' : unit := tt. End Rnd.
+ Module Rnd_inst : Rnd. Definition t' : unit := tt. End Rnd_inst.
+ Definition t : unit.
+ exact Rnd_inst.t'.
+ Qed.
+ End Bar.
+
+ Print Assumptions Bar.t.
+End Test1.
+
+Module Test2.
+ Module Type Foo.
+ Parameter t1 : unit.
+ Parameter t2 : unit.
+ End Foo.
+
+ Module Bar : Foo.
+ Inductive ind := .
+ Definition t' : unit := tt.
+ Definition t1 : unit.
+ Proof.
+ exact ((fun (_ : ind -> False) => tt) (fun H => match H with end)).
+ Qed.
+ Definition t2 : unit.
+ Proof.
+ exact t'.
+ Qed.
+ End Bar.
+
+ Print Assumptions Bar.t1.
+ Print Assumptions Bar.t1.
+End Test2.
diff --git a/toplevel/assumptions.ml b/toplevel/assumptions.ml
index 45c539e229..a865ee987f 100644
--- a/toplevel/assumptions.ml
+++ b/toplevel/assumptions.ml
@@ -42,6 +42,11 @@ let rec search_cst_label lab = function
| (l, SFBconst cb) :: _ when Label.equal l lab -> cb
| _ :: fields -> search_cst_label lab fields
+let rec search_mind_label lab = function
+ | [] -> raise Not_found
+ | (l, SFBmind mind) :: _ when Label.equal l lab -> mind
+ | _ :: fields -> search_mind_label lab fields
+
(* TODO: using [empty_delta_resolver] below is probably slightly incorrect. But:
a) I don't see currently what should be used instead
b) this shouldn't be critical for Print Assumption. At worse some
@@ -133,6 +138,18 @@ let lookup_constant cst =
else lookup_constant_in_impl cst (Some cb)
with Not_found -> lookup_constant_in_impl cst None
+let lookup_mind_in_impl mind =
+ try
+ let mp,dp,lab = repr_kn (canonical_mind mind) in
+ let fields = memoize_fields_of_mp mp in
+ search_mind_label lab fields
+ with Not_found ->
+ anomaly (str "Print Assumption: unknown inductive " ++ MutInd.print mind)
+
+let lookup_mind mind =
+ try Global.lookup_mind mind
+ with Not_found -> lookup_mind_in_impl mind
+
(** Graph traversal of an object, collecting on the way the dependencies of
traversed objects *)
@@ -204,7 +221,7 @@ and traverse_inductive (curr, data, ax2ty) mind obj =
where I_0, I_1, ... are in the same mutual definition and c_ij
are all their constructors. *)
if Refmap_env.mem firstind_ref data then data, ax2ty else
- let mib = Global.lookup_mind mind in
+ let mib = lookup_mind mind in
(* Collects references of parameters *)
let param_ctx = mib.mind_params_ctxt in
let nparam = List.length param_ctx in
@@ -308,7 +325,7 @@ let assumptions ?(add_opaque=false) ?(add_transparent=false) st gr t =
else
accu
| IndRef (m,_) | ConstructRef ((m,_),_) ->
- let mind = Global.lookup_mind m in
+ let mind = lookup_mind m in
if mind.mind_typing_flags.check_guarded then
accu
else
--
cgit v1.2.3
From 1fe73e6af81759fa8b78c8660745492ed886d477 Mon Sep 17 00:00:00 2001
From: Pierre Courtieu
Date: Wed, 26 Apr 2017 11:57:58 +0200
Subject: Adding an option "Printing Unfocused".
Off by default.
+ small refactoring of emacs hacks in printer.ml.
---
ide/coq.ml | 6 ++-
ide/coq.mli | 2 +
ide/coqide_ui.ml | 1 +
ide/ide_slave.ml | 3 +-
ide/wg_ProofView.ml | 35 +++++++++----
printing/printer.ml | 139 ++++++++++++++++++++++++++++++++++-----------------
printing/printer.mli | 21 +++++++-
toplevel/coqtop.ml | 1 +
8 files changed, 147 insertions(+), 61 deletions(-)
diff --git a/ide/coq.ml b/ide/coq.ml
index 3a1d877872..cd45e2fcdc 100644
--- a/ide/coq.ml
+++ b/ide/coq.ml
@@ -519,6 +519,7 @@ struct
let all_basic = ["Printing"; "All"]
let existential = ["Printing"; "Existential"; "Instances"]
let universes = ["Printing"; "Universes"]
+ let unfocused = ["Printing"; "Unfocused"]
type bool_descr = { opts : t list; init : bool; label : string }
@@ -534,7 +535,8 @@ struct
label = "Display _existential variable instances" };
{ opts = [universes]; init = false; label = "Display _universe levels" };
{ opts = [all_basic;existential;universes]; init = false;
- label = "Display all _low-level contents" }
+ label = "Display all _low-level contents" };
+ { opts = [unfocused]; init = false; label = "Display _unfocused goals" }
]
(** The current status of the boolean options *)
@@ -549,6 +551,8 @@ struct
let _ = reset ()
+ let printing_unfocused () = Hashtbl.find current_state unfocused
+
(** Transmitting options to coqtop *)
let enforce h k =
diff --git a/ide/coq.mli b/ide/coq.mli
index ab8c12a6f1..e8e2f5239e 100644
--- a/ide/coq.mli
+++ b/ide/coq.mli
@@ -140,6 +140,8 @@ sig
val set : t -> bool -> unit
+ val printing_unfocused: unit -> bool
+
(** [enforce] transmits to coq the current option values.
It is also called by [goals] and [evars] above. *)
diff --git a/ide/coqide_ui.ml b/ide/coqide_ui.ml
index 91c281c8d8..717c4000f5 100644
--- a/ide/coqide_ui.ml
+++ b/ide/coqide_ui.ml
@@ -85,6 +85,7 @@ let init () =
\n \
\n \
\n \
+\n \
\n \
\n