diff options
Diffstat (limited to 'theories/Logic')
29 files changed, 97 insertions, 97 deletions
diff --git a/theories/Logic/Berardi.v b/theories/Logic/Berardi.v index 2b388687..38377573 100644 --- a/theories/Logic/Berardi.v +++ b/theories/Logic/Berardi.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -45,7 +45,7 @@ Lemma AC_IF : (B -> Q e1) -> (~ B -> Q e2) -> Q (IFProp B e1 e2). Proof. intros P B e1 e2 Q p1 p2. -unfold IFProp in |- *. +unfold IFProp. case (EM B); assumption. Qed. @@ -76,7 +76,7 @@ Record retract_cond : Prop := Lemma AC : forall r:retract_cond, retract -> forall a:A, j2 r (i2 r a) = a. Proof. intros r. -case r; simpl in |- *. +case r; simpl. trivial. Qed. @@ -113,7 +113,7 @@ Lemma retract_pow_U_U : retract (pow U) U. Proof. exists g f. intro a. -unfold f, g in |- *; simpl in |- *. +unfold f, g; simpl. apply AC. exists (fun x:pow U => x) (fun x:pow U => x). trivial. @@ -130,8 +130,8 @@ Definition R : U := g (fun u:U => Not_b (u U u)). Lemma not_has_fixpoint : R R = Not_b (R R). Proof. -unfold R at 1 in |- *. -unfold g in |- *. +unfold R at 1. +unfold g. rewrite AC with (r := L1 U U) (a := fun u:U => Not_b (u U u)). trivial. exists (fun x:pow U => x) (fun x:pow U => x); trivial. @@ -141,7 +141,7 @@ Qed. Theorem classical_proof_irrelevence : T = F. Proof. generalize not_has_fixpoint. -unfold Not_b in |- *. +unfold Not_b. apply AC_IF. intros is_true is_false. elim is_true; elim is_false; trivial. diff --git a/theories/Logic/ChoiceFacts.v b/theories/Logic/ChoiceFacts.v index fb7898c6..1a32d518 100644 --- a/theories/Logic/ChoiceFacts.v +++ b/theories/Logic/ChoiceFacts.v @@ -1,7 +1,7 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -344,7 +344,7 @@ Lemma rel_choice_and_proof_irrel_imp_guarded_rel_choice : RelationalChoice -> ProofIrrelevance -> GuardedRelationalChoice. Proof. intros rel_choice proof_irrel. - red in |- *; intros A B P R H. + red; intros A B P R H. destruct (rel_choice _ _ (fun (x:sigT P) (y:B) => R (projT1 x) y)) as (R',(HR'R,H0)). intros (x,HPx). destruct (H x HPx) as (y,HRxy). @@ -580,7 +580,7 @@ Lemma classical_denumerable_description_imp_fun_choice : (forall x y, decidable (R x y)) -> FunctionalChoice_on_rel R. Proof. intros A Descr. - red in |- *; intros R Rdec H. + red; intros R Rdec H. set (R':= fun x y => R x y /\ forall y', R x y' -> y <= y'). destruct (Descr R') as (f,Hf). intro x. diff --git a/theories/Logic/Classical.v b/theories/Logic/Classical.v index 9362a11f..d25e0e21 100644 --- a/theories/Logic/Classical.v +++ b/theories/Logic/Classical.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Logic/ClassicalChoice.v b/theories/Logic/ClassicalChoice.v index 6bc0be1d..479056c9 100644 --- a/theories/Logic/ClassicalChoice.v +++ b/theories/Logic/ClassicalChoice.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Logic/ClassicalDescription.v b/theories/Logic/ClassicalDescription.v index d35ed138..2fd6e68e 100644 --- a/theories/Logic/ClassicalDescription.v +++ b/theories/Logic/ClassicalDescription.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -20,7 +20,7 @@ Require Export Classical. (* Axiomatize classical reasoning *) Require Export Description. (* Axiomatize constructive form of Church's iota *) Require Import ChoiceFacts. -Notation Local inhabited A := A (only parsing). +Local Notation inhabited A := A (only parsing). (** The idea for the following proof comes from [ChicliPottierSimpson02] *) diff --git a/theories/Logic/ClassicalEpsilon.v b/theories/Logic/ClassicalEpsilon.v index ae32b127..7ab991f8 100644 --- a/theories/Logic/ClassicalEpsilon.v +++ b/theories/Logic/ClassicalEpsilon.v @@ -1,7 +1,7 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Logic/ClassicalFacts.v b/theories/Logic/ClassicalFacts.v index bcec657a..34ae1cd5 100644 --- a/theories/Logic/ClassicalFacts.v +++ b/theories/Logic/ClassicalFacts.v @@ -1,7 +1,7 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -117,7 +117,7 @@ Qed. *) -Notation Local inhabited A := A (only parsing). +Local Notation inhabited A := A (only parsing). Lemma prop_ext_A_eq_A_imp_A : prop_extensionality -> forall A:Prop, inhabited A -> (A -> A) = A. @@ -148,7 +148,7 @@ Proof. case (prop_ext_retract_A_A_imp_A Ext A a); intros g1 g2 g1_o_g2. exists (fun f => (fun x:A => f (g1 x x)) (g2 (fun x => f (g1 x x)))). intro f. - pattern (g1 (g2 (fun x:A => f (g1 x x)))) at 1 in |- *. + pattern (g1 (g2 (fun x:A => f (g1 x x)))) at 1. rewrite (g1_o_g2 (fun x:A => f (g1 x x))). reflexivity. Qed. @@ -191,13 +191,13 @@ Section Proof_irrelevance_gen. intros Ext Ind. case (ext_prop_fixpoint Ext bool true); intros G Gfix. set (neg := fun b:bool => bool_elim bool false true b). - generalize (refl_equal (G neg)). - pattern (G neg) at 1 in |- *. + generalize (eq_refl (G neg)). + pattern (G neg) at 1. apply Ind with (b := G neg); intro Heq. rewrite (bool_elim_redl bool false true). - change (true = neg true) in |- *; rewrite Heq; apply Gfix. + change (true = neg true); rewrite Heq; apply Gfix. rewrite (bool_elim_redr bool false true). - change (neg false = false) in |- *; rewrite Heq; symmetry in |- *; + change (neg false = false); rewrite Heq; symmetry ; apply Gfix. Qed. @@ -207,9 +207,9 @@ Section Proof_irrelevance_gen. intros Ext Ind A a1 a2. set (f := fun b:bool => bool_elim A a1 a2 b). rewrite (bool_elim_redl A a1 a2). - change (f true = a2) in |- *. + change (f true = a2). rewrite (bool_elim_redr A a1 a2). - change (f true = f false) in |- *. + change (f true = f false). rewrite (aux Ext Ind). reflexivity. Qed. @@ -228,9 +228,9 @@ Section Proof_irrelevance_Prop_Ext_CC. Definition FalseP : BoolP := fun C c1 c2 => c2. Definition BoolP_elim C c1 c2 (b:BoolP) := b C c1 c2. Definition BoolP_elim_redl (C:Prop) (c1 c2:C) : - c1 = BoolP_elim C c1 c2 TrueP := refl_equal c1. + c1 = BoolP_elim C c1 c2 TrueP := eq_refl c1. Definition BoolP_elim_redr (C:Prop) (c1 c2:C) : - c2 = BoolP_elim C c1 c2 FalseP := refl_equal c2. + c2 = BoolP_elim C c1 c2 FalseP := eq_refl c2. Definition BoolP_dep_induction := forall P:BoolP -> Prop, P TrueP -> P FalseP -> forall b:BoolP, P b. @@ -263,9 +263,9 @@ Section Proof_irrelevance_CIC. | trueP : boolP | falseP : boolP. Definition boolP_elim_redl (C:Prop) (c1 c2:C) : - c1 = boolP_ind C c1 c2 trueP := refl_equal c1. + c1 = boolP_ind C c1 c2 trueP := eq_refl c1. Definition boolP_elim_redr (C:Prop) (c1 c2:C) : - c2 = boolP_ind C c1 c2 falseP := refl_equal c2. + c2 = boolP_ind C c1 c2 falseP := eq_refl c2. Scheme boolP_indd := Induction for boolP Sort Prop. Lemma ext_prop_dep_proof_irrel_cic : prop_extensionality -> proof_irrelevance. @@ -344,8 +344,8 @@ Section Proof_irrelevance_EM_CC. Lemma p2p1 : forall A:Prop, A -> b2p (p2b A). Proof. - unfold p2b in |- *; intro A; apply or_dep_elim with (b := em A); - unfold b2p in |- *; intros. + unfold p2b; intro A; apply or_dep_elim with (b := em A); + unfold b2p; intros. apply (or_elim_redl A (~ A) B (fun _ => b1) (fun _ => b2)). destruct (b H). Qed. @@ -353,8 +353,8 @@ Section Proof_irrelevance_EM_CC. Lemma p2p2 : b1 <> b2 -> forall A:Prop, b2p (p2b A) -> A. Proof. intro not_eq_b1_b2. - unfold p2b in |- *; intro A; apply or_dep_elim with (b := em A); - unfold b2p in |- *; intros. + unfold p2b; intro A; apply or_dep_elim with (b := em A); + unfold b2p; intros. assumption. destruct not_eq_b1_b2. rewrite <- (or_elim_redr A (~ A) B (fun _ => b1) (fun _ => b2)) in H. @@ -392,9 +392,9 @@ Section Proof_irrelevance_CCI. Hypothesis em : forall A:Prop, A \/ ~ A. Definition or_elim_redl (A B C:Prop) (f:A -> C) (g:B -> C) - (a:A) : f a = or_ind f g (or_introl B a) := refl_equal (f a). + (a:A) : f a = or_ind f g (or_introl B a) := eq_refl (f a). Definition or_elim_redr (A B C:Prop) (f:A -> C) (g:B -> C) - (b:B) : g b = or_ind f g (or_intror A b) := refl_equal (g b). + (b:B) : g b = or_ind f g (or_intror A b) := eq_refl (g b). Scheme or_indd := Induction for or Sort Prop. Theorem proof_irrelevance_cci : forall (B:Prop) (b1 b2:B), b1 = b2. diff --git a/theories/Logic/ClassicalUniqueChoice.v b/theories/Logic/ClassicalUniqueChoice.v index ebb73b19..4a4fc23f 100644 --- a/theories/Logic/ClassicalUniqueChoice.v +++ b/theories/Logic/ClassicalUniqueChoice.v @@ -1,7 +1,7 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Logic/Classical_Pred_Set.v b/theories/Logic/Classical_Pred_Set.v index 7d8bde71..cda9d22c 100644 --- a/theories/Logic/Classical_Pred_Set.v +++ b/theories/Logic/Classical_Pred_Set.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Logic/Classical_Pred_Type.v b/theories/Logic/Classical_Pred_Type.v index 9d57fe88..7e1a4096 100644 --- a/theories/Logic/Classical_Pred_Type.v +++ b/theories/Logic/Classical_Pred_Type.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -42,7 +42,7 @@ Qed. Lemma not_ex_all_not : forall P:U -> Prop, ~ (exists n : U, P n) -> forall n:U, ~ P n. Proof. (* Intuitionistic *) -unfold not in |- *; intros P notex n abs. +unfold not; intros P notex n abs. apply notex. exists n; trivial. Qed. @@ -52,20 +52,20 @@ Lemma not_ex_not_all : Proof. intros P H n. apply NNPP. -red in |- *; intro K; apply H; exists n; trivial. +red; intro K; apply H; exists n; trivial. Qed. Lemma ex_not_not_all : forall P:U -> Prop, (exists n : U, ~ P n) -> ~ (forall n:U, P n). Proof. (* Intuitionistic *) -unfold not in |- *; intros P exnot allP. +unfold not; intros P exnot allP. elim exnot; auto. Qed. Lemma all_not_not_ex : forall P:U -> Prop, (forall n:U, ~ P n) -> ~ (exists n : U, P n). Proof. (* Intuitionistic *) -unfold not in |- *; intros P allnot exP; elim exP; intros n p. +unfold not; intros P allnot exP; elim exP; intros n p. apply allnot with n; auto. Qed. diff --git a/theories/Logic/Classical_Prop.v b/theories/Logic/Classical_Prop.v index d2b35da2..1f6b05f5 100644 --- a/theories/Logic/Classical_Prop.v +++ b/theories/Logic/Classical_Prop.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -20,7 +20,7 @@ Axiom classic : forall P:Prop, P \/ ~ P. Lemma NNPP : forall p:Prop, ~ ~ p -> p. Proof. -unfold not in |- *; intros; elim (classic p); auto. +unfold not; intros; elim (classic p); auto. intro NP; elim (H NP). Qed. @@ -35,7 +35,7 @@ Qed. Lemma not_imply_elim : forall P Q:Prop, ~ (P -> Q) -> P. Proof. -intros; apply NNPP; red in |- *. +intros; apply NNPP; red. intro; apply H; intro; absurd P; trivial. Qed. @@ -68,7 +68,7 @@ Qed. Lemma or_not_and : forall P Q:Prop, ~ P \/ ~ Q -> ~ (P /\ Q). Proof. -simple induction 1; red in |- *; simple induction 2; auto. +simple induction 1; red; simple induction 2; auto. Qed. Lemma not_or_and : forall P Q:Prop, ~ (P \/ Q) -> ~ P /\ ~ Q. @@ -112,7 +112,7 @@ Module Eq_rect_eq. Lemma eq_rect_eq : forall (U:Type) (p:U) (Q:U -> Type) (x:Q p) (h:p = p), x = eq_rect p Q x p h. Proof. -intros; rewrite proof_irrelevance with (p1:=h) (p2:=refl_equal p); reflexivity. +intros; rewrite proof_irrelevance with (p1:=h) (p2:=eq_refl p); reflexivity. Qed. End Eq_rect_eq. diff --git a/theories/Logic/Classical_Type.v b/theories/Logic/Classical_Type.v index 9b28a6ab..86fdd69f 100644 --- a/theories/Logic/Classical_Type.v +++ b/theories/Logic/Classical_Type.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Logic/ConstructiveEpsilon.v b/theories/Logic/ConstructiveEpsilon.v index 33550389..89d3eebc 100644 --- a/theories/Logic/ConstructiveEpsilon.v +++ b/theories/Logic/ConstructiveEpsilon.v @@ -1,12 +1,12 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(*i $Id: ConstructiveEpsilon.v 14628 2011-11-03 23:22:45Z herbelin $ i*) +(*i $Id: ConstructiveEpsilon.v 15714 2012-08-08 18:54:37Z herbelin $ i*) (** This provides with a proof of the constructive form of definite and indefinite descriptions for Sigma^0_1-formulas (hereafter called @@ -112,7 +112,7 @@ of our searching algorithm. *) Let R (x y : nat) : Prop := x = S y /\ ~ P y. -Notation Local acc x := (Acc R x). +Local Notation acc x := (Acc R x). Lemma P_implies_acc : forall x : nat, P x -> acc x. Proof. diff --git a/theories/Logic/Decidable.v b/theories/Logic/Decidable.v index fec7904e..aaf1813b 100644 --- a/theories/Logic/Decidable.v +++ b/theories/Logic/Decidable.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Logic/Description.v b/theories/Logic/Description.v index b74ebcc8..3e5d4ef0 100644 --- a/theories/Logic/Description.v +++ b/theories/Logic/Description.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Logic/Diaconescu.v b/theories/Logic/Diaconescu.v index 8569e55e..87b27987 100644 --- a/theories/Logic/Diaconescu.v +++ b/theories/Logic/Diaconescu.v @@ -1,7 +1,7 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -61,7 +61,7 @@ Variable pred_extensionality : PredicateExtensionality. Lemma prop_ext : forall A B:Prop, (A <-> B) -> A = B. Proof. intros A B H. - change ((fun _ => A) true = (fun _ => B) true) in |- *. + change ((fun _ => A) true = (fun _ => B) true). rewrite pred_extensionality with (P := fun _:bool => A) (Q := fun _:bool => B). reflexivity. @@ -134,8 +134,8 @@ right. intro HP. assert (Hequiv : forall b:bool, class_of_true b <-> class_of_false b). intro b; split. -unfold class_of_false in |- *; right; assumption. -unfold class_of_true in |- *; right; assumption. +unfold class_of_false; right; assumption. +unfold class_of_true; right; assumption. assert (Heq : class_of_true = class_of_false). apply pred_extensionality with (1 := Hequiv). apply diff_true_false. @@ -188,8 +188,8 @@ Lemma projT1_injective : a1=a2 -> a1'=a2'. Proof. intro Heq ; unfold a1', a2', A'. rewrite Heq. - replace (or_introl (a2=a2) (refl_equal a2)) - with (or_intror (a2=a2) (refl_equal a2)). + replace (or_introl (a2=a2) (eq_refl a2)) + with (or_intror (a2=a2) (eq_refl a2)). reflexivity. apply proof_irrelevance. Qed. @@ -265,7 +265,7 @@ End ProofIrrel_RelChoice_imp_EqEM. (** Proof sketch from Bell [Bell93] (with thanks to P. Castéran) *) -Notation Local inhabited A := A (only parsing). +Local Notation inhabited A := A (only parsing). Section ExtensionalEpsilon_imp_EM. @@ -279,7 +279,7 @@ Hypothesis epsilon_extensionality : forall (A:Type) (i:inhabited A) (P Q:A->Prop), (forall a, P a <-> Q a) -> epsilon A i P = epsilon A i Q. -Notation Local eps := (epsilon bool true) (only parsing). +Local Notation eps := (epsilon bool true) (only parsing). Theorem extensional_epsilon_imp_EM : forall P:Prop, P \/ ~ P. Proof. diff --git a/theories/Logic/Epsilon.v b/theories/Logic/Epsilon.v index cb8f8a73..da3e5b08 100644 --- a/theories/Logic/Epsilon.v +++ b/theories/Logic/Epsilon.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Logic/Eqdep.v b/theories/Logic/Eqdep.v index b8e99036..6841334f 100644 --- a/theories/Logic/Eqdep.v +++ b/theories/Logic/Eqdep.v @@ -1,7 +1,7 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Logic/EqdepFacts.v b/theories/Logic/EqdepFacts.v index d84cd824..a22f286e 100644 --- a/theories/Logic/EqdepFacts.v +++ b/theories/Logic/EqdepFacts.v @@ -1,7 +1,7 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -101,7 +101,7 @@ Section Dependent_Equality. forall (p q:U) (x:P p) (y:P q), eq_dep p x q y -> eq_dep1 p x q y. Proof. destruct 1. - apply eq_dep1_intro with (refl_equal p). + apply eq_dep1_intro with (eq_refl p). simpl; trivial. Qed. @@ -121,7 +121,7 @@ Proof. apply eq_dep_intro. Qed. -Notation eq_sigS_eq_dep := eq_sigT_eq_dep (only parsing). (* Compatibility *) +Notation eq_sigS_eq_dep := eq_sigT_eq_dep (compat "8.2"). (* Compatibility *) Lemma eq_dep_eq_sigT : forall (U:Type) (P:U -> Type) (p q:U) (x:P p) (y:P q), @@ -250,12 +250,12 @@ Section Equivalences. (** Uniqueness of Reflexive Identity Proofs *) Definition UIP_refl_ := - forall (x:U) (p:x = x), p = refl_equal x. + forall (x:U) (p:x = x), p = eq_refl x. (** Streicher's axiom K *) Definition Streicher_K_ := - forall (x:U) (P:x = x -> Prop), P (refl_equal x) -> forall p:x = x, P p. + forall (x:U) (P:x = x -> Prop), P (eq_refl x) -> forall p:x = x, P p. (** Injectivity of Dependent Equality is a consequence of *) (** Invariance by Substitution of Reflexive Equality Proof *) @@ -389,14 +389,14 @@ Proof (eq_dep_eq__UIP U eq_dep_eq). (** Uniqueness of Reflexive Identity Proofs is a direct instance of UIP *) -Lemma UIP_refl : forall (x:U) (p:x = x), p = refl_equal x. +Lemma UIP_refl : forall (x:U) (p:x = x), p = eq_refl x. Proof (UIP__UIP_refl U UIP). (** Streicher's axiom K is a direct consequence of Uniqueness of Reflexive Identity Proofs *) Lemma Streicher_K : - forall (x:U) (P:x = x -> Prop), P (refl_equal x) -> forall p:x = x, P p. + forall (x:U) (P:x = x -> Prop), P (eq_refl x) -> forall p:x = x, P p. Proof (UIP_refl__Streicher_K U UIP_refl). End Axioms. diff --git a/theories/Logic/Eqdep_dec.v b/theories/Logic/Eqdep_dec.v index 59088aa7..3a6f6a23 100644 --- a/theories/Logic/Eqdep_dec.v +++ b/theories/Logic/Eqdep_dec.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -9,7 +9,7 @@ (* Created by Bruno Barras, Jan 1998 *) (* Made a module instance for EqdepFacts by Hugo Herbelin, Mar 2006 *) -(** We prove that there is only one proof of [x=x], i.e [refl_equal x]. +(** We prove that there is only one proof of [x=x], i.e [eq_refl x]. This holds if the equality upon the set of [x] is decidable. A corollary of this theorem is the equality of the right projections of two equal dependent pairs. @@ -43,7 +43,7 @@ Section EqdepDec. Let comp (x y y':A) (eq1:x = y) (eq2:x = y') : y = y' := eq_ind _ (fun a => a = y') eq2 _ eq1. - Remark trans_sym_eq : forall (x y:A) (u:x = y), comp u u = refl_equal y. + Remark trans_sym_eq : forall (x y:A) (u:x = y), comp u u = eq_refl y. Proof. intros. case u; trivial. @@ -61,7 +61,7 @@ Section EqdepDec. Let nu_constant : forall (y:A) (u v:x = y), nu u = nu v. intros. - unfold nu in |- *. + unfold nu. case (eq_dec x y); intros. reflexivity. @@ -69,13 +69,13 @@ Section EqdepDec. Qed. - Let nu_inv (y:A) (v:x = y) : x = y := comp (nu (refl_equal x)) v. + Let nu_inv (y:A) (v:x = y) : x = y := comp (nu (eq_refl x)) v. Remark nu_left_inv : forall (y:A) (u:x = y), nu_inv (nu u) = u. Proof. intros. - case u; unfold nu_inv in |- *. + case u; unfold nu_inv. apply trans_sym_eq. Qed. @@ -90,10 +90,10 @@ Section EqdepDec. Qed. Theorem K_dec : - forall P:x = x -> Prop, P (refl_equal x) -> forall p:x = x, P p. + forall P:x = x -> Prop, P (eq_refl x) -> forall p:x = x, P p. Proof. intros. - elim eq_proofs_unicity with x (refl_equal x) p. + elim eq_proofs_unicity with x (eq_refl x) p. trivial. Qed. @@ -115,7 +115,7 @@ Section EqdepDec. Proof. intros. cut (proj (ex_intro P x y) y = proj (ex_intro P x y') y). - simpl in |- *. + simpl. case (eq_dec x x). intro e. elim e using K_dec; trivial. @@ -135,7 +135,7 @@ Require Import EqdepFacts. Theorem K_dec_type : forall A:Type, (forall x y:A, {x = y} + {x <> y}) -> - forall (x:A) (P:x = x -> Prop), P (refl_equal x) -> forall p:x = x, P p. + forall (x:A) (P:x = x -> Prop), P (eq_refl x) -> forall p:x = x, P p. Proof. intros A eq_dec x P H p. elim p using K_dec; intros. @@ -146,7 +146,7 @@ Qed. Theorem K_dec_set : forall A:Set, (forall x y:A, {x = y} + {x <> y}) -> - forall (x:A) (P:x = x -> Prop), P (refl_equal x) -> forall p:x = x, P p. + forall (x:A) (P:x = x -> Prop), P (eq_refl x) -> forall p:x = x, P p. Proof fun A => K_dec_type (A:=A). (** We deduce the [eq_rect_eq] axiom for (decidable) types *) @@ -212,13 +212,13 @@ Module DecidableEqDep (M:DecidableType). (** Uniqueness of Reflexive Identity Proofs *) - Lemma UIP_refl : forall (x:U) (p:x = x), p = refl_equal x. + Lemma UIP_refl : forall (x:U) (p:x = x), p = eq_refl x. Proof (UIP__UIP_refl U UIP). (** Streicher's axiom K *) Lemma Streicher_K : - forall (x:U) (P:x = x -> Prop), P (refl_equal x) -> forall p:x = x, P p. + forall (x:U) (P:x = x -> Prop), P (eq_refl x) -> forall p:x = x, P p. Proof (K_dec_type eq_dec). (** Injectivity of equality on dependent pairs in [Type] *) @@ -281,13 +281,13 @@ Module DecidableEqDepSet (M:DecidableSet). (** Uniqueness of Reflexive Identity Proofs *) - Lemma UIP_refl : forall (x:U) (p:x = x), p = refl_equal x. + Lemma UIP_refl : forall (x:U) (p:x = x), p = eq_refl x. Proof N.UIP_refl. (** Streicher's axiom K *) Lemma Streicher_K : - forall (x:U) (P:x = x -> Prop), P (refl_equal x) -> forall p:x = x, P p. + forall (x:U) (P:x = x -> Prop), P (eq_refl x) -> forall p:x = x, P p. Proof N.Streicher_K. (** Proof-irrelevance on subsets of decidable sets *) @@ -301,7 +301,7 @@ Module DecidableEqDepSet (M:DecidableSet). Lemma inj_pair2 : forall (P:U -> Type) (p:U) (x y:P p), - existS P p x = existS P p y -> x = y. + existT P p x = existT P p y -> x = y. Proof eq_dep_eq__inj_pair2 U N.eq_dep_eq. (** Injectivity of equality on dependent pairs with second component diff --git a/theories/Logic/ExtensionalityFacts.v b/theories/Logic/ExtensionalityFacts.v index f5e71ef4..9cbf756d 100644 --- a/theories/Logic/ExtensionalityFacts.v +++ b/theories/Logic/ExtensionalityFacts.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Logic/FunctionalExtensionality.v b/theories/Logic/FunctionalExtensionality.v index 35db160f..ecb7428e 100644 --- a/theories/Logic/FunctionalExtensionality.v +++ b/theories/Logic/FunctionalExtensionality.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Logic/Hurkens.v b/theories/Logic/Hurkens.v index bb03c666..1dce51b2 100644 --- a/theories/Logic/Hurkens.v +++ b/theories/Logic/Hurkens.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -46,7 +46,7 @@ Lemma Omega : forall i:U -> bool, induct i -> b2p (i WF). Proof. intros i y. apply y. -unfold le, WF, induct in |- *. +unfold le, WF, induct. apply p2p2. intros x H0. apply y. @@ -55,7 +55,7 @@ Qed. Lemma lemma1 : induct (fun u => p2b (I u)). Proof. -unfold induct in |- *. +unfold induct. intros x p. apply (p2p2 (I x)). intro q. diff --git a/theories/Logic/IndefiniteDescription.v b/theories/Logic/IndefiniteDescription.v index 8badc07c..5424eea8 100644 --- a/theories/Logic/IndefiniteDescription.v +++ b/theories/Logic/IndefiniteDescription.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Logic/JMeq.v b/theories/Logic/JMeq.v index 753009e6..530e0555 100644 --- a/theories/Logic/JMeq.v +++ b/theories/Logic/JMeq.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Logic/ProofIrrelevance.v b/theories/Logic/ProofIrrelevance.v index 36508969..7d6d0cf8 100644 --- a/theories/Logic/ProofIrrelevance.v +++ b/theories/Logic/ProofIrrelevance.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Logic/ProofIrrelevanceFacts.v b/theories/Logic/ProofIrrelevanceFacts.v index 6accc480..2e9f0c19 100644 --- a/theories/Logic/ProofIrrelevanceFacts.v +++ b/theories/Logic/ProofIrrelevanceFacts.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -25,7 +25,7 @@ Module ProofIrrelevanceTheory (M:ProofIrrelevance). forall (U:Type) (p:U) (Q:U -> Type) (x:Q p) (h:p = p), x = eq_rect p Q x p h. Proof. - intros; rewrite M.proof_irrelevance with (p1:=h) (p2:=refl_equal p). + intros; rewrite M.proof_irrelevance with (p1:=h) (p2:=eq_refl p). reflexivity. Qed. End Eq_rect_eq. diff --git a/theories/Logic/RelationalChoice.v b/theories/Logic/RelationalChoice.v index d0d58e37..efec03d4 100644 --- a/theories/Logic/RelationalChoice.v +++ b/theories/Logic/RelationalChoice.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Logic/SetIsType.v b/theories/Logic/SetIsType.v index f0876fbc..c0a6f9ed 100644 --- a/theories/Logic/SetIsType.v +++ b/theories/Logic/SetIsType.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) |