├── README.md └── rocq ├── Commands.md ├── README.md ├── Tactics.md ├── aczel.v ├── arec.v ├── ars.v ├── bctt.v ├── bsat.v ├── certsig.v ├── certsum.v ├── cty.v ├── demo-gs.v ├── ediv.v ├── elim.v ├── ewo.v ├── exquant.v ├── ext.v ├── finty.v ├── gs.v ├── ids_gcd.v ├── ids_less.v ├── ids_less_binary.v ├── indeq.v ├── indind.v ├── la.v ├── leibniz.v ├── list.v ├── lwo.v ├── more.v ├── numeral.v ├── pairing.v ├── pat.v ├── pds.v ├── post.v ├── provability.v ├── regexp.v ├── sdec.v ├── solver.v ├── syntax.v ├── teaser-bctt.v ├── teaser-cert.v ├── teaser-elim.v ├── teaser-ewo.v ├── teaser-exquant.v ├── teaser-gs.v ├── teaser-indind.v ├── teaser-la.v ├── teaser-leibniz.v ├── teaser-list.v ├── teaser-pat.v ├── teaser-sigma.v ├── teaser-sizerec.v ├── undec.v ├── vector.v ├── wfrec.v └── xm.v /README.md: -------------------------------------------------------------------------------- 1 | # MPCTT Textbook Project 2 | Modeling and Proving in Computational Type Theory 3 | Using the Rocq Prover 4 | Gert Smolka 5 | 6 | [pdf](https://www.ps.uni-saarland.de/~smolka/drafts/mpctt.pdf) 7 | [Rocq](rocq) 8 | -------------------------------------------------------------------------------- /rocq/Commands.md: -------------------------------------------------------------------------------- 1 | # Commands used 2 | ## Definitions 3 | Definition, Fixpoint 4 | Fact, Lemma, Theorem, Corollary, Example, Goal 5 | Proof, Qed, Defined, Admitted, Abort, Show Proof 6 | ## Implicit arguments, implicit types, simplification 7 | Arguments, Implicit Types, Opaque 8 | ## Sections 9 | Section, Variable, End 10 | ## Modules 11 | Module, End, Import, Parameter 12 | From Coq Require Import Lia 13 | ## Notation 14 | Notation 15 | ## Services 16 | Check, Compute, Eval, Scheme 17 | Print, Print All, About, Locate, Search, Search concl: 18 | Print Assumptions 19 | ## System settings 20 | Set Printing All, Implicit, Universes 21 | Unset Printing ... 22 | Unset Elimination Schemes 23 | ## Demo 24 | Fail, Restart 25 | 26 | 27 | -------------------------------------------------------------------------------- /rocq/README.md: -------------------------------------------------------------------------------- 1 | # Directory of Rocq Files 2 | - Getting Started [gs.v](gs.v) 3 | - Basic Computational Type Theory [bctt.v](bctt.v) 4 | - Propositions as Types [pat.v](pat.v) 5 | - Leibniz Equality [leibniz.v](leibniz.v) 6 | - Inductive Eliminators [elim.v](elim.v) 7 | - Existential Quantification [exquant.v](exquant.v) 8 | - Arithmetic Pairing [pairing.v](pairing.v) 9 | - Abstract Syntax [syntax.v](syntax.v) 10 | - Certifying Functions and Sum Types [certsum.v](certsum.v) 11 | - Certifying Functions and Sigma Types [certsig.v](certsig.v) 12 | - Linear Arithmetic [la.v](la.v) 13 | - More Types [more.v](more.v) 14 | - Indexed Inductives [indind.v](indind.v) 15 | - Extensionality [ext.v](ext.v) 16 | - Excluded Middle and Double Negation [xm.v](xm.v) 17 | - Provability [provability.v](provability.v) 18 | - Least Witness Operators [lwo.v](lwo.v) 19 | - Arithmetic Recursion [arec.v](arec.v) 20 | - Euclidean Division [ediv.v](ediv.v) 21 | - Lists [list.v](list.v) 22 | - EWOs [ewo.v](ewo.v) 23 | - Finite Types [finty.v](finty.v) 24 | - Countable Types [cty.v](cty.v) 25 | - Propositional Deduction [pds.v](pds.v), [solver.v](solver.v) 26 | - Boolean Satisfiability [bsat.v](bsat.v) 27 | - Regular Expression Matching [regexp.v](regexp.v) 28 | - Abstract Reduction Systems [ars.v](ars.v) 29 | - Axiom CT and Semidecidability [undec.v](undec.v) 30 | - Inductive Equality [indeq.v](indeq.v) 31 | - Well-Founded Recursion [wfrec.v](wfrec.v) 32 | - Aczel Trees and Hierarchy Theorems [aczel.v](aczel.v) 33 | - [List of tactics used](Tactics.md) 34 | - [List of commands used](Commands.md) 35 | 36 | Some chapters have teaser files 37 | that may be used for demos in lectures. 38 | For instance, 39 | [teaser-gs.v](teaser-gs.v) and [demo-gs.v](demo-gs.v) for "Getting Started". 40 | -------------------------------------------------------------------------------- /rocq/Tactics.md: -------------------------------------------------------------------------------- 1 | # Tactics used 2 | ## Basic 3 | exact, refine, intros, apply 4 | ## Conversion 5 | hnf, cbn, cbv, unfold, fold, change, pattern 6 | ## Equality 7 | reflexivity, rewrite, f_equal, symmetry, transitivity, replace 8 | ## Inductive types 9 | constructor, destruct, induction, discriminate, injection 10 | ## Logic 11 | split, left, right, exfalso, contradict, exists 12 | ## Goal management 13 | revert, clear, assert, enough, specialize, generalize, pose, set 14 | ## Basic automation 15 | assumption, trivial, easy, subst, auto 16 | ## More Automation 17 | lia, nia, tauto, congruence, intuition, firstorder, decide equality 18 | ## Existential Variables 19 | eapply, eexists, eassumption, eauto 20 | ## Tacticals 21 | ";", now, repeat, try, unshelve 22 | -------------------------------------------------------------------------------- /rocq/aczel.v: -------------------------------------------------------------------------------- 1 | Unset Elimination Schemes. 2 | 3 | Module Aczel. 4 | Set Printing Universes. 5 | Inductive tree : Type := T (X: Type) (f: X -> tree). 6 | Implicit Types s t: tree. 7 | 8 | Fail Check T tree (fun s => s). 9 | 10 | Lemma tree_ind (p: tree -> Type) : 11 | (forall X f, (forall x, p (f x)) -> p (T X f)) -> forall s, p s. 12 | Proof. 13 | intros F. 14 | refine (fix G s := match s with T X f => _ end). 15 | exact (F X f (fun x => (G (f x)))). 16 | Qed. 17 | 18 | Definition sub s t : Prop := match t with T X f => exists x, f x = s end. 19 | 20 | Lemma sub_irrefl : 21 | forall s, ~sub s s. 22 | Proof. 23 | apply tree_ind. intros X f IH [x H]. 24 | apply (IH x). 25 | replace (sub (f x) (f x)) with (sub (f x) (T X f)) by congruence. 26 | unfold sub. eauto. 27 | Qed. 28 | 29 | End Aczel. 30 | 31 | Module Demo. 32 | Set Printing Universes. 33 | Definition Type_i := Type. 34 | Fail Inductive tree : Type_i := T (X: Type_i) (f: X -> tree). 35 | Inductive tree : Prop := T (X: Type_i) (f: X -> tree). 36 | Module T. 37 | Inductive tree : Type_i := T_i (X: Prop) (f: X -> tree). 38 | End T. 39 | Fail Inductive L (X: Type_i) : Set := C (_: X). 40 | Inductive L (X: Type_i) : Set := N | C (_: L (prod nat X)). 41 | End Demo. 42 | 43 | Module XM_PI. 44 | Inductive tree: Prop := T (X: Prop) (f: X -> tree). 45 | Implicit Types s t: tree. 46 | 47 | Definition u: tree := T tree (fun s => s). 48 | 49 | Lemma tree_ind (p: tree -> Prop) : 50 | (forall X f, (forall x, p (f x)) -> p (T X f)) -> forall s, p s. 51 | Proof. 52 | intros F. 53 | refine (fix G s := match s with T X f => _ end). 54 | exact (F X f (fun x => (G (f x)))). 55 | Qed. 56 | 57 | Lemma sub_contra (R: tree -> tree -> Prop) : 58 | ~ (forall s X f, R s (T X f) <-> exists x, f x = s). 59 | Proof. 60 | intros H. 61 | assert (H_irrefl: forall s, ~R s s). 62 | { apply tree_ind. intros X f IH [x H1]%H. 63 | apply (IH x). 64 | replace (R (f x) (f x)) with (R (f x) (T X f)) by congruence. 65 | apply H. eauto. } 66 | apply (H_irrefl u), H. eauto. 67 | Qed. 68 | 69 | Theorem Rocquand (A: Prop) (E: Prop -> A) (D: A -> Prop) : 70 | ~ forall P, D (E P) <-> P. 71 | Proof. 72 | intros H. 73 | apply (sub_contra (fun s t => D (match t with T X f => E (exists x, f x = s) end))). 74 | intros s X f. apply H. 75 | Qed. 76 | 77 | Lemma elim_or (X Y: Prop) (p: X \/ Y -> Prop) : 78 | (forall x, p (or_introl x)) -> 79 | (forall y, p (or_intror y)) -> 80 | forall a, p a. 81 | Proof. 82 | exact (fun f g a => match a with or_introl x => f x | or_intror y => g y end). 83 | Qed. 84 | 85 | Theorem xm_pi : 86 | (forall P: Prop, P \/ ~P) -> forall (P: Prop) (a b: P), a = b. 87 | Proof. 88 | intros H A a b. 89 | destruct (H (a=b)) as [H1|H1]. exact H1. exfalso. 90 | apply (Rocquand A (fun P => if H P then a else b) (fun c => a = c)). 91 | intros P. 92 | (* destruct (H P) as [H2|H2]; tauto. *) 93 | pattern (H P). apply elim_or; tauto. 94 | Qed. 95 | End XM_PI. 96 | 97 | Module Hierarchy. 98 | Definition embeds X Y := 99 | exists (E: X -> Y) (D: Y -> X), forall x, D (E x) = x. 100 | 101 | Fact embeds_refl X : 102 | embeds X X. 103 | Proof. 104 | exists (fun x => x), (fun x => x). reflexivity. 105 | Qed. 106 | 107 | Definition Tyi := Type. 108 | 109 | Section Hier. 110 | Variables (A: Tyi) (D: A -> Tyi). 111 | 112 | Inductive tree: Tyi:= T (a: A) (f: D a -> tree). 113 | 114 | Implicit Types s t : tree. 115 | 116 | Lemma tree_ind (p: tree -> Type) : 117 | (forall a f, (forall x, p (f x)) -> p (T a f)) -> forall s, p s. 118 | Proof. 119 | intros F. 120 | refine (fix G s := match s with T a f => _ end). 121 | exact (F a f (fun x => (G (f x)))). 122 | Qed. 123 | 124 | Definition sub s t : Prop := 125 | match t with T a f => exists x, f x = s end. 126 | 127 | Fact sub_irrefl : 128 | forall s, ~ sub s s. 129 | Proof. 130 | apply tree_ind. intros a f IH [x H]. 131 | apply (IH x). 132 | replace (sub (f x) (f x)) with (sub (f x) (T a f)) by congruence. 133 | unfold sub. eauto. 134 | Qed. 135 | 136 | Lemma hier (E: Tyi -> A) : 137 | ~ embeds tree (D (E tree)). 138 | Proof. 139 | intros (F&G&H). 140 | apply (sub_irrefl (T (E tree) G)). 141 | cbn. exists (F (T (E tree) G)). apply H. 142 | Qed. 143 | End Hier. 144 | 145 | Theorem hierarchy : 146 | forall X:Tyi, ~embeds Tyi X. 147 | Proof. 148 | intros X (E&D&H). 149 | apply (hier X D E). rewrite H. apply embeds_refl. 150 | Qed. 151 | 152 | Goal forall X:Tyi, Tyi <> X. 153 | Proof. 154 | intros X H. apply (hierarchy X). rewrite H. apply embeds_refl. 155 | Qed. 156 | 157 | #[bypass_check(universes)] 158 | Example test: False. 159 | Proof. 160 | apply (hierarchy Tyi). 161 | apply embeds_refl. 162 | Qed. 163 | 164 | Print Assumptions test. 165 | 166 | End Hierarchy. 167 | 168 | 169 | 170 | 171 | -------------------------------------------------------------------------------- /rocq/arec.v: -------------------------------------------------------------------------------- 1 | (*** MPCTT, Chapter Arithmetic Recursion *) 2 | From Stdlib Require Import Lia. 3 | Definition dec (X: Type) : Type := X + (X -> False). 4 | Definition eqdec X := forall x y: X, dec (x = y). 5 | 6 | (*** Complete induction *) 7 | 8 | Definition complete_ind {p: nat -> Type} : 9 | (forall x, (forall y, y < x -> p y) -> p x) -> forall x, p x. 10 | Proof. 11 | intros H x. apply H. 12 | induction x as [|x IH]. lia. 13 | intros y H1. apply H. intros z H3. 14 | apply IH. lia. 15 | Qed. 16 | 17 | Definition Fib f n := 18 | if n - 1 then n else f (n - 2) + f (n - 1). 19 | 20 | Fact Fib_unique f f' : 21 | (forall n, f n = Fib f n) -> 22 | (forall n, f' n = Fib f' n) -> 23 | forall n, f n = f' n. 24 | Proof. 25 | intros H1 H2. 26 | refine (complete_ind _). intros n IH. 27 | rewrite H1, H2. 28 | destruct n. reflexivity. 29 | destruct n. reflexivity. 30 | cbn. rewrite IH. rewrite IH. 31 | reflexivity. lia. lia. 32 | Qed. 33 | 34 | (*** Size Induction *) 35 | 36 | Definition size_ind {X} (sigma: X -> nat) {p: X -> Type} : 37 | (forall x, (forall y, sigma y < sigma x -> p y) -> p x) -> 38 | forall x, p x. 39 | Proof. 40 | intros H. 41 | enough (forall n x, sigma x < n -> p x) by eauto. 42 | induction n as [|n IH]. lia. 43 | intros x H1. apply H. intros y H2. 44 | apply IH. lia. 45 | Qed. 46 | 47 | Definition size_ind2 {X Y} (sigma: X -> Y -> nat) {p: X -> Y -> Type} : 48 | (forall x y, (forall x' y', sigma x' y' < sigma x y -> p x' y') -> p x y) -> 49 | forall x y, p x y. 50 | Proof. 51 | intros H. 52 | enough (forall n x y, sigma x y < n -> p x y) by eauto. 53 | induction n as [|n IH]. lia. 54 | intros x y H1. apply H. intros x' y' H2. 55 | apply IH. lia. 56 | Qed. 57 | 58 | (*** Euclidean Quotient *) 59 | 60 | Definition DIV f x y := 61 | if x - y then 0 else S (f (x - S y) y). 62 | 63 | Notation "f == f'" := (forall x y, f x y = f' x y) (at level 70). 64 | 65 | Fact DIV_unique f f' : 66 | f == DIV f -> f' == DIV f' -> f == f'. 67 | Proof. 68 | intros H1 H2 x y. 69 | revert x. apply complete_ind. intros x IH. 70 | rewrite H1, H2. unfold DIV. 71 | destruct (x - y) eqn:?. reflexivity. 72 | f_equal. apply IH. lia. 73 | Qed. 74 | 75 | Fact DIV_unique' f f' : 76 | f == DIV f -> f' == DIV f' -> f == f'. 77 | Proof. 78 | intros H1 H2 x y. revert x. 79 | enough (forall n x, x < n -> f x y = f' x y) by eauto. 80 | induction n as [|n IH]; intros x H. 81 | - exfalso; lia. 82 | - rewrite H1, H2. unfold DIV. 83 | destruct (x - y) eqn:?. 84 | + reflexivity. 85 | + f_equal. apply IH. lia. 86 | Qed. 87 | 88 | (*** Step-Indexed Construction *) 89 | 90 | Fixpoint Div n x y := 91 | match n with 92 | | 0 => 0 93 | | S n => DIV (Div n) x y 94 | end. 95 | 96 | Fact Div_index_independence n1 n2 x y : 97 | n1 > x -> n2 > x -> Div n1 x y = Div n2 x y. 98 | Proof. 99 | induction n1 as [|n1 IH] in n2, x|-*; intros H1 H2. 100 | - exfalso; lia. 101 | - destruct n2. exfalso; lia. 102 | cbn. unfold DIV. 103 | destruct (x - y) eqn:?. reflexivity. 104 | f_equal. apply IH; lia. 105 | Qed. 106 | 107 | Definition D x := Div (S x) x. 108 | 109 | Fact D_correct : 110 | D == DIV D. 111 | Proof. 112 | intros x y. cbn. unfold DIV. 113 | destruct (x - y) eqn:?. reflexivity. 114 | f_equal. apply Div_index_independence; lia. 115 | Qed. 116 | 117 | Compute D 33 4. 118 | 119 | (*** GCD *) 120 | 121 | Definition GCD f (x y: nat) : nat := 122 | if x then y else 123 | if x - y then f x (y - x) else f y x. 124 | 125 | Definition sigma x y := 2*x + y. 126 | 127 | Fact GCD_unique f f' : 128 | f == GCD f -> f' == GCD f' -> f == f'. 129 | Proof. 130 | intros H1 H2. 131 | apply (size_ind2 sigma). 132 | intros x y IH. 133 | rewrite H1, H2. unfold GCD. 134 | destruct x. reflexivity. 135 | destruct (S x - y) as [|d] eqn:H3. 136 | - apply IH. unfold sigma; lia. 137 | - apply IH. unfold sigma; lia. 138 | Qed. 139 | 140 | Fixpoint gcd_index n x y := 141 | match n with 142 | | 0 => 0 143 | | S n => GCD (gcd_index n) x y 144 | end. 145 | 146 | Fact gcd_index_independence n1 n2 x y : 147 | n1 > sigma x y -> n2 > sigma x y -> gcd_index n1 x y = gcd_index n2 x y. 148 | Proof. 149 | induction n1 as [|n1 IH] in n2, x, y|-*; intros H1 H2. 150 | - exfalso; lia. 151 | - destruct n2. exfalso; lia. 152 | cbn. unfold GCD. 153 | destruct x. reflexivity. 154 | destruct (S x - y) as [|d] eqn:H3. 155 | + apply IH; unfold sigma in *; lia. 156 | + apply IH; unfold sigma in *; lia. 157 | Qed. 158 | 159 | Definition gcd x y := gcd_index (S (sigma x y)) x y. 160 | 161 | Compute gcd 16 24. 162 | Compute gcd 60 48. 163 | Compute gcd 175 5. 164 | 165 | Fact gcd_correct : 166 | gcd == GCD gcd. 167 | Proof. 168 | intros x y. cbn. unfold GCD. 169 | destruct x. reflexivity. 170 | destruct (S x - y) eqn:?. 171 | - apply gcd_index_independence; unfold sigma; lia. 172 | - apply gcd_index_independence; unfold sigma; lia. 173 | Qed. 174 | 175 | Fact GCD_unique' f f' : 176 | f == GCD f -> f' == GCD f' -> f == f'. 177 | Proof. 178 | intros H1 H2. 179 | enough (forall n x y, sigma x y < n -> f x y = f' x y) by eauto. 180 | induction n as [|n IH]; intros x y H. 181 | - exfalso; lia. 182 | - rewrite H1, H2. unfold GCD. 183 | destruct x. reflexivity. 184 | destruct (S x - y) eqn:?. 185 | + apply IH. unfold sigma in * ; lia. 186 | + apply IH. unfold sigma in * ; lia. 187 | Qed. 188 | 189 | Fact gcd_rec (p: nat -> nat -> Type) : 190 | (forall y, p 0 y) -> 191 | (forall x y, x <= y -> p x (y - x) -> p x y) -> 192 | (forall x y, p x y -> p y x) -> 193 | forall x y, p x y. 194 | Proof. 195 | intros H1 H2 H3. 196 | apply (size_ind2 sigma). 197 | intros x y IH. 198 | destruct x. {apply H1.} 199 | destruct (S x - y) eqn:?. 200 | - apply H2. lia. apply IH. unfold sigma; lia. 201 | - apply H3, IH. unfold sigma; lia. 202 | Qed. 203 | 204 | (*** Course-of-values Recursion *) 205 | 206 | Fixpoint vec n X : Type := 207 | match n with 208 | | 0 => unit 209 | | S n => X * vec n X 210 | end. 211 | 212 | Compute vec 3 nat. 213 | Check (1,(2,(3,tt))) : vec 3 nat. 214 | 215 | Fixpoint vecrec {X} (f: forall n, vec n X -> X) n : vec n X := 216 | match n with 217 | | 0 => tt 218 | | S n => let v := vecrec f n in (f n v, v) 219 | end. 220 | 221 | Definition covrec {X} f n : X := fst (vecrec f (S n)). 222 | 223 | Definition vecnat : forall n, vec n nat := vecrec (fun n _ => n). 224 | 225 | Compute vecnat 5. 226 | Compute covrec (fun n _ => n*n) 5. 227 | 228 | Definition facstep n : vec n nat -> nat := 229 | match n with 230 | | 0 => fun _ => 1 231 | | S _ => fun v => n * fst v 232 | end. 233 | 234 | Compute covrec facstep 5. 235 | 236 | Goal let fac := covrec facstep in 237 | forall n, fac n = if n then 1 else n * fac (n - 1). 238 | Proof. 239 | intros fac n. 240 | destruct n. reflexivity. 241 | replace (S n - 1) with n by lia. 242 | reflexivity. 243 | Qed. 244 | 245 | Definition fibstep n : vec n nat -> nat := 246 | match n with 247 | | 0 => fun _ => 0 248 | | 1 => fun _ => 1 249 | | S (S _) => fun v => fst (snd v) + fst v 250 | end. 251 | 252 | Definition fib := covrec fibstep. 253 | 254 | Compute fib 5. 255 | 256 | Fact fib_correct n : 257 | fib n = if n - 1 then n else fib (n - 2) + fib (n - 1). 258 | Proof. 259 | destruct n. reflexivity. 260 | destruct n. reflexivity. 261 | replace (S (S n) - 2) with n by lia. 262 | reflexivity. 263 | Qed. 264 | -------------------------------------------------------------------------------- /rocq/ars.v: -------------------------------------------------------------------------------- 1 | From Stdlib Require Import Lia. 2 | Definition iffT (X Y: Type) : Type := (X -> Y) * (Y -> X). 3 | Notation "X <=> Y" := (iffT X Y) (at level 95, no associativity). 4 | Ltac refl := reflexivity. 5 | 6 | Section Path. 7 | Variable X: Type. 8 | Variable R: X -> X -> Type. 9 | Implicit Types x y z: X. 10 | 11 | Inductive path (x: X) : X -> Type := 12 | | P1 : path x x 13 | | P2 : forall x' y, R x x' -> path x' y -> path x y. 14 | 15 | Definition len 16 | : forall {x y}, path x y -> nat 17 | := fix f x _ a := match a with 18 | | P1 _ => 0 19 | | P2 _ x' y r a => S (f x' y a) 20 | end. 21 | 22 | Definition app 23 | : forall {z x y}, path x y -> path y z -> path x z 24 | := fun z => fix f x _ a := match a with 25 | | P1 _ => fun b => b 26 | | P2 _ x' y r a' => fun b => P2 x x' z r (f x' y a' b) 27 | end. 28 | 29 | Fact app_P1_eq x z (b: path x z) : 30 | app (P1 x) b = b. 31 | Proof. 32 | refl. 33 | Qed. 34 | 35 | Fact app_P2_eq x x' y z (r: R x x') (a: path x' y) (b: path y z) : 36 | app (P2 x x' y r a) b = P2 x x' z r (app a b). 37 | Proof. 38 | refl. 39 | Qed. 40 | 41 | (* See the types of the pattern variables using refine *) 42 | Goal forall {z x y}, path x y -> path y z -> path x z. 43 | Proof. 44 | refine (fun z => fix f x _ a := 45 | match a with 46 | | P1 _ => fun b => _ 47 | | P2 _ x' y r a' => fun b => _ 48 | end). 49 | - exact b. 50 | - exact (P2 x x' z r (f x' y a' b)). 51 | Qed. 52 | 53 | Definition path_elim (p: forall x y, path x y -> Type) 54 | : (forall x, p x x (P1 x)) -> 55 | (forall x x' y r a, p x' y a -> p x y (P2 x x' y r a)) -> 56 | forall x y a, p x y a 57 | := fun e1 e2 => fix f x _ a := match a with 58 | | P1 _ => e1 x 59 | | P2 _ x' y r a' => e2 x x' y r a' (f x' y a') 60 | end. 61 | 62 | Goal forall x y z (a: path x y) (b: path y z), 63 | len (app a b) = len a + len b. 64 | Proof. 65 | intros x y z. revert x y. 66 | refine (path_elim _ _ _); cbn. 67 | - intros x b. refl. 68 | - intros x x' y r a IH b. f_equal. apply IH. 69 | Qed. 70 | 71 | Goal forall x y z (a: path x y) (b: path y z), 72 | len (app a b) = len a + len b. 73 | Proof. 74 | induction a as [x|x x' y r a IH]; cbn. 75 | - intros b. refl. 76 | - intros b. f_equal. apply IH. 77 | Qed. 78 | 79 | Fact path_step x y : 80 | R x y -> path x y. 81 | Proof. 82 | intros H. eapply P2. exact H. apply P1. 83 | Qed. 84 | 85 | Definition path_index (p: X -> X -> Type) 86 | : (forall x, p x x) -> 87 | (forall x x' y, R x x' -> p x' y -> p x y ) -> 88 | forall x y, path x y -> p x y 89 | := fun e1 e2 => fix f x _ a := match a with 90 | | P1 _ => e1 x 91 | | P2 _ x' y r a' => e2 x x' y r (f x' y a') 92 | end. 93 | 94 | Fact path_trans x y z : 95 | path x y -> path y z -> path x z. 96 | Proof. 97 | revert x y. 98 | refine (path_index _ _ _). 99 | - easy. 100 | - intros x x' y r IH H. 101 | eapply P2. exact r. apply IH, H. 102 | Qed. 103 | End Path. 104 | Arguments path {X}. 105 | Arguments path_index {X R p}. 106 | 107 | Module Exercise. 108 | Definition R x y := (S x = y). 109 | Goal forall x y, path R x y <=> x <= y. 110 | Proof. 111 | split. 112 | - revert x y. 113 | refine (path_index _ _). 114 | + lia. 115 | + intros x y' y <-. lia. 116 | - enough (forall k, path R x (k + x)) as H. 117 | + intros H1. generalize (H (y - x)). 118 | replace ((y - x) + x) with y by lia. 119 | easy. 120 | + induction k as [|k IH] in x |-*. 121 | * apply P1. 122 | * apply (P2 _ _ x (S x)). 123 | -- refl. 124 | -- replace (S k + x) with (k + S x) by lia. 125 | apply IH. 126 | Qed. 127 | End Exercise. 128 | 129 | Section Star. 130 | Variable X: Type. 131 | Implicit Type p: X -> X -> Type. 132 | 133 | Definition refl p := forall x, p x x. 134 | Definition trans p := forall x y z, p x y -> p y z -> p x z. 135 | Definition incl p p' := forall x y, p x y -> p' x y. 136 | 137 | Fact path_star p R : 138 | refl p -> trans p -> incl R p -> forall x y, path R x y -> p x y. 139 | Proof. 140 | intros H1 H2 H3. 141 | refine (path_index _ _). 142 | - exact H1. 143 | - intros x x' y H4 IH. 144 | eapply H2. 2: exact IH. apply H3, H4. 145 | Qed. 146 | 147 | Fact path_fun_char R x y : 148 | path R x y <=> forall p, refl p -> trans p -> incl R p -> p x y. 149 | Proof. 150 | split. 151 | - intros H1 p H2 H3 H4. 152 | revert x y H1. apply path_star; assumption. 153 | - intros H. apply H; hnf. 154 | + apply P1. 155 | + eauto using app. 156 | + apply path_step. 157 | Qed. 158 | 159 | Fact path_path R (x y: X) : 160 | path R x y <=> path (path R) x y. 161 | Proof. 162 | split. 163 | - apply path_step. 164 | - revert x y. apply path_star. 165 | + constructor. 166 | + intros x y z. apply app. 167 | + easy. 168 | Qed. 169 | End Star. 170 | 171 | 172 | 173 | 174 | 175 | 176 | 177 | 178 | 179 | 180 | 181 | -------------------------------------------------------------------------------- /rocq/bctt.v: -------------------------------------------------------------------------------- 1 | (*** MPCTT, Chapter 2 *) 2 | 3 | (* We show how Rocq realizes basic CTT *) 4 | 5 | (* We start with the definitions of the inductive types 6 | for booleans, numbers, and pairs. We also define 7 | some inductive functions for these types. 8 | We enclose the definitions in a module so that 9 | the predefined definitions can be used afterwards. *) 10 | 11 | Module Definitions. 12 | 13 | Inductive bool : Type := 14 | | true : bool 15 | | false : bool. 16 | 17 | Inductive nat : Type := 18 | | O : nat 19 | | S : nat -> nat. 20 | 21 | Inductive Pair (X Y: Type) : Type := 22 | | pair : X -> Y -> Pair X Y. 23 | 24 | Definition negb (x: bool) : bool := 25 | match x with 26 | | true => false 27 | | false => true 28 | end. 29 | 30 | Fixpoint add (x y: nat) : nat := 31 | match x with 32 | | O => y 33 | | S x => S (add x y) 34 | end. 35 | 36 | Fixpoint mul (x y: nat) : nat := 37 | match x with 38 | | O => O 39 | | S x => add y (mul x y) 40 | end. 41 | 42 | Fixpoint sub (x y: nat) : nat := 43 | match x, y with 44 | | O, _ => O 45 | | S x', O => x 46 | | S x', S y' => sub x' y' 47 | end. 48 | 49 | Check pair. 50 | Check pair nat. 51 | Check pair nat bool. 52 | Check pair nat bool (S O). 53 | Check pair nat bool (S O) false. 54 | Check pair _ _ (S O) false. 55 | 56 | (* declare implicit arguments *) 57 | Arguments pair {X Y}. 58 | Check pair (S O) false. 59 | Check @pair nat bool (S O). 60 | Check @pair. 61 | Check @pair nat. 62 | Check @pair nat bool. 63 | Check pair. 64 | 65 | Definition swap {X Y: Type} (a: Pair X Y) : Pair Y X := 66 | match a with 67 | | pair x y => pair y x 68 | end. 69 | 70 | Definition fst {X Y: Type} (a: Pair X Y) : X := 71 | match a with 72 | | pair x _ => x 73 | end. 74 | 75 | Check negb. 76 | Compute negb (negb (negb true)). 77 | Check mul. 78 | Compute mul (S (S O)) (S (S (S O))). 79 | Compute sub (S (S (S O))) (S (S O)). 80 | Check swap. 81 | Check swap (pair O true). 82 | Compute swap (pair O true). 83 | End Definitions. 84 | 85 | (** We now look at predefined inductive types and inductive functions *) 86 | 87 | Print bool. 88 | (** Read "Set" as "Type" *) 89 | Print nat. 90 | Locate "*". 91 | Print prod. 92 | Check S (S O). 93 | Compute 2 + 7. 94 | Locate "+". 95 | Check Nat.add. 96 | Check Nat.sub. 97 | 98 | (** Beta and zeta reduction *) 99 | 100 | Eval cbv beta in 101 | (fun x => x + x) (2 * 3). 102 | 103 | Eval cbv zeta in 104 | let x := 2 * 5 in x + x. 105 | 106 | (** Match reduction *) 107 | 108 | Eval cbv match in 109 | fun x:nat => match 0 with 0 => x | S y => x + y end. 110 | 111 | Eval cbv match in 112 | fun x:nat => match S (2 * x) with 0 => x | S y => x + y end. 113 | 114 | Eval cbv match in 115 | fun x:nat => match x with 0 => x | S y => x + y end. 116 | 117 | (** Delta reduction *) 118 | 119 | Definition T : nat -> nat := 120 | fun x => S (S x). 121 | 122 | Eval cbv delta in 123 | T 1 = 3. 124 | 125 | (* We switch to proof mode so that we can do reduction chains *) 126 | 127 | Goal 128 | T 1 = 3. 129 | Proof. 130 | cbv delta [T]. 131 | cbv beta. 132 | Abort. 133 | 134 | Definition P : nat -> nat := 135 | fun x => match x with 0 => 0 | S x' => x' end. 136 | 137 | Goal 138 | P 2 = 1. 139 | Proof. 140 | cbv delta [P]. 141 | cbv beta. 142 | cbv match. 143 | Abort. 144 | 145 | (** Fix reduction *) 146 | 147 | Definition D: nat -> nat := 148 | fix f x := match x with 0 => 0 | S x' => S (S (f x')) end. 149 | 150 | Goal 151 | D 1 = 2. 152 | Proof. 153 | cbv delta [D]. 154 | cbv fix. cbv beta. 155 | cbv match. 156 | cbv fix. cbv beta. 157 | cbv match. 158 | Abort. 159 | 160 | (** Readability can be improved by folding in the constant $D$ 161 | again after reduction of the recursive abstraction. 162 | This is routinely done by Rocq's simplification tactics. *) 163 | 164 | Goal 165 | D 1 = 2. 166 | Proof. 167 | cbv delta [D]. 168 | cbv fix. fold D. cbv beta. 169 | cbv match. 170 | cbv delta [D]. 171 | cbv fix. fold D. cbv beta. 172 | cbv match. 173 | Abort. 174 | 175 | Example demo n : 176 | D (S n) = S (S (D n)). 177 | Proof. 178 | (* We use set/subst so that delta affects only the left occurrence od D *) 179 | set (a:= S (S (D n))). 180 | cbv delta [D]. cbv fix. fold D. cbv beta. 181 | cbv match. cbv beta. 182 | subst a. 183 | Abort. 184 | 185 | Locate "+". 186 | Import Nat. (* Write add for Nat.add *) 187 | Print add. 188 | 189 | Example demo x : 190 | 2 + x = S (S x). 191 | Proof. 192 | cbv delta [add]. cbv fix. fold add. cbv beta. 193 | cbv match. cbv beta. 194 | cbv delta [add]. cbv fix. fold add. cbv beta. 195 | cbv match. cbv beta. 196 | cbv delta [add]. cbv fix. fold add. cbv beta. 197 | cbv match. 198 | Abort. 199 | 200 | Example demo x y : 201 | S (S x) + y = S (S (x + y)). 202 | Proof. 203 | set (a:= x+y). 204 | cbv delta [add]. cbv fix. fold add. cbv beta. 205 | cbv match. cbv beta. 206 | cbv delta [add]. cbv fix. fold add. cbv beta. 207 | cbv match. cbv beta. 208 | subst a. 209 | Abort. 210 | 211 | (** We compute normal forms of terms describing functions *) 212 | 213 | Compute add 0. 214 | Compute add 1. 215 | Compute add 2. 216 | Compute fun x => 3 + x - 2. 217 | 218 | (** Alpha renaming *) 219 | 220 | Goal 221 | (fun n: nat => n) = (fun x => x). 222 | Proof. 223 | reflexivity. 224 | Abort. 225 | 226 | (** Rocq has also eta conversion, explained in Chapter 4 of MPCTT. *) 227 | 228 | Goal forall X Y (f: X -> Y), 229 | (fun x => f x) = f. 230 | Proof. 231 | intros X Y f. 232 | cbv. 233 | reflexivity. 234 | Abort. 235 | 236 | Goal 237 | (fun x => 3 + x - 2) = S. 238 | Proof. 239 | cbv. 240 | reflexivity. 241 | Abort. 242 | -------------------------------------------------------------------------------- /rocq/certsig.v: -------------------------------------------------------------------------------- 1 | (*** MPCTT, Chapter Sigma Types *) 2 | 3 | From Stdlib Require Import Lia. 4 | 5 | Notation "~ X" := (X -> False) (at level 75, right associativity) : type_scope. 6 | Definition iffT (X Y: Type) : Type := (X -> Y) * (Y -> X). 7 | Notation "X <=> Y" := (iffT X Y) (at level 95, no associativity). 8 | Definition dec (X: Type) : Type := X + ~ X. 9 | Definition eqdec X := forall x y: X, dec (x = y). 10 | Definition decider {X} (p: X -> Type) := forall x, dec (p x). 11 | 12 | (*** Sigma types *) 13 | 14 | Module SigmaTypes. 15 | Inductive sig (X: Type) (p: X -> Type) : Type := 16 | | Sig : forall x, p x -> sig X p. 17 | 18 | Arguments sig {X}. 19 | Arguments Sig {X}. 20 | 21 | Definition elim_sig {X: Type} {p: X -> Type} (q: sig p -> Type) 22 | : (forall x y, q (Sig p x y)) -> forall a, q a 23 | := fun f a => match a with Sig _ x y => f x y end. 24 | 25 | Definition pi1 {X: Type} {p: X -> Type} 26 | : sig p -> X 27 | := fun a => match a with Sig _ x _ => x end. 28 | 29 | Definition pi2 {X: Type} {p: X -> Type} 30 | : forall a: sig p, p (pi1 a) 31 | := fun a => match a with Sig _ _ y => y end. 32 | 33 | Fact eta_law X (p: X -> Type) : 34 | forall a, a = Sig p (pi1 a) (pi2 a). 35 | Proof. 36 | apply elim_sig. cbn. reflexivity. 37 | Qed. 38 | 39 | Fact pi1_injective X (p: X -> Type) x y x' y' : 40 | Sig p x y = Sig p x' y' -> x = x'. 41 | Proof. 42 | intros H. 43 | change x with (pi1 (Sig p x y)). 44 | rewrite H. reflexivity. 45 | Qed. 46 | 47 | Fact pi2_injective_nondep X Y x y x' y' : 48 | Sig (fun _: X => Y) x y = Sig _ x' y' -> y = y'. 49 | Proof. 50 | intros H. 51 | change y with (pi2 (Sig _ x y)). 52 | rewrite H. reflexivity. 53 | Qed. 54 | 55 | (* Injectivity in the second component cannot be shown in general *) 56 | 57 | Fail 58 | Fact pi2_injective X (p: X -> Type) x y x' y' : 59 | Sig p x y = Sig p x' y' -> y = y'. 60 | 61 | Fact pi2_injective X (p: X -> Type) x y y' : 62 | Sig p x y = Sig p x y' -> y = y'. 63 | Proof. 64 | (* cannot be shown without assumptions *) 65 | intros H. 66 | change y with (pi2 (Sig p x y)). 67 | Fail pattern (Sig p x y). 68 | Fail rewrite H. 69 | Abort. 70 | 71 | Goal forall X (p: X -> Type), 72 | @pi1 X p = @elim_sig X p (fun _ => X) (fun x _ => x). 73 | Proof. 74 | reflexivity. 75 | Qed. 76 | 77 | Goal forall X (p: X -> Type), 78 | @pi2 X p = @elim_sig X p (fun a => p (pi1 a)) (fun _ y => y). 79 | Proof. 80 | reflexivity. 81 | Qed. 82 | 83 | End SigmaTypes. 84 | 85 | (* We shall use Rocq's predefined sigma types from now on. 86 | We rename the constructors and projections to better fit MPCTT. 87 | We also define the big sigma notation 88 | (replacing Rocq's curly braces notation *) 89 | 90 | Notation sig := sigT. 91 | Notation Sig := existT. 92 | Notation pi1 := projT1. 93 | Notation pi2 := projT2. 94 | Notation "'Sigma' x .. y , p" := 95 | (sig (fun x => .. (sig (fun y => p%type)) ..)) 96 | (at level 200, x binder, right associativity, 97 | format "'[' 'Sigma' '/ ' x .. y , '/ ' p ']'") 98 | : type_scope. 99 | 100 | (*** Certifying Division *) 101 | 102 | Goal 103 | forall x, Sigma y, (x = 2*y)%nat + (x = S (2*y)). 104 | Proof. 105 | induction x as [|x [y [IH|IH]]]. 106 | - exists 0. auto. 107 | - exists y. right. congruence. 108 | - exists (S y). left. lia. 109 | Qed. 110 | 111 | (* Note the "%nat" annotation in the statement of div2. It is 112 | needed to help with the overloading of "+" and "*" 113 | for numbers and types. *) 114 | 115 | Goal (forall x, Sigma y, (x = 2*y)%nat + (x = S (2*y))) <=> 116 | (Sigma D M, (forall x, x = 2 * D x + M x /\ M x < 2) %nat). 117 | Proof. 118 | split. 119 | - intros F. 120 | exists (fun x => pi1 (F x)). 121 | exists (fun x => if pi2 (F x) then 0 else 1). 122 | intros x. 123 | destruct (F x) as [y [H|H]]; cbn; lia. 124 | - intros (D&M&H) x. 125 | specialize (H x). 126 | exists (D x). 127 | destruct (M x) as [|b]. 128 | + left. lia. 129 | + right. lia. 130 | Qed. 131 | 132 | (*** Translation Theorems *) 133 | 134 | Fact trans_eqdec X : 135 | eqdec X <=> Sigma f: X -> X -> bool, forall x y, x = y <-> f x y = true. 136 | Proof. 137 | split. 138 | - intros d. 139 | exists (fun x y => if d x y then true else false). 140 | intros x y. 141 | destruct (d x y) as [H|H]; intuition congruence. 142 | - intros [f H] x y. specialize (H x y). 143 | destruct (f x y); unfold dec; intuition easy. 144 | Qed. 145 | 146 | Fact trans_eqdec' X : 147 | eqdec X <=> Sigma f: X -> X -> bool, forall x y, if f x y then x = y else x <> y. 148 | Proof. 149 | split. 150 | - intros d. 151 | exists (fun x y => if d x y then true else false). 152 | intros x y. 153 | destruct (d x y) as [H|H]; easy. 154 | - intros [f H] x y. specialize (H x y). 155 | destruct (f x y); unfold dec; intuition. 156 | Qed. 157 | 158 | Fact trans_p_dec X (p: X -> Type) : 159 | decider p <=> Sigma f: X -> bool, forall x, p x <=> f x = true. 160 | Proof. 161 | split. 162 | - intros d. 163 | exists (fun x => if d x then true else false). 164 | intros x. 165 | destruct (d x) as [H|H]; unfold iffT; intuition easy. 166 | - intros [f H] x. specialize (H x). 167 | destruct (f x); unfold dec, iffT in *; intuition easy. 168 | Qed. 169 | 170 | Fact trans_p_dec' X (p: X -> Type) : 171 | decider p <=> Sigma f: X -> bool, forall x, if f x then p x else ~ p x. 172 | Proof. 173 | split. 174 | - intros d. 175 | exists (fun x => if d x then true else false). 176 | intros x. 177 | destruct (d x) as [H|H]; unfold iffT; easy. 178 | - intros [f H] x. specialize (H x). 179 | destruct (f x); unfold dec, iffT in *; auto. 180 | Qed. 181 | 182 | Fact trans_skolem X Y (p: X -> Y -> Type) : 183 | (forall x, Sigma y, p x y) <=> Sigma f: X -> Y, forall x, p x (f x). 184 | Proof. 185 | split. 186 | - intros F. 187 | exists (fun x => pi1 (F x)). 188 | intros x. destruct (F x) as [y H]. cbn. exact H. 189 | - intros [f H] x. exists (f x). apply H. 190 | Qed. 191 | 192 | (*** Truncation *) 193 | 194 | Definition trunc X := forall Z:Prop, (X -> Z) -> Z. 195 | Notation "□ X" := (trunc X) (at level 30, right associativity). 196 | 197 | Goal forall P Q, P /\ Q <-> □ (P * Q). 198 | Proof. 199 | intros *; split. 200 | - intros H Z H1. tauto. 201 | - intros H. apply H. tauto. 202 | Qed. 203 | 204 | Goal forall P Q, P \/ Q <-> □ (P + Q). 205 | Proof. 206 | intros *; split. 207 | - intros H Z. tauto. 208 | - intros H. apply H. tauto. 209 | Qed. 210 | 211 | 212 | Goal forall X p, @ex X p <-> □ @sig X p. 213 | Proof. 214 | intros *; split. 215 | - intros [x H] Z H1. apply H1. eauto. 216 | - intros H. apply H. intros [x Hx]. eauto. 217 | Qed. 218 | 219 | Goal forall X, ~X <-> ~ □ X. 220 | Proof. 221 | intros *; split; intros H. 222 | - intros H2. apply H2. easy. 223 | - intros x. apply H. intros Z. auto. 224 | Qed. 225 | 226 | Inductive truncation X : Prop := Truncation (_: X). 227 | 228 | Goal forall X, truncation X <-> □ X. 229 | Proof. 230 | intros X. split. 231 | - intros [x] Z. auto. 232 | - intros H. apply H. intros x. constructor. exact x. 233 | Qed. 234 | 235 | (*** Exercises *) 236 | 237 | (** Certifying Distance *) 238 | Fact distance : 239 | forall x y, Sigma z, (x + z = y)%nat + (y + z = x)%nat. 240 | Proof. 241 | induction x as [|x IH]; cbn. 242 | - intros y. exists y. auto. 243 | - destruct y; cbn. 244 | + exists (S x). auto. 245 | + specialize (IH y) as [z [<-|<-]]; exists z; auto. 246 | Qed. 247 | 248 | Section Distance. 249 | Variable D: forall x y: nat, Sigma z, (x + z = y)%nat + (y + z = x)%nat. 250 | 251 | Fact D_sub x y : 252 | x - y = if pi2 (D x y) then 0 else pi1 (D x y). 253 | Proof. 254 | destruct (D x y) as [z [<-|<-]]; cbn; lia. 255 | Qed. 256 | 257 | Fact D_pi1 x y : 258 | pi1 (D x y) = (x - y) + (y - x). 259 | Proof. 260 | destruct (D x y) as [z [<-|<-]]; cbn; lia. 261 | Qed. 262 | 263 | Goal pi1 (D 3 7) = 4. 264 | Proof. 265 | rewrite D_pi1. reflexivity. 266 | Qed. 267 | End Distance. 268 | 269 | Fact skolem_prop_bool Y (p: bool -> Y -> Prop) : 270 | (forall x, exists y, p x y) <=> (exists f, forall x, p x (f x)). 271 | Proof. 272 | split. 273 | - intros F. 274 | destruct (F true) as [y1 H1]. 275 | destruct (F false) as [y2 H2]. 276 | exists (fun x: bool => if x then y1 else y2). 277 | destruct x; assumption. 278 | - intros [f H] x. exists (f x). apply H. 279 | Qed. 280 | 281 | Fact skolem_prop_prop X (Y: Prop) (p: X -> Y -> Prop) : 282 | (forall x, exists y, p x y) <=> (exists f, forall x, p x (f x)). 283 | Proof. 284 | split. 285 | - intros F. 286 | exists (fun x => let (y,_) := F x in y). 287 | intros x. destruct (F x) as [y H]. exact H. 288 | - intros [f H] x. exists (f x). apply H. 289 | Qed. 290 | 291 | Module Ex_eta. 292 | Section M. 293 | Variables (P: Prop) (p: P -> Prop). 294 | Definition pi1 (a: ex p) : P := 295 | match a with ex_intro _ x c => x end. 296 | Definition pi2 (a: ex p) : p (pi1 a) := 297 | match a with ex_intro _ x c => c end. 298 | Goal forall a, a = ex_intro p (pi1 a) (pi2 a). 299 | Proof. 300 | intros [x a]. reflexivity. 301 | Qed. 302 | End M. 303 | End Ex_eta. 304 | 305 | Goal forall x, Sigma a b, (x = 2 * a + b /\ b < 2)%nat. 306 | Proof. 307 | induction x as [|x (a&b&H)]. 308 | - exists 0, 0. lia. 309 | - destruct b. 310 | + exists a, 1. lia. 311 | + exists (S a), 0. lia. 312 | Qed. 313 | 314 | Goal 315 | (forall x, Sigma a b, (x = 2 * a + b /\ b < 2)%nat) <=> 316 | (forall x, Sigma a, (x = 2 * a)%nat + (x = S (2 * a))). 317 | Proof. 318 | split; intros F x. 319 | - specialize (F x) as (a&b&H). 320 | exists a. destruct b; intuition lia. 321 | - specialize (F x) as [a [H|H]]; exists a. 322 | + exists 0. lia. 323 | + exists 1. lia. 324 | Qed. 325 | 326 | Goal 327 | (forall x, exists a b, x = 2 * a + b /\ b < 2) <-> 328 | (forall x, exists a, x = 2 * a \/ x = S (2 * a)). 329 | Proof. 330 | split; intros F x. 331 | - specialize (F x) as (a&b&H). 332 | exists a. destruct b; intuition lia. 333 | - specialize (F x) as [a [H|H]]; exists a. 334 | + exists 0. lia. 335 | + exists 1. lia. 336 | Qed. 337 | 338 | -------------------------------------------------------------------------------- /rocq/certsum.v: -------------------------------------------------------------------------------- 1 | (*** MPCTT, Chapter Sum Types *) 2 | 3 | Notation "~ X" := (X -> False) (at level 75, right associativity) : type_scope. 4 | 5 | (*** Sum types *) 6 | 7 | Module SumTypes. 8 | Inductive sum (X Y: Type) : Type := L (x : X) | R (y : Y). 9 | 10 | Arguments L {X Y}. 11 | Arguments R {X Y}. 12 | 13 | Definition match_sum (X Y Z: Type) 14 | : sum X Y -> (X -> Z) -> (Y -> Z) -> Z 15 | := fun a e1 e2 => match a with L x => e1 x | R y => e2 y end. 16 | 17 | Definition elim_sum (X Y: Type) (p: sum X Y -> Type) 18 | : (forall x, p (L x)) -> (forall y, p (R y)) -> forall a, p a 19 | := fun e1 e2 a => match a with L x => e1 x | R y => e2 y end. 20 | 21 | Fact L_injective (X Y: Type) (x x': X) : 22 | @L X Y x = L x' -> x = x'. 23 | Proof. 24 | intros H. 25 | change (match_sum X Y X (L x) (fun x => x) (fun _ => x) = x'). 26 | rewrite H. cbn. 27 | reflexivity. 28 | Qed. 29 | 30 | Fact L_R_disjoint X Y (x:X) (y:Y) : 31 | L x <> R y. 32 | Proof. 33 | intros H. 34 | change (match_sum X Y Prop (L x) (fun x => False) (fun y => True)). 35 | rewrite H. cbn. 36 | exact I. 37 | Qed. 38 | End SumTypes. 39 | 40 | (* From now on we use Rocq's predefined sum types *) 41 | 42 | Locate "+". 43 | Print sum. 44 | Print unit. 45 | 46 | Goal forall a : False + unit, 47 | a = inr tt. 48 | Proof. 49 | destruct a as [x|y]. 50 | - destruct x. 51 | - destruct y. reflexivity. 52 | Qed. 53 | 54 | Goal forall a : (False + unit) + unit, 55 | a = inr tt \/ a = inl (inr tt). 56 | Proof. 57 | destruct a as [x|y]. 58 | - right. destruct x as [x|y]. 59 | + destruct x. 60 | + destruct y. reflexivity. 61 | - left. destruct y. reflexivity. 62 | Qed. 63 | 64 | Goal forall a : ((False + unit) + unit) + unit, 65 | a = inr tt \/ a = inl (inr tt) \/ a = inl (inl (inr tt)). 66 | Proof. 67 | destruct a as [x|y]. 68 | - right. destruct x as [x|y]. 69 | + right. destruct x as [x|y]. 70 | * destruct x. 71 | * destruct y. reflexivity. 72 | + left. destruct y. reflexivity. 73 | - left. destruct y. reflexivity. 74 | Qed. 75 | 76 | Section Exercise. 77 | Variables X Y : Type. 78 | Let T (a: X + Y) := if a then true else false. 79 | Goal forall a, if T a then exists x, a = inl x else exists y, a = inr y. 80 | Proof. 81 | intros [x|y]. 82 | - cbn. exists x. reflexivity. 83 | - cbn. exists y. reflexivity. 84 | Qed. 85 | End Exercise. 86 | 87 | (*** Proving at Type Level *) 88 | 89 | Definition iffT (X Y: Type) : Type := (X -> Y) * (Y -> X). 90 | Notation "X <=> Y" := (iffT X Y) (at level 95, no associativity). 91 | 92 | Section Exercise. 93 | Implicit Types X Y Z : Type. 94 | Goal forall X Y Z, 95 | (X + Y -> Z) <=> (X -> Z) * (Y -> Z). 96 | Proof. 97 | unfold iffT. tauto. 98 | Qed. 99 | 100 | Goal forall X Y Z, 101 | (X + Y -> Z) <=> (X -> Z) * (Y -> Z). 102 | Proof. 103 | split. 104 | - intros f. split. 105 | + intros x. apply f. left. exact x. 106 | + intros y. apply f. right. exact y. 107 | - intros [f g]. intros [x|y]. 108 | + apply f,x. 109 | + apply g,y. 110 | Qed. 111 | 112 | Goal forall X Y, 113 | X + Y <=> forall Z, (X -> Z ) -> (Y -> Z) -> Z. 114 | Proof. 115 | split. 116 | - intros [x|y] Z f g; auto. 117 | - intros F. apply F; auto. 118 | Qed. 119 | End Exercise. 120 | 121 | From Stdlib Require Import Bool. 122 | Module Exercise. 123 | Goal forall x y : bool, 124 | x && y = false <=> (x = false) + (y = false). 125 | Proof. 126 | destruct x, y; cbn; unfold iffT; tauto. 127 | Qed. 128 | End Exercise. 129 | 130 | Goal forall X Y, 131 | X + Y <=> forall Z, (X -> Z) -> (Y -> Z) -> Z. 132 | Proof. 133 | intros *. split. 134 | - intros [x|y]; auto. 135 | - intros F. apply F; auto. 136 | Qed. 137 | 138 | (*** Decision Types *) 139 | 140 | Definition dec (X: Type) : Type := X + (X -> False). 141 | 142 | Goal forall X Y, 143 | dec X -> dec Y -> dec (X + Y). 144 | Proof. 145 | unfold dec. tauto. 146 | Qed. 147 | 148 | Goal forall X Y, 149 | (X <=> Y) -> dec X -> dec Y. 150 | Proof. 151 | unfold dec, iffT. tauto. 152 | Qed. 153 | 154 | Goal forall X (f: X -> bool) x, 155 | dec (f x = true). 156 | Proof. 157 | intros *. destruct (f x) as [|]. 158 | - left. reflexivity. 159 | - right. easy. 160 | Qed. 161 | 162 | (*** Certifying Equality Deciders *) 163 | 164 | Goal forall x y: nat, (x = y) + (x <> y). 165 | Proof. 166 | induction x as [|x IH]; destruct y. 167 | - left. reflexivity. 168 | - right. easy. 169 | - right. easy. 170 | - destruct (IH y) as [H|H]. 171 | + left. congruence. 172 | + right. congruence. 173 | Qed. 174 | 175 | Definition eqdec X := forall x y: X, dec (x = y). 176 | 177 | Fact eqdec_bot : eqdec False. 178 | Proof. 179 | intros []. 180 | Qed. 181 | 182 | Fact nat_eqdec : eqdec nat. 183 | Proof. 184 | hnf. induction x as [|x IH]; destruct y. 185 | - left. reflexivity. 186 | - right. congruence. 187 | - right. congruence. 188 | - destruct (IH y) as [H|H]. 189 | + left. congruence. 190 | + right. congruence. 191 | Qed. 192 | 193 | Fact eqdec_prod X Y : 194 | eqdec X -> eqdec Y -> eqdec (X*Y). 195 | Proof. 196 | intros dX dY [x y] [x' y']. 197 | destruct (dX x x') as [H1|H1]. 198 | - destruct (dY y y') as [H2|H2]. 199 | + left. congruence. 200 | + right. congruence. 201 | - right. congruence. 202 | Qed. 203 | 204 | Fact eqdec_sum X Y : 205 | eqdec X * eqdec Y <=> eqdec (X + Y). 206 | Proof. 207 | split. 208 | - intros [dX dY] [x1|y1] [x2|y2]. 209 | + destruct (dX x1 x2); unfold dec in *; intuition congruence. 210 | + unfold dec in *; intuition congruence. 211 | + unfold dec in *; intuition congruence. 212 | + destruct (dY y1 y2); unfold dec in *; intuition congruence. 213 | - intros d; split. 214 | + intros x1 x2. 215 | destruct (d (inl x1) (inl x2)); unfold dec in *; intuition congruence. 216 | + intros y1 y2. 217 | destruct (d (inr y1) (inr y2)); unfold dec in *; intuition congruence. 218 | Qed. 219 | 220 | Definition injective {X Y} (f: X -> Y) := 221 | forall x x', f x = f x' -> x = x'. 222 | 223 | Fact eqdec_injective {X Y f} : 224 | @injective X Y f -> eqdec Y -> eqdec X. 225 | Proof. 226 | intros H d x x'. specialize (H x x'). 227 | destruct (d (f x) (f x')) as [H1|H1]; 228 | unfold dec in *; intuition congruence. 229 | Qed. 230 | 231 | Fixpoint nat_eqb (x y: nat) : bool := 232 | match x, y with 233 | | 0, 0 => true 234 | | 0, _ => false 235 | | S x, 0 => false 236 | | S x, S y => nat_eqb x y 237 | end. 238 | 239 | Fact nat_eqb_correct x y : 240 | if nat_eqb x y then x = y else x <> y. 241 | Proof. 242 | revert y. 243 | induction x as [|x IH]; destruct y; cbn. 244 | - reflexivity. 245 | - congruence. 246 | - congruence. 247 | - specialize (IH y). 248 | destruct (nat_eqb x y). 249 | + congruence. 250 | + congruence. 251 | Qed. 252 | 253 | 254 | 255 | 256 | -------------------------------------------------------------------------------- /rocq/demo-gs.v: -------------------------------------------------------------------------------- 1 | Check 3+4. 2 | Compute 3+4. 3 | Check Nat.add. 4 | 5 | Goal 3 + 4 = 10 - 3. 6 | Proof. 7 | cbn. reflexivity. 8 | Qed. 9 | 10 | Goal forall y, 0 + y = y. 11 | Proof. 12 | intros y. cbn. reflexivity. 13 | Qed. 14 | 15 | Goal forall x, x + 0 = x. 16 | Proof. 17 | intros x; cbn. 18 | easy. 19 | Qed. 20 | 21 | Goal forall x, x + 0 = x. 22 | Proof. 23 | induction x. 24 | - reflexivity. 25 | - cbn. rewrite IHx. reflexivity. 26 | Qed. 27 | 28 | Fact add0 x : 29 | x + 0 = x. 30 | Proof. 31 | induction x. 32 | - reflexivity. 33 | - cbn. rewrite IHx. reflexivity. 34 | Qed. 35 | 36 | Fact addS x y: 37 | x + S y = S (x + y). 38 | Proof. 39 | induction x. 40 | - reflexivity. 41 | - cbn. rewrite IHx. reflexivity. 42 | Qed. 43 | 44 | Fact add_comm x y: 45 | x + y = y + x. 46 | Proof. 47 | induction x; cbn. 48 | - rewrite add0. reflexivity. 49 | - rewrite IHx. rewrite addS. reflexivity. 50 | Qed. 51 | 52 | Goal forall x y, x + y = y + x. 53 | Proof. 54 | intros x. intros y. 55 | revert x. revert y. 56 | intros a b. apply add_comm. 57 | Qed. 58 | 59 | (** Distance, quantified inductive hypothesis *) 60 | 61 | Fixpoint D x y := 62 | match x, y with 63 | | 0, y => y 64 | | S x, 0 => S x 65 | | S x, S y => D x y 66 | end. 67 | 68 | Compute D 3 7. 69 | Compute D 7 3. 70 | 71 | Arguments D : simpl nomatch. 72 | 73 | Fact D_comm x y : 74 | D x y = D y x. 75 | Proof. 76 | induction x; cbn. 77 | - destruct y; cbn; reflexivity. 78 | - destruct y; cbn. 79 | + reflexivity. 80 | + 81 | Abort. 82 | 83 | Fact D_comm x y : 84 | D x y = D y x. 85 | Proof. 86 | revert y. 87 | induction x; cbn. 88 | - destruct y; cbn; reflexivity. 89 | - destruct y; cbn. 90 | + reflexivity. 91 | + apply IHx. 92 | Qed. 93 | 94 | Fact D_sub x y : 95 | D x y = (x - y) + (y - x). 96 | Proof. 97 | revert y. 98 | induction x as [|x IH]; intros y. 99 | - destruct y; reflexivity. 100 | - destruct y; cbn. 101 | + rewrite add0. reflexivity. 102 | + apply IH. 103 | Qed. 104 | 105 | (** Automatic prover for linear arithmetic *) 106 | 107 | From Stdlib Require Import Lia. 108 | 109 | Goal forall x y z, x - (y + z) = x - y - z. 110 | Proof. 111 | lia. 112 | Qed. 113 | 114 | 115 | (** Iteration *) 116 | 117 | Fixpoint iter (X: Type) (f: X -> X) (n:nat) (x:X) : X := 118 | match n with 119 | | 0 => x 120 | | S n => f (iter X f n x) 121 | end. 122 | 123 | Fact iter_add n x : 124 | n + x = iter nat S n x. 125 | Proof. 126 | induction n as [|n IH]; cbn. 127 | - reflexivity. 128 | - f_equal. exact IH. 129 | Qed. 130 | 131 | Fact iter_mul n x : 132 | n * x = iter nat (Nat.add x) n 0. 133 | Proof. 134 | induction n as [|n IH]; cbn. 135 | - reflexivity. 136 | - f_equal. exact IH. 137 | Qed. 138 | 139 | (* demo "_" *) 140 | 141 | (* goto gs.v for implicit arguments *) 142 | 143 | Inductive Pair (X Y: Type) : Type := 144 | | pair : X -> Y -> Pair X Y. 145 | 146 | Check Pair. 147 | Check pair. 148 | 149 | Definition swap X Y (a: Pair X Y) := 150 | match a with 151 | | pair _ _ x y => pair _ _ y x 152 | end. 153 | 154 | Goal forall X Y (a: Pair X Y), 155 | swap _ _ (swap _ _ a) = a. 156 | Proof. 157 | intros X Y. 158 | destruct a as [x y]. 159 | reflexivity. 160 | Qed. 161 | -------------------------------------------------------------------------------- /rocq/elim.v: -------------------------------------------------------------------------------- 1 | (*** MPCTT, Chapter 5 *) 2 | 3 | (* We define and apply inductive eliminators *) 4 | 5 | (*** Void and Unit *) 6 | 7 | Module VoidUnit. 8 | Inductive bot : Prop := . 9 | Inductive top : Prop := I. 10 | 11 | Definition elim_bot 12 | : forall Z: Type, bot -> Z 13 | := fun Z a => match a with end. 14 | 15 | Definition elim_top 16 | : forall p: top -> Type, p I -> forall a, p a 17 | := fun p e a => match a with I => e end. 18 | 19 | Goal bot <-> False. 20 | Proof. 21 | split. 22 | - apply elim_bot. 23 | - apply False_ind. 24 | Show Proof. 25 | Qed. 26 | 27 | Goal forall x: top, x = I. 28 | Proof. 29 | apply elim_top. 30 | reflexivity. 31 | Show Proof. 32 | Qed. 33 | 34 | Goal forall x: top, x = I. 35 | Proof. 36 | exact (fun x => match x with I => eq_refl I end). 37 | Qed. 38 | 39 | Inductive F: Prop := C (_: F). 40 | 41 | Definition elim_F 42 | : forall Z: Type, F -> Z 43 | := fun Z => fix f a := match a with C a => f a end. 44 | 45 | Goal F <-> bot. 46 | Proof. 47 | split. 48 | - apply elim_F. 49 | - apply elim_bot. 50 | Show Proof. 51 | Qed. 52 | End VoidUnit. 53 | 54 | (*** Bool *) 55 | 56 | Definition elim_bool 57 | : forall p: bool -> Type, p true -> p false -> forall x, p x 58 | := fun p e1 e2 x => match x with true => e1 | false => e2 end. 59 | 60 | Print elim_bool. 61 | (* Note: Rocq derives the return type function of the match *) 62 | 63 | Goal forall x, x = true \/ x = false. 64 | Proof. 65 | refine (elim_bool _ _ _). 66 | Show Proof. 67 | - left. reflexivity. 68 | - right. reflexivity. 69 | Qed. 70 | 71 | Check fun p: bool -> Prop => elim_bool p. 72 | Check fun p: bool -> Type => elim_bool p. 73 | Check fun p: bool -> Prop => p true. 74 | Check fun p: bool -> Type => p true. 75 | 76 | Goal forall x, x = true \/ x = false. 77 | Proof. 78 | intros [|]; auto. 79 | Show Proof. 80 | Qed. 81 | 82 | Goal forall (f: bool -> bool) x, f (f (f x)) = f x. 83 | Proof. 84 | intros f. 85 | enough (forall x y z, f true = y -> f false = z -> f (f (f x)) = f x) by eauto. 86 | refine (elim_bool _ _ _). 87 | all: refine (elim_bool _ _ _). 88 | all: refine (elim_bool _ _ _). 89 | all: congruence. 90 | Qed. 91 | 92 | Goal forall (f: bool -> bool) x, f (f (f x)) = f x. 93 | Proof. 94 | intros f. 95 | enough (forall x y z, f true = y -> f false = z -> f (f (f x)) = f x) by eauto. 96 | intros [|] [|] [|]; congruence. 97 | Qed. 98 | 99 | Goal forall (f: bool -> bool) x, f (f (f x)) = f x. 100 | Proof. 101 | destruct x; 102 | destruct (f true) eqn:H1; 103 | destruct (f false) eqn:H2. 104 | all: congruence. 105 | Qed. 106 | 107 | (*** Nat *) 108 | 109 | Definition elim_nat 110 | : forall p: nat -> Type, p 0 -> (forall n, p n -> p (S n)) -> forall n, p n 111 | := fun p e1 e2 => fix F n := 112 | match n with 0 => e1 | S n' => e2 n' (F n') end. 113 | 114 | Definition match_nat 115 | : forall p: nat -> Type, p 0 -> (forall n, p (S n)) -> forall n, p n 116 | := fun p e1 e2 n => 117 | match n with 0 => e1 | S n' => e2 n' end. 118 | 119 | Goal forall x, x + 0 = x. 120 | Proof. 121 | refine (elim_nat _ _ _). 122 | - reflexivity. 123 | - intros n IH. cbn. rewrite IH. reflexivity. 124 | Qed. 125 | 126 | Goal forall x y, 127 | x + y = elim_nat (fun _ => nat) y (fun _ => S) x. 128 | Proof. 129 | intros *. 130 | induction x as [|x IH]; cbn. 131 | - reflexivity. 132 | - f_equal. exact IH. 133 | Qed. 134 | 135 | Goal forall x y, 136 | x + y = elim_nat (fun _ => nat -> nat) (fun y => y) (fun _ a y => S (a y)) x y. 137 | Proof. 138 | intros *. 139 | induction x as [|x IH]; cbn. 140 | - reflexivity. 141 | - f_equal. exact IH. 142 | Qed. 143 | 144 | Fixpoint plus (x: nat) : nat -> nat := 145 | match x with 146 | | 0 => fun y => y 147 | | S x' => fun y => S (plus x' y) 148 | end. 149 | 150 | Goal plus = elim_nat (fun _ => nat -> nat) (fun y => y) (fun x a y => S (a y)). 151 | Proof. 152 | cbv. reflexivity. 153 | Qed. 154 | 155 | Goal forall x y: nat, x = y \/ x <> y. 156 | Proof. 157 | refine (elim_nat _ _ _). 158 | - refine (match_nat _ _ _). 159 | all: auto. (* auto includes discriminate *) 160 | - intros x IH. 161 | refine (match_nat _ _ _). 162 | + auto. 163 | + intros y. specialize (IH y). 164 | destruct IH; auto. (* auto includes injectivity *) 165 | Qed. 166 | 167 | Goal forall x y: nat, x = y \/ x <> y. 168 | Proof. 169 | induction x as [|x IH]; destruct y. 170 | 1-3: auto. 171 | specialize (IH y) as [IH|IH]. 172 | - left. congruence. 173 | - right. congruence. 174 | Qed. 175 | 176 | Fixpoint eq_nat (x y: nat) : bool := 177 | match x, y with 178 | | 0, 0 => true 179 | | 0, S _ => false 180 | | S _, 0 => false 181 | | S x, S y => eq_nat x y 182 | end. 183 | 184 | Fact eq_nat_correct x y : 185 | eq_nat x y = true <-> x = y. 186 | Proof. 187 | revert x y. 188 | refine (elim_nat _ _ _). 189 | - refine (elim_nat _ _ _). 190 | + cbn. easy. 191 | + intros y _. cbn. easy. 192 | - intros x IH. 193 | refine (elim_nat _ _ _). 194 | + cbn. easy. 195 | + intros y _. cbn. specialize (IH y). split. 196 | * intros H. f_equal. apply IH, H. 197 | * intros H. apply IH. congruence. 198 | Qed. 199 | 200 | Goal forall x y: nat, x = y \/ x <> y. 201 | Proof. 202 | intros x y. 203 | assert (H:= eq_nat_correct x y). 204 | destruct (eq_nat x y). 205 | - left. apply H. easy. 206 | - right. intros H1. 207 | enough (false = true) by easy. 208 | apply H, H1. 209 | Qed. 210 | 211 | Module Exercise. 212 | Notation "x <= y" := (x - y = 0) : nat_scope. 213 | Fact antisymmetry x y : 214 | x <= y -> y <= x -> x = y. 215 | Proof. 216 | revert y. 217 | induction x as [|x IH]; destruct y. 218 | all:auto. 219 | Qed. 220 | End Exercise. 221 | 222 | (*** Pairs *) 223 | 224 | Definition elim_pair 225 | : forall (X Y: Type) (p: X * Y -> Type), (forall x y, p(x,y)) -> forall a, p a 226 | := fun X Y p e => fix F a := 227 | match a with (x,y) => e x y end. 228 | 229 | Definition fst 230 | : forall X Y, X * Y -> X 231 | := fun X Y => elim_pair X Y _ (fun x _ => x). 232 | 233 | Definition snd 234 | : forall X Y, X * Y -> Y 235 | := fun X Y => elim_pair X Y _ (fun _ y => y). 236 | 237 | Goal forall X Y (a: X * Y), a = (fst _ _ a, snd _ _ a). 238 | Proof. 239 | intros X Y. 240 | apply elim_pair. 241 | cbn. 242 | reflexivity. 243 | Qed. 244 | 245 | (*** [nat <> bool] *) 246 | 247 | Goal nat <> bool. 248 | Proof. 249 | pose (p X := forall x y z : X, x = y \/ x = z \/ y = z). 250 | enough (p bool /\ ~ p nat) as [H1 H2]. 251 | - intros H. apply H2. rewrite H. exact H1. 252 | - split; unfold p. 253 | + intros [|] [|] [|]; auto. 254 | + intros H. specialize (H 0 1 2) as [H|[H|H]]; congruence. 255 | Qed. 256 | 257 | (*** Rocq's Set considered harmful *) 258 | 259 | (* Rocq has a subuniverse Set of Type and type inference uses 260 | Set rather than type if it can. In particular, the predefined 261 | types bool and nat are typed with Set. This can lead to 262 | annoying problems. An example follows. *) 263 | 264 | Check nat. 265 | Check bool. 266 | Set Printing All. 267 | Check nat <> bool. 268 | 269 | Lemma eq_not (X: Type) (x y : X) (p: X -> Prop) : 270 | ~ p x -> p y -> x <> y. 271 | Proof. 272 | intros H H1. contradict H. rewrite H. exact H1. 273 | Qed. 274 | 275 | Definition card_le2 (X: Type) := 276 | forall x y z : X, x = y \/ x = z \/ y = z. 277 | 278 | Goal nat <> bool. 279 | Proof. 280 | Fail apply (eq_not Type nat bool card_le2). 281 | enough (not (@eq Type nat bool)) as H. 282 | - contradict H. rewrite H. reflexivity. 283 | - apply (eq_not Type _ _ card_le2). 284 | + intros H. specialize (H 0 1 2) as [H|[H|H]]; congruence. 285 | + intros [|] [|] [|]; auto. 286 | Qed. 287 | Unset Printing All. 288 | 289 | (*** Rocq derives eliminators *) 290 | Check bool_rect. 291 | Check nat_rect. 292 | Check prod_rect. 293 | Check False_rect. 294 | Check True_rect. 295 | (* Rocq doesn't derive most general eliminator for True by default *) 296 | Check and_ind. 297 | Check and_rect. 298 | Check or_ind. 299 | 300 | Check False_rect nat. 301 | Check False_rect (nat -> nat). 302 | Check False_rect (nat -> nat -> nat). 303 | 304 | -------------------------------------------------------------------------------- /rocq/ewo.v: -------------------------------------------------------------------------------- 1 | (*** MPCTT, Chapter EWOs *) 2 | From Stdlib Require Import Lia. 3 | Definition iffT (X Y: Type) : Type := (X -> Y) * (Y -> X). 4 | Notation "X <=> Y" := (iffT X Y) (at level 95, no associativity). 5 | Definition dec (X: Type) : Type := X + (X -> False). 6 | Definition eqdec X := forall x y: X, dec (x = y). 7 | Definition decider {X} (p: X -> Type) := forall x, dec (p x). 8 | Notation sig := sigT. 9 | Notation Sig := existT. 10 | Notation pi1 := projT1. 11 | Notation pi2 := projT2. 12 | Notation "'Sigma' x .. y , p" := 13 | (sigT (fun x => .. (sigT (fun y => p%type)) ..)) 14 | (at level 200, x binder, right associativity, 15 | format "'[' 'Sigma' '/ ' x .. y , '/ ' p ']'") 16 | : type_scope. 17 | Definition inv {X Y: Type} (g: Y -> X) (f: X -> Y) := forall x, g (f x) = x. 18 | Inductive injection (X Y: Type) : Type := 19 | | Injection {f: X -> Y} {g: Y -> X} (_: inv g f). 20 | Definition injective {X Y} (f: X -> Y) := 21 | forall x x', f x = f x' -> x = x'. 22 | Definition surjective {X Y} (f: X -> Y) := 23 | forall y, exists x, f x = y. 24 | Definition bijective {X Y} (f: X -> Y) := 25 | injective f /\ surjective f. 26 | 27 | 28 | (*** Linear Search Types *) 29 | 30 | Section EWO. 31 | Variable p: nat -> Prop. 32 | 33 | Inductive T (n: nat) : Prop := C (phi: ~p n -> T (S n)). 34 | 35 | Definition T_elim (q: nat -> Type) 36 | : (forall n, (~p n -> q (S n)) -> q n) -> 37 | forall n, T n -> q n 38 | := fun e => fix f n a := 39 | let (phi) := a in 40 | e n (fun h => f (S n) (phi h)). 41 | 42 | (*** EWO for Numbers *) 43 | 44 | Lemma TI n : 45 | p n -> T n. 46 | Proof. 47 | intros H. constructor. intros H1. destruct (H1 H). 48 | Qed. 49 | 50 | Lemma TD n : 51 | T (S n) -> T n. 52 | Proof. 53 | intros H. constructor. intros _. exact H. 54 | Qed. 55 | 56 | Variable p_dec: decider p. 57 | 58 | Definition TE (q: nat -> Type) 59 | : (forall n, p n -> q n) -> 60 | (forall n, ~p n -> q (S n) -> q n) -> 61 | forall n, T n -> q n. 62 | Proof. 63 | intros e1 e2. 64 | apply T_elim. intros n IH. 65 | destruct (p_dec n); auto. 66 | Qed. 67 | 68 | (** From now on T will only be used through TI, TD, and TE *) 69 | 70 | 71 | Lemma T_zero n : 72 | T n -> T 0. 73 | Proof. 74 | induction n as [|n IH]. 75 | - auto. 76 | - intros H. apply IH. apply TD, H. 77 | Qed. 78 | 79 | Definition ewo_nat : 80 | ex p -> sig p. 81 | Proof. 82 | intros H. 83 | refine (TE (fun _ => sig p) _ _ 0 _). 84 | - eauto. 85 | - auto. 86 | - destruct H as [n H]. apply (T_zero n), TI, H. 87 | Qed. 88 | 89 | (* T_zero generalized *) 90 | 91 | Fact T_lower n m : 92 | m <= n -> T n -> T m. 93 | Proof. 94 | induction n as [|n IH]. 95 | - intros H1 H2. 96 | assert (m = 0) as -> by lia. 97 | exact H2. 98 | - intros H1 H2. 99 | assert (m = S n \/ m <= n) as [H|H] by lia. 100 | + congruence. 101 | + apply IH. exact H. apply TD, H2. 102 | Qed. 103 | 104 | Fact T_sig n : 105 | T n -> Sigma m, m >= n /\ p m. 106 | Proof. 107 | revert n. 108 | refine (TE _ _ _). 109 | - intros n IH. exists n. split. lia. exact IH. 110 | - intros n _ (m&IH1&IH2). exists m. split. lia. exact IH2. 111 | Qed. 112 | 113 | Fact T_equiv n : 114 | T n <=> Sigma m, m >= n /\ p m. 115 | Proof. 116 | split. 117 | - apply T_sig. 118 | - intros (m&H1&H2). 119 | eapply T_lower. exact H1. apply TI, H2. 120 | Qed. 121 | End EWO. 122 | 123 | (*** General EWOs *) 124 | 125 | Definition ewo (X:Type) := 126 | forall p: X -> Prop, decider p -> ex p -> sig p. 127 | 128 | Fact bot_ewo: 129 | ewo False. 130 | Proof. 131 | intros p _ [[] _]. 132 | (* note computational falsity elimination *) 133 | Qed. 134 | 135 | Goal ewo True. 136 | Proof. 137 | intros p d H. 138 | destruct (d I) as [H1|H1]. 139 | - eauto. 140 | - exfalso. destruct H as [[] H]. auto. 141 | (* note computational falsity elimination *) 142 | Qed. 143 | 144 | Goal ewo bool. 145 | Proof. 146 | intros p d H. 147 | destruct (d true) as [H1|H1]. 148 | - eauto. 149 | - destruct (d false) as [H2|H2]. 150 | + eauto. 151 | + exfalso. destruct H as [[|] H]; auto. 152 | (* note computational falsity elimination *) 153 | Qed. 154 | 155 | Goal ewo nat. 156 | Proof. 157 | exact ewo_nat. 158 | Qed. 159 | 160 | Theorem ewo_or X (p q: X -> Prop) : 161 | ewo X -> decider p -> decider q -> ex p \/ ex q -> sig p + sig q. 162 | Proof. 163 | intros E dp dq H. 164 | destruct (E (fun x => p x \/ q x)) as [x H1]. 165 | - intros x. unfold dec. 166 | destruct (dp x) as [H1|H1]. { auto. } 167 | destruct (dq x) as [H2|H2]. { auto. } 168 | tauto. 169 | - destruct H as [[x H]|[x H]]; eauto. 170 | - destruct (dp x) as [H2|H2]. { eauto. } 171 | destruct (dq x) as [H3|H3]. { eauto. } 172 | exfalso. tauto. 173 | Qed. 174 | 175 | Definition option_ewo {X} : 176 | ewo X -> ewo (option X). 177 | Proof. 178 | intros E p p_dec H. 179 | destruct (p_dec None) as [H1|H1]. 180 | - eauto. 181 | - destruct (E (fun x => p (Some x))) as [x H2]. 182 | + easy. 183 | + destruct H as [[x|] H]. 184 | * eauto. 185 | * easy. 186 | + eauto. 187 | Qed. 188 | 189 | Definition option_ewo' {X} : 190 | ewo (option X) -> ewo X. 191 | Proof. 192 | intros E p p_dec H. 193 | destruct (E (fun a => match a with Some x => p x | none => False end)) as [[x|] H1]. 194 | - intros [x|]. 195 | + easy. 196 | + right; easy. 197 | - destruct H as [x H]. exists (Some x); exact H. 198 | - eauto. 199 | - easy. 200 | Qed. 201 | 202 | Fixpoint Fin n : Type := 203 | match n with 0 => False | S n' => option (Fin n') end. 204 | 205 | Fact Fin_ewo : 206 | forall n, ewo (Fin n). 207 | Proof. 208 | induction n as [|n IH]; cbn. 209 | - apply bot_ewo. 210 | - apply option_ewo, IH. 211 | Qed. 212 | 213 | Fact injection_ewo X Y : 214 | injection X Y -> ewo Y -> ewo X. 215 | Proof. 216 | intros [f g H] E p p_dec H1. 217 | destruct (E (fun y => p (g y))) as [y H2]. 218 | - easy. 219 | - destruct H1 as [x H1]. exists (f x). congruence. 220 | - eauto. 221 | Qed. 222 | 223 | Fact injection_nat_ewo X : 224 | injection X nat -> ewo X. 225 | Proof. 226 | intros H. 227 | eapply injection_ewo. exact H. exact ewo_nat. 228 | Qed. 229 | 230 | Fact ewo_binary : 231 | injection (nat * nat) nat -> 232 | forall p: nat -> nat -> Prop, 233 | forall d: forall x y, dec (p x y), 234 | (exists x y, p x y) -> Sigma x y, p x y. 235 | Proof. 236 | intros E %injection_ewo. 2:exact ewo_nat. 237 | intros p d H. 238 | pose (q a := p (fst a) (snd a)). 239 | specialize (E (fun a => p (fst a) (snd a))) as [[x y] E]. 240 | - intros [x y]. apply d. 241 | - destruct H as (x&y&H). exists (x,y). exact H. 242 | - eauto. 243 | Qed. 244 | 245 | (*** EWO Applications *) 246 | 247 | Fact surjective_inv {X Y f} : 248 | @surjective X Y f -> ewo X -> eqdec Y -> Sigma g, inv f g. 249 | Proof. 250 | intros H e d. 251 | enough (G: forall y, Sigma x, f x = y). 252 | { exists (fun y => pi1 (G y)). intros y. apply (pi2 (G y)). } 253 | intros y. apply e. 254 | - intros x. apply d. 255 | - apply H. 256 | Qed. 257 | 258 | Fact bijective_inv X Y f : 259 | @bijective X Y f -> ewo X -> eqdec Y -> Sigma g, inv g f /\ inv f g. 260 | Proof. 261 | intros [H1 H2] e d. 262 | destruct (surjective_inv H2 e d) as [g H3]. 263 | exists g. split. 2:exact H3. 264 | intros x. apply H1. congruence. 265 | Qed. 266 | 267 | Section Step_indexed_eqdec. 268 | Variable X: Type. 269 | Variable f: X -> X -> nat -> bool. 270 | Variable f_prop: forall x y, x = y <-> exists n, f x y n = true. 271 | Goal eqdec X. 272 | Proof. 273 | intros x y. 274 | enough (Sigma n, f x x n = true) as [n H]. 275 | { destruct (f x y n) eqn:H1. 276 | - left. apply f_prop. exists n. exact H1. 277 | - right. intros <-. congruence. } 278 | apply ewo_nat. 279 | - intros n. 280 | destruct (f x x n). 281 | + left. auto. 282 | + right. intros [=]. 283 | - apply f_prop. reflexivity. 284 | Qed. 285 | End Step_indexed_eqdec. 286 | 287 | (*** EWO with interface *) 288 | 289 | Definition safe p n := forall k, k < n -> ~p k. 290 | Definition least p n := p n /\ safe p n. 291 | 292 | Fact safe_S p n : 293 | safe p n -> ~p n -> safe p (S n). 294 | Proof. 295 | unfold safe; intros H1 H2 k H3. 296 | assert (k < n \/ k = n) as [H| ->] by lia; auto. 297 | Qed. 298 | 299 | Module EWO_Nat_Inter. 300 | 301 | Section EWO_nat_above. 302 | Variable p : nat -> Prop. 303 | Variable T' : nat -> Prop. 304 | Variable I : forall n, p n -> T' n. 305 | Variable D : forall n, T' (S n) -> T' n. 306 | Variable E' : 307 | decider p -> 308 | forall q: nat -> Type, 309 | (forall n, p n -> q n) -> 310 | (forall n, q (S n) -> q n) -> 311 | (forall n, T' n -> q n). 312 | 313 | Lemma T'_sig : 314 | decider p -> forall n, T' n -> sig p. 315 | Proof. 316 | intros d. apply (E' d); eauto. 317 | Qed. 318 | 319 | Lemma T'_zero : 320 | forall n, T' n -> T' 0. 321 | Proof. 322 | induction n as [|n IH]. easy. 323 | intros H. apply IH, D, H. 324 | Qed. 325 | 326 | Lemma ex_T'_zero : 327 | ex p -> T' 0. 328 | Proof. 329 | intros [x H]. eapply T'_zero, I, H. 330 | Qed. 331 | 332 | Fact ewo_nat : 333 | decider p -> ex p -> sig p. 334 | Proof. 335 | intros d H. apply (T'_sig d 0). 336 | destruct H as [n H]. 337 | apply (T'_zero n), I, H. 338 | Qed. 339 | 340 | Lemma T'_least : 341 | decider p -> forall n, T' n -> safe p n -> sig (least p). 342 | Proof. 343 | intros d. 344 | refine (E' d _ _ _). 345 | - intros n H1 H2. exists n. easy. 346 | - intros n H1 H2. 347 | destruct (d n) as [H3|H3]. 348 | + exists n. easy. 349 | + apply H1, safe_S; easy. 350 | Qed. 351 | 352 | Fact ewo_nat_least : 353 | decider p -> ex p -> sig (least p). 354 | Proof. 355 | intros d H. 356 | apply (T'_least d 0). 2:easy. 357 | apply ex_T'_zero, H. 358 | Qed. 359 | 360 | End EWO_nat_above. 361 | 362 | Section EWO_nat_below. 363 | Variable p : nat -> Prop. 364 | 365 | Definition I : forall n, p n -> T p n. 366 | Proof. intros n H. apply C. easy. Qed. 367 | 368 | Definition D : forall n, T p (S n) -> T p n. 369 | Proof. intros n H. apply C. easy. Qed. 370 | 371 | Fact E': 372 | decider p -> 373 | forall q: nat -> Type, 374 | (forall n, p n -> q n) -> 375 | (forall n, q (S n) -> q n) -> 376 | (forall n, T p n -> q n). 377 | Proof. 378 | intros d q e1 e2. 379 | apply T_elim. 380 | intros n IH. 381 | destruct (d n) as [H|H]; auto. 382 | Qed. 383 | End EWO_nat_below. 384 | End EWO_Nat_Inter. 385 | 386 | -------------------------------------------------------------------------------- /rocq/exquant.v: -------------------------------------------------------------------------------- 1 | (*** MPCTT, Chapter existential quantification *) 2 | 3 | Module Demo. 4 | Inductive ex {X: Type} (p: X -> Prop) : Prop := E (x:X) (a: p x). 5 | 6 | (* X implicit for ex and E *) 7 | 8 | Definition match_ex {X: Type} (p: X -> Prop) (Z: Prop) 9 | : ex p -> (forall x, p x -> Z) -> Z 10 | := fun a e => match a with E _ x b => e x b end. 11 | 12 | Lemma deMorgan X (p: X -> Prop) : 13 | ~ ex (fun x => p x) <-> forall x, ~ p x. 14 | Proof. 15 | split. 16 | - intros f x a. 17 | apply f. 18 | exact (E p x a). (* note eta conversion *) 19 | - intros f a. 20 | apply (match_ex p False a). 21 | exact f. 22 | Show Proof. 23 | Qed. 24 | End Demo. 25 | 26 | Locate "exists". 27 | Print ex. 28 | 29 | Lemma deMorgan X (p: X -> Prop) : 30 | ~ (exists x, p x) <-> forall x, ~ p x. 31 | Proof. 32 | split. 33 | - intros f x a. apply f. exists x. exact a. 34 | - intros f [x a]. exact (f x a). 35 | Show Proof. 36 | Qed. 37 | 38 | Goal forall X (p: X -> Prop), 39 | ~ (exists x, p x) <-> forall x, ~ p x. 40 | Proof. 41 | refine (fun X p => conj (fun f x a => _) (fun f b => _)). 42 | - refine (f (ex_intro p x a)). 43 | - refine (match b with ex_intro _ x a => f x a end). 44 | Show Proof. 45 | Qed. 46 | 47 | Theorem Barber X (p: X -> X -> Prop) : 48 | ~ (exists x, forall y, p x y <-> ~ p y y). 49 | Proof. 50 | intros [x H]. specialize (H x). tauto. 51 | Qed. 52 | 53 | (** Lawvere *) 54 | 55 | Fact negb_no_fp : 56 | ~ exists x, negb x = x. 57 | Proof. 58 | intros [[|] H]; discriminate. 59 | Qed. 60 | 61 | Fact not_no_fp : 62 | ~ exists P: Prop, (~P) = P. 63 | Proof. 64 | intros [P H]. 65 | enough (H1: ~(P <-> P)). 66 | - tauto. 67 | - pattern P at 2. rewrite <-H. tauto. 68 | Qed. 69 | 70 | Definition surjective {X Y} (f: X -> Y) := 71 | forall y, exists x, f x = y. 72 | 73 | Theorem Lawvere X Y (f: X -> X -> Y) (g: Y -> Y) : 74 | surjective f -> exists y, g y = y. 75 | Proof. 76 | intros H. 77 | specialize (H (fun x => g (f x x))) as [x H]. 78 | apply (f_equal (fun f => f x)) in H. 79 | exists (f x x). 80 | easy. 81 | Qed. 82 | 83 | Corollary Lawvere_bool X : 84 | ~ exists f: X -> X -> bool, surjective f. 85 | Proof. 86 | intros [f H]. 87 | apply negb_no_fp. 88 | revert H. apply Lawvere. 89 | Qed. 90 | 91 | Corollary Lawvere_Prop X : 92 | ~ exists f: X -> X -> Prop, surjective f. 93 | Proof. 94 | intros [f H]. 95 | apply not_no_fp. 96 | revert H. apply Lawvere. 97 | Qed. 98 | 99 | (** Exercise: Equational proof of not_no_fp, tricky *) 100 | 101 | Corollary not_no_fp' X : 102 | (~X) <> X. 103 | Proof. 104 | intros H. 105 | pose (id:= fun a: False => a). 106 | enough (exists a, id a = a) as [[] _]. 107 | enough (exists f: X -> X -> False, surjective f) as [f H1]. 108 | - revert H1. apply Lawvere. 109 | - pattern (X -> False). rewrite H. 110 | exists (fun x => x). intros x. exists x. reflexivity. 111 | Qed. 112 | -------------------------------------------------------------------------------- /rocq/ext.v: -------------------------------------------------------------------------------- 1 | Definition inv {X Y: Type} (g: Y -> X) (f: X -> Y) := forall x, g (f x) = x. 2 | Inductive bijection (X Y: Type) : Type := 3 | | Bijection: forall (f: X -> Y) (g: Y -> X), inv g f -> inv f g -> bijection X Y. 4 | Arguments Bijection {X Y}. 5 | Definition dec (X: Type) := sum X (X -> False). 6 | Definition eqdec X := forall x y: X, dec (x = y). 7 | 8 | 9 | Definition FE : Prop := 10 | forall X (p: X -> Type) (f g: forall x, p x), (forall x, f x = g x) -> f = g. 11 | Definition PE : Prop := 12 | forall P Q: Prop, P <-> Q -> P = Q. 13 | Definition unique (X: Type) : Prop := 14 | forall a b: X, a = b. 15 | Definition PI : Prop := 16 | forall P: Prop, unique P. 17 | 18 | Fact unique_bot : 19 | unique False. 20 | Proof. 21 | intros []. 22 | Qed. 23 | 24 | Fact unique_top : 25 | unique True. 26 | Proof. 27 | intros [] []. reflexivity. 28 | Qed. 29 | 30 | Fact PE_PI : 31 | PE -> PI. 32 | Proof. 33 | intros H P a. 34 | enough (unique P) as H1. 35 | { apply H1. } 36 | enough (P = True) as H1. 37 | { rewrite H1. apply unique_top. } 38 | apply H. tauto. 39 | Qed. 40 | 41 | Fact bool_bool_enum : 42 | FE -> forall f, 43 | f = (fun _ => true) \/ 44 | f = (fun _ => false) \/ 45 | f = (fun b => b) \/ 46 | f = negb. 47 | Proof. 48 | intros H f. 49 | assert (forall g, f true = g true -> f false = g false -> f = g). 50 | - intros g ? ?. apply H. intros []; easy. 51 | - destruct (f true) eqn:?, (f false) eqn:?; auto. 52 | Qed. 53 | 54 | Goal 55 | FE -> bijection (True -> True) True. 56 | Proof. 57 | intros H. 58 | exists (fun _ => I) (fun _ x => x). 59 | - intros f. apply H. intros x. apply unique_top. 60 | - intros x. apply unique_top. 61 | Qed. 62 | 63 | Goal 64 | FE -> bijection (True -> bool) bool. 65 | Proof. 66 | intros H. 67 | exists (fun f => f I) (fun b x => b). 68 | - intros f. apply H. intros []. reflexivity. 69 | - intros b. reflexivity. 70 | Qed. 71 | 72 | Goal 73 | FE -> bijection (bool -> bool) (bool * bool). 74 | Proof. 75 | intros H. 76 | exists (fun f => (f true, f false)) 77 | (fun a => fun b: bool => if b then fst a else snd a) 78 | ; hnf. 79 | - intros f. apply H. cbn. intros [|]; reflexivity. 80 | - intros [x y]. reflexivity. 81 | Qed. 82 | 83 | Goal forall X, unique X -> eqdec X. 84 | Proof. 85 | intros X H x y. left. apply H. 86 | Qed. 87 | 88 | Goal forall X, bijection X True -> unique X. 89 | Proof. 90 | intros X [f g H1 H2] x y. 91 | rewrite <-(H1 x), <-(H1 y). f_equal. 92 | apply unique_top. 93 | Qed. 94 | 95 | Goal forall X, X -> unique X -> bijection X True. 96 | Proof. 97 | intros X x. 98 | exists (fun _ => I) (fun _ => x). 99 | - intros y. apply H. 100 | - intros a. apply unique_top. 101 | Qed. 102 | 103 | Fact unique_Rocquand (A: Prop) (E: Prop -> A) (D: A -> Prop) : 104 | unique A -> (forall P: Prop, D (E P) <-> P) -> False. 105 | Proof. 106 | intros HA H. 107 | apply H. rewrite (HA (E False) (E True)). apply H, I. 108 | Qed. 109 | 110 | Section Elim_restriction_or. 111 | Variables (f: True \/ True -> bool) 112 | (feqT: f (or_introl I) = true) 113 | (feqF: f (or_intror I) = false). 114 | Goal PI -> False. 115 | Proof. 116 | intros H. 117 | assert (f (or_introl I) = f (or_intror I)) 118 | as H1 by apply f_equal, H. 119 | revert H1. rewrite feqT, feqF. discriminate. 120 | Qed. 121 | End Elim_restriction_or. 122 | 123 | Section Elim_restriction_ex. 124 | Variables (f: @ex bool (fun _ => True) -> bool) 125 | (feq: forall x, f (ex_intro _ x I) = x). 126 | 127 | Goal PI -> False. 128 | Proof. 129 | intros H. 130 | assert (f (ex_intro _ true I) = f (ex_intro _ false I)) 131 | as H1 by apply f_equal, H. 132 | revert H1. rewrite !feq. discriminate. 133 | Qed. 134 | End Elim_restriction_ex. 135 | -------------------------------------------------------------------------------- /rocq/gs.v: -------------------------------------------------------------------------------- 1 | (*** MPCTT, Chapter 1 *) 2 | 3 | (* Step through the commands and proofs 4 | and try to understand what is happening. 5 | The main point is to understand 6 | how theorems are stated and proved. *) 7 | 8 | (** Booleans *) 9 | 10 | Check negb. 11 | Compute negb true. 12 | 13 | Fact negb_involutive b : 14 | negb (negb b) = b. 15 | Proof. 16 | destruct b. 17 | - reflexivity. 18 | - reflexivity. 19 | Qed. 20 | 21 | (** Commutativity of addition *) 22 | 23 | Compute 4 + 3. 24 | Check 4 + 3. 25 | 26 | Fact add0 x : 27 | x + 0 = x. 28 | Proof. 29 | induction x as [|x IH]. 30 | - reflexivity. 31 | - cbn. f_equal. exact IH. 32 | Qed. 33 | 34 | Fact addS x y: 35 | x + S y = S (x + y). 36 | Proof. 37 | induction x as [|x IH]. 38 | - reflexivity. 39 | - cbn. f_equal. exact IH. 40 | Qed. 41 | 42 | Fact add_comm x y: 43 | x + y = y + x. 44 | Proof. 45 | induction x as [|x IH]; cbn. 46 | - symmetry. apply add0. 47 | - rewrite addS. f_equal. exact IH. 48 | Qed. 49 | 50 | (** Subtraction *) 51 | 52 | Compute 7 - 3. 53 | Compute 3 - 7. 54 | 55 | Fact sub_xx x : 56 | x - x = 0. 57 | Proof. 58 | induction x as [|x IH]; cbn. 59 | - reflexivity. 60 | - exact IH. 61 | Qed. 62 | 63 | Fact add_sub x y : 64 | (x + y) - x = y. 65 | Proof. 66 | induction x as [|x IH]; cbn. 67 | - destruct y; reflexivity. 68 | - exact IH. 69 | Qed. 70 | 71 | Fact add_sub' x y : 72 | (x + y) - y = x. 73 | Proof. 74 | rewrite add_comm. apply add_sub. 75 | Qed. 76 | 77 | (** Distance, quantified inductive hypothesis *) 78 | 79 | Fixpoint D x y := 80 | match x, y with 81 | | 0, 0 => 0 82 | | 0, S y => S y 83 | | S x, 0 => S x 84 | | S x, S y => D x y 85 | end. 86 | 87 | (* Ensure that cbn does not show matches *) 88 | Arguments D : simpl nomatch. 89 | 90 | Compute D 7 3. 91 | Compute D 3 7. 92 | 93 | Fact D_comm x y : 94 | D x y = D y x. 95 | Proof. 96 | revert y. 97 | induction x as [|x IH]; intros y. 98 | - destruct y; reflexivity. 99 | - destruct y; cbn. 100 | + reflexivity. 101 | + apply IH. 102 | Qed. 103 | 104 | Fact D_sub x y : 105 | D x y = (x - y) + (y - x). 106 | Proof. 107 | revert y. 108 | induction x as [|x IH]; intros y. 109 | - destruct y; reflexivity. 110 | - destruct y; cbn. 111 | + rewrite add0. reflexivity. 112 | + apply IH. 113 | Qed. 114 | 115 | (** Pairs *) 116 | 117 | Definition swap {X Y} (a: X * Y) : Y * X := 118 | match a with (x, y) => (y, x) end. 119 | 120 | Compute swap (3,5). 121 | 122 | Fact swap_swap X Y (a: X * Y) : 123 | swap (swap a) = a. 124 | Proof. 125 | destruct a as [x y]. reflexivity. 126 | Qed. 127 | 128 | Fact pair_eta X Y (a: X * Y) : 129 | (fst a, snd a) = a. 130 | Proof. 131 | destruct a as [x y]. reflexivity. 132 | Qed. 133 | 134 | (* Printing of implicit arguments can be forced *) 135 | Set Printing Implicit. 136 | Check swap (3,5). 137 | Check swap_swap. 138 | Check pair_eta. 139 | Unset Printing Implicit. 140 | 141 | (* Everything can be printed *) 142 | Set Printing All. 143 | Check 1+2. 144 | Check swap (1,2). 145 | Check pair_eta. 146 | Check nat -> nat. 147 | Unset Printing All. 148 | 149 | (* Printing of defined notations can be swiched off *) 150 | Unset Printing Notations. 151 | Check 2+3. 152 | Check swap (1,2). 153 | Check pair_eta. 154 | Set Printing Notations. 155 | 156 | (* ADVICE: Rocq comes with lots of notational conveniences, 157 | including infix operators, argument inference, and implicit arguments. 158 | This can be confusing. It is important to understand 159 | what a phrase elaborates to once all notational conveniences are removed. *) 160 | 161 | (** Iteration *) 162 | 163 | Fixpoint iter {X: Type} (f: X -> X) (n:nat) (x:X) : X := 164 | match n with 165 | | 0 => x 166 | | S n => f (iter f n x) 167 | end. 168 | 169 | Fact iter_add n x : 170 | n + x = iter S n x. 171 | Proof. 172 | induction n as [|n IH]; cbn. 173 | - reflexivity. 174 | - f_equal. exact IH. 175 | Qed. 176 | 177 | Fact iter_mul n x : 178 | n * x = iter (Nat.add x) n 0. 179 | Proof. 180 | induction n as [|n IH]; cbn. 181 | - reflexivity. 182 | - f_equal. exact IH. 183 | Qed. 184 | 185 | Fact iter_pow n x : 186 | Nat.pow x n = iter (Nat.mul x) n 1. 187 | Proof. 188 | induction n as [|n IH]; cbn. 189 | - reflexivity. 190 | - f_equal. exact IH. 191 | Qed. 192 | 193 | Fact iter_shift X (f: X -> X) n x : 194 | iter f (S n) x = iter f n (f x). 195 | Proof. 196 | induction n as [|n IH]. 197 | - cbn. reflexivity. 198 | - cbn. f_equal. exact IH. 199 | Qed. 200 | 201 | Fact iter_shift' X (f: X -> X) n x : 202 | f (iter f n x) = iter f n (f x). 203 | Proof. 204 | rewrite <-iter_shift. reflexivity. 205 | Qed. 206 | 207 | (* Warning: Predefined iter has different argument order *) 208 | Check Nat.iter. 209 | 210 | (** Factorial *) 211 | 212 | Fixpoint fac (n: nat) : nat := 213 | match n with 214 | | 0 => 1 215 | | S n => S n * fac n 216 | end. 217 | 218 | Definition fac_step (a: nat * nat) : nat * nat := 219 | (S (fst a), S (fst a) * snd a). 220 | 221 | Fact iter_fact n : 222 | (n, fac n) = iter fac_step n (0,1). 223 | Proof. 224 | induction n as [|n IH]. 225 | - reflexivity. 226 | - cbn. rewrite <-IH. reflexivity. 227 | Qed. 228 | 229 | (** Even *) 230 | 231 | Definition Even f n := 232 | match n with 233 | | 0 => true 234 | | 1 => false 235 | | S (S n) => f n 236 | end. 237 | 238 | Fixpoint even n := 239 | match n with 240 | | 0 => true 241 | | S n => negb (even n) 242 | end. 243 | 244 | Fact even_correct n : 245 | even n = Even even n. 246 | Proof. 247 | destruct n. reflexivity. 248 | destruct n. reflexivity. 249 | cbn. apply negb_involutive. 250 | Qed. 251 | 252 | Definition even_iter n := iter negb n true. 253 | 254 | Fact even_iter_correct n : 255 | even_iter n = Even even_iter n. 256 | Proof. 257 | destruct n. reflexivity. 258 | destruct n. reflexivity. 259 | cbn. apply negb_involutive. 260 | Qed. 261 | 262 | (** Fibonacci function *) 263 | 264 | Definition Fib f n := 265 | match n with 266 | | 0 => 0 267 | | 1 => 1 268 | | S (S n) => f n + f (S n) 269 | end. 270 | 271 | Fixpoint fib_bool' n b := 272 | match n, b with 273 | | 0, false => 0 274 | | 0, true => 1 275 | | S n, false => fib_bool' n true 276 | | S n, true => fib_bool' n false + fib_bool' n true 277 | end. 278 | 279 | Arguments fib_bool' : simpl nomatch. 280 | 281 | Definition fib_bool n := fib_bool' n false. 282 | 283 | Check fib_bool. 284 | Compute fib_bool 10. 285 | 286 | Fact fib_bool_correct n : 287 | fib_bool n = Fib fib_bool n. 288 | Proof. 289 | destruct n. reflexivity. 290 | destruct n. reflexivity. 291 | reflexivity. 292 | Qed. 293 | 294 | Definition fib_step (a: nat * nat) := (snd a, fst a + snd a). 295 | Definition fib_iter n := fst (iter fib_step n (0,1)). 296 | 297 | Check fib_iter. 298 | Compute fib_iter 10. 299 | 300 | Fact fib_iter_correct n : 301 | fib_iter n = Fib fib_iter n. 302 | Proof. 303 | destruct n. 304 | - cbn. reflexivity. 305 | - destruct n. 306 | + cbn. reflexivity. 307 | + cbn. unfold fib_iter. reflexivity. 308 | Qed. 309 | 310 | (** Ackermann function *) 311 | 312 | Definition Acker f x y := 313 | match x, y with 314 | | 0, y => S y 315 | | S x, 0 => f x 1 316 | | S x, S y => f x (f (S x) y) 317 | end. 318 | 319 | Fixpoint acker' (f: nat -> nat) (x:nat) : nat := 320 | match x with 321 | | 0 => f 1 322 | | S x => f (acker' f x) 323 | end. 324 | 325 | Fixpoint acker (x:nat) : nat -> nat := 326 | match x with 327 | | 0 => S 328 | | S x => acker' (acker x) 329 | end. 330 | 331 | Check acker. 332 | Compute acker 3 5. 333 | 334 | Fact acker_correct x y : 335 | acker x y = Acker acker x y. 336 | Proof. 337 | destruct x. reflexivity. 338 | destruct y. reflexivity. 339 | reflexivity. 340 | Qed. 341 | 342 | Definition B f y := iter f (S y) 1. 343 | Definition A x := iter B x S. 344 | 345 | Check A. 346 | Compute A 3 5. 347 | 348 | Fact acker_iter_correct x y : 349 | A x y = Acker A x y. 350 | Proof. 351 | destruct x. reflexivity. 352 | destruct y. reflexivity. 353 | reflexivity. 354 | Qed. 355 | 356 | (** Automation available for arithmetic proofs. *) 357 | From Stdlib Require Import Lia. 358 | 359 | Goal forall x y, x + y - y = x. 360 | Proof. 361 | lia. 362 | Qed. 363 | Goal forall x y, x * y = y * x. 364 | Proof. 365 | lia. 366 | Qed. 367 | Goal forall x y z, x * y * z = (x * y) * z. 368 | Proof. 369 | lia. 370 | Qed. 371 | Goal forall x y z, x*(y + z) = x*y + x*z. 372 | Proof. 373 | lia. 374 | Qed. 375 | 376 | (** Tactics used: 377 | reflexivity, symmetry, f_equal, cbn, rewrite, 378 | apply, exact, destruct, induction, intros, 379 | lia *) 380 | -------------------------------------------------------------------------------- /rocq/ids_gcd.v: -------------------------------------------------------------------------------- 1 | From Stdlib Require Import Arith Lia. 2 | Notation "'Sigma' x .. y , p" := 3 | (sigT (fun x => .. (sigT (fun y => p)) ..)) 4 | (at level 200, x binder, right associativity, 5 | format "'[' 'Sigma' '/ ' x .. y , '/ ' p ']'") 6 | : type_scope. 7 | Notation pi1 := projT1. 8 | Notation pi2 := projT2. 9 | 10 | Definition size_rec {X: Type} (sigma: X -> nat) {p: X -> Type} : 11 | (forall x, (forall y, sigma y < sigma x -> p y) -> p x) -> 12 | (forall x, p x). 13 | Proof. 14 | intros F. 15 | enough (forall n x, sigma x < n -> p x) as H. 16 | { intros x. apply (H (S (sigma x))). lia. } 17 | induction n as [|n IH]; intros x H. 18 | - exfalso. lia. 19 | - apply F. intros y H1. apply IH. lia. 20 | Defined. 21 | Definition size_rec2 {X Y: Type} (sigma: X -> Y -> nat) {p: X -> Y -> Type} : 22 | (forall x y, (forall x' y', sigma x' y' < sigma x y -> p x' y') -> p x y) -> 23 | (forall x y, p x y). 24 | Proof. 25 | intros F. 26 | enough (forall '(x,y), p x y) as H. 27 | { intros x y. apply (H (x,y)). } 28 | refine (size_rec (fun '(x,y) => sigma x y) (fun '(x,y) IH => _)). cbn in IH. 29 | apply F. intros x' y' H. apply (IH (x',y')), H. 30 | Defined. 31 | 32 | Notation functional p := (forall x y z z', p x y z -> p x y z' -> z = z'). 33 | 34 | Definition gcd_rel (p: nat -> nat -> nat -> Prop) : Prop := 35 | (forall y, p 0 y y) /\ 36 | (forall x y z, p x y z -> p y x z) /\ 37 | (forall x y z, x <= y -> p x (y - x) z -> p x y z). 38 | 39 | Inductive G : nat -> nat -> nat -> Prop := 40 | | G1 y : G 0 y y 41 | | G2 x y z: G x y z -> G y x z 42 | | G3 x y z : x <= y -> G x (y - x) z -> G x y z. 43 | 44 | Fact G_gcd_rel : 45 | gcd_rel G. 46 | Proof. 47 | split. exact G1. split. exact G2. exact G3. 48 | Qed. 49 | 50 | Fact G_cont p x y z : 51 | gcd_rel p -> G x y z -> p x y z. 52 | Proof. 53 | intros (H1&H2&H3). hnf. 54 | induction 1 as [y|x y z _ IH|x y z H _ IH]. 55 | - apply H1. 56 | - apply H2, IH. 57 | - apply H3; easy. 58 | Qed. 59 | 60 | Fact G_total : 61 | forall x y, Sigma z, G x y z. 62 | Proof. 63 | refine (size_rec2 (fun x y => x + y) _). 64 | intros x y IH. 65 | destruct x. 66 | - exists y. apply G1. 67 | - destruct y. 68 | + exists (S x). apply G2, G1. 69 | + destruct (le_lt_dec x y) as [H|H]. 70 | * specialize (IH (S x) (y - x)) as [z IH]. lia. 71 | exists z. apply G3. lia. exact IH. 72 | * specialize (IH (S y) (x - y)) as [z IH]. lia. 73 | exists z. apply G2, G3. lia. apply IH. 74 | Defined. 75 | 76 | Definition gcd x y := pi1 (G_total x y). 77 | Compute gcd 63 49. 78 | 79 | Fact G_agree x y z p : 80 | gcd_rel p -> functional p -> (G x y z <-> p x y z). 81 | Proof. 82 | intros H1 H2. split. 83 | - apply G_cont, H1. 84 | - intros H3. 85 | destruct (G_total x y) as [z' H]. 86 | enough (z = z') as <- by exact H. 87 | eapply H2. exact H3. 88 | apply G_cont; assumption. 89 | Qed. 90 | 91 | Inductive G': nat -> nat -> nat -> Prop := 92 | | G'1: forall y, G' 0 y y 93 | | G'2: forall x, G' (S x) 0 (S x) 94 | | G'3: forall x y z, x <= y -> G' (S x) (y - x) z -> G' (S x) (S y) z 95 | | G'4: forall x y z, y < x -> G' (x - y) (S y) z -> G' (S x) (S y) z. 96 | 97 | Fact G'_sym x y z : 98 | G' x y z -> G' y x z. 99 | Proof. 100 | induction 1 as [y|x|x y z H H1 IH|x y z H _ IH]. 101 | - destruct y; constructor. 102 | - apply G'1. 103 | - destruct (Nat.eq_dec x y) as [<-|H2]. 104 | + apply G'3. exact H. exact H1. 105 | + apply G'4. lia. exact IH. 106 | - apply G'3. lia. exact IH. 107 | Qed. 108 | 109 | Fact G'_gcd_rel : 110 | gcd_rel G'. 111 | Proof. 112 | split. constructor. 113 | split. exact G'_sym. 114 | intros x y z H. 115 | destruct x. 116 | - replace (y-0) with y by lia. easy. 117 | - destruct y. 118 | + lia. 119 | + cbn. apply G'3. lia. 120 | Qed. 121 | 122 | Fact G'_inv : 123 | forall x y z, G' x y z -> 124 | match x, y return Prop with 125 | | 0, y => z = y 126 | | S x, 0 => z = S x 127 | | S x, S y => if le_lt_dec x y 128 | then G' (S x) (y - x) z 129 | else G' (x - y) (S y) z 130 | end. 131 | Proof. 132 | destruct 1 as [y|x|x y z H1 H2|x y z H1 H2]. 133 | - reflexivity. 134 | - reflexivity. 135 | - destruct le_lt_dec as [H|H]. easy. exfalso. lia. 136 | - destruct le_lt_dec as [H|H]. exfalso. lia. easy. 137 | Defined. 138 | 139 | Fact G'_fun : 140 | functional G'. 141 | Proof. 142 | intros x y z z'. 143 | induction 1 as [y|x|x y z H _ IH|x y z H _ IH]. 144 | all:intros H1%G'_inv; revert H1. 145 | - easy. 146 | - easy. 147 | - destruct le_lt_dec as [H1|H1]. exact IH. lia. 148 | - destruct le_lt_dec as [H1|H1]. lia. exact IH. 149 | Qed. 150 | 151 | Fact G_G' {x y z} : 152 | G x y z <-> G' x y z. 153 | Proof. 154 | apply G_agree. 155 | - apply G'_gcd_rel. 156 | - apply G'_fun. 157 | Qed. 158 | 159 | Fact G_fun : 160 | functional G. 161 | Proof. 162 | hnf. intros * H1%G_G' H2%G_G'. 163 | eapply G'_fun; eassumption. 164 | Qed. 165 | -------------------------------------------------------------------------------- /rocq/ids_less.v: -------------------------------------------------------------------------------- 1 | Definition iffT (X Y: Type) : Type := (X -> Y) * (Y -> X). 2 | Notation "X <=> Y" := (iffT X Y) (at level 95, no associativity). 3 | From Stdlib Require Import Arith Lia Eqdep_dec. 4 | Notation sig := sigT. 5 | Notation Sig := existT. 6 | Notation pi1 := projT1. 7 | Notation pi2 := projT2. 8 | Notation "'Sigma' x .. y , p" := 9 | (sig (fun x => .. (sig (fun y => p)) ..)) 10 | (at level 200, x binder, right associativity, 11 | format "'[' 'Sigma' '/ ' x .. y , '/ ' p ']'") 12 | : type_scope. 13 | 14 | Fact inr_injective X Y (y y': Y) : 15 | inr X y = inr y' -> y = y'. 16 | Proof. 17 | intros H % (f_equal (fun a: X + Y => match a with 18 | | inl x => y 19 | | inr y => y 20 | end)). 21 | exact H. 22 | Qed. 23 | 24 | Definition cast {X} {x y: X} {p: X -> Type} 25 | : x = y -> p x -> p y 26 | := fun e a => match e with eq_refl => a end. 27 | 28 | Fact UIP_nat : 29 | forall {x: nat} (e: x = x), e = eq_refl. 30 | Proof. 31 | exact UIP_refl_nat. 32 | Qed. 33 | Fact cast_nat {p: nat -> Type} {x: nat} {y: p x} : 34 | forall e: x = x, cast (p:= p) e y = y. 35 | Proof. 36 | intros e. rewrite (UIP_nat e). reflexivity. 37 | Qed. 38 | Fact DPI_nat : 39 | forall (p: nat -> Type) n a b, Sig p n a = Sig p n b -> a = b. 40 | Proof. 41 | intros p. 42 | enough (forall a b: sig p, a = b -> forall e: pi1 a = pi1 b, cast e (pi2 a) = pi2 b) as H. 43 | { intros x u v H1. refine (H _ _ H1 eq_refl). } 44 | intros a b []. apply cast_nat. 45 | Qed. 46 | 47 | (*** Binary Variant *) 48 | Module Binary. 49 | Inductive L (x: nat) : nat -> Type := 50 | | L1: L x (S x) 51 | | L2: forall y z, L x y -> L y z -> L x z. 52 | 53 | Fact L_complete x y : 54 | x < y -> L x y. 55 | Proof. 56 | induction y as [|y IH]. 57 | - lia. 58 | - intros H. 59 | destruct (Nat.eq_dec x y) as [<-|H1]. 60 | + constructor. 61 | + eapply L2. 2:exact (L1 y). 62 | apply IH. lia. 63 | Qed. 64 | 65 | Definition L_elim (p: forall x y, L x y -> Type) 66 | : (forall x, p x (S x) (L1 x)) -> 67 | (forall x y z a b, p x y a -> p y z b -> p x z (L2 x y z a b)) -> 68 | forall x y a, p x y a 69 | := fun e1 e2 => fix f x _ a := 70 | match a with 71 | | L1 _ => e1 x 72 | | L2 _ y z a b => e2 x y z a b (f x y a) (f y z b) 73 | end. 74 | 75 | Fact L_sound {x y} : 76 | L x y -> x < y. 77 | Proof. 78 | induction 1; lia. 79 | Qed. 80 | 81 | Fixpoint depth_left {x y} (a: L x y) : nat := 82 | match a with 83 | | L1 _ => 1 84 | | L2 _ y z a b => S (depth_left a) 85 | end. 86 | 87 | Goal 88 | L2 3 4 6 (L1 3) (L2 4 5 6 (L1 4) (L1 5)) 89 | <> L2 3 5 6 (L2 3 4 5 (L1 3) (L1 4)) (L1 5) . 90 | Proof. 91 | intros [=] % (f_equal depth_left). 92 | Qed. 93 | 94 | 95 | Definition L_inv {x y} : 96 | L x y -> sum (y = S x) (Sigma z, prod (L x z) (L z y)). 97 | Proof. 98 | destruct 1 as [|z y a b]. 99 | - left; reflexivity. 100 | - right. exists z. easy. 101 | Defined. 102 | 103 | Compute fun x => L_inv (L1 x). 104 | Compute fun x y z a b => L_inv (L2 x y z a b). 105 | 106 | Fact L2_injective x y z a b a' b': 107 | L2 x z y a b = L2 x z y a' b' -> (a,b) = (a',b'). 108 | Proof. 109 | intros H. 110 | apply (f_equal L_inv) in H. cbn in H. 111 | apply inr_injective in H. 112 | apply DPI_nat in H. 113 | exact H. 114 | Qed. 115 | 116 | Fact L_inv_full {x y} : 117 | forall a: L x y, 118 | match y return L x y -> Type with 119 | | 0 => fun _ => False 120 | | S y => fun a => sum 121 | (Sigma e: y = x, cast (p:= fun z => L x (S z)) e a = L1 x) 122 | (Sigma z a1 a2, a = L2 x z (S y) a1 a2) 123 | end a. 124 | Proof. 125 | destruct a as [|y z a1 a2]. 126 | - left. exists eq_refl. reflexivity. 127 | - destruct z. 128 | + generalize (L_sound a2). lia. 129 | + right. exists y, a1, a2. reflexivity. 130 | Qed. 131 | End Binary. 132 | 133 | (*** Linear Indexed-Determined Variant *) 134 | Module Linear. 135 | Inductive L : nat -> nat -> Type := 136 | | L1 : forall y, L 0 (S y) 137 | | L2 : forall x y, L x y -> L (S x) (S y). 138 | 139 | Definition L_elim (p: forall x y, L x y -> Type) 140 | : (forall y, p 0 (S y) (L1 y)) -> 141 | (forall x y a, p x y a -> p (S x) (S y) (L2 x y a)) -> 142 | forall x y a, p x y a 143 | := fun e1 e2 => fix f x y a := match a with 144 | | L1 y => e1 y 145 | | L2 x y a => e2 x y a (f x y a) 146 | end. 147 | 148 | Goal forall x y, 149 | L x y <=> x < y. 150 | Proof. 151 | split. 152 | - induction 1 as [y|x y a IH]; lia. 153 | - induction x as [|x IH] in y |-*. 154 | all: (destruct y; try lia). 155 | all: intros H; constructor. 156 | apply IH. lia. 157 | Qed. 158 | 159 | Definition L_inv {x y} (a: L x y) : 160 | match x, y return L x y -> Type with 161 | | 0, S y => fun a => a = L1 y 162 | | S x, S y => fun a => Sigma a', a = L2 x y a' 163 | | _, _ => fun a => False 164 | end a. 165 | Proof. 166 | destruct a as [y|x y a]. 167 | - reflexivity. 168 | - exists a. reflexivity. 169 | Defined. 170 | 171 | Compute fun x => L_inv (L1 x). 172 | Compute fun x y a => L_inv (L2 x y a). 173 | 174 | Fact L_unique x y : 175 | forall a b: L x y, a = b. 176 | Proof. 177 | induction a as [y|x y a IH]; 178 | intros b; generalize (L_inv b). 179 | - easy. 180 | - intros [a' ->]. f_equal. apply IH. 181 | Qed. 182 | 183 | Goal forall x, L x 0 -> False. 184 | Proof. 185 | intros [|x] a; generalize (L_inv a); easy. 186 | Qed. 187 | 188 | Goal forall x, L x x -> False. 189 | Proof. 190 | induction x as [|x IH]. 191 | - intros a. contradict (L_inv a). 192 | - intros a. destruct (L_inv a) as [a' ->]. easy. 193 | Qed. 194 | 195 | End Linear. 196 | 197 | (*** Linear Left-Anchored Variant *) 198 | Module Linear_Anchored. 199 | Inductive L (x: nat) : nat -> Type := 200 | | L1 : L x (S x) 201 | | L2 : forall y, L x y -> L x (S y). 202 | 203 | Goal forall x y, 204 | L x y <=> x < y. 205 | Proof. 206 | split. 207 | - induction 1 as [|y a IH]; lia. 208 | - induction y as [|y IH]. 209 | + lia. 210 | + destruct (Nat.eq_dec x y) as [<-|H]. 211 | * intros _. constructor. 212 | * intros H1. constructor. apply IH. lia. 213 | Qed. 214 | 215 | Definition L_inv {x y} : 216 | L x y -> sum (y = S x) (Sigma y', prod (y = S y') (L x y')). 217 | Proof. 218 | destruct 1 as [|y a]. 219 | - left. reflexivity. 220 | - right. exists y. easy. 221 | Defined. 222 | 223 | Compute fun x => L_inv (L1 x). 224 | Compute fun x y a => L_inv (L2 x y a). 225 | 226 | Fact L2_injective x y a b : 227 | L2 x y a = L2 x y b -> a = b. 228 | Proof. 229 | intros H % (f_equal L_inv) 230 | % inr_injective 231 | % DPI_nat 232 | % (f_equal snd). 233 | exact H. 234 | Qed. 235 | 236 | Fact L_inv_full {x y} : 237 | forall a: L x y, 238 | match y return L x y -> Type with 239 | | 0 => fun _ => False : Type 240 | | S y => fun a => sum 241 | (Sigma e: y = x, cast (p:= fun z => L x (S z)) e a = L1 x) 242 | (Sigma a', a = L2 x y a') 243 | end a. 244 | Proof. 245 | destruct a as [|y a]. 246 | - left. exists eq_refl. reflexivity. 247 | - right. exists a. reflexivity. 248 | Defined. 249 | 250 | Fact L_lt {x y} : 251 | L x y -> x < y. 252 | Proof. 253 | induction 1 as [|y a IH]; lia. 254 | Qed. 255 | 256 | Fact L_unique x : 257 | forall y (a b: L x y), a = b. 258 | Proof. 259 | induction a as [|y a IH]; 260 | intros b; destruct (L_inv_full b) as [[e H]|[b' H]]. 261 | - rewrite <-H. refine (cast_nat _). 262 | - exfalso. specialize (L_lt b'). lia. 263 | - exfalso. specialize (L_lt a). lia. 264 | - rewrite H. f_equal. apply IH. 265 | Qed. 266 | End Linear_Anchored. 267 | -------------------------------------------------------------------------------- /rocq/ids_less_binary.v: -------------------------------------------------------------------------------- 1 | Definition iffT (X Y: Type) : Type := (X -> Y) * (Y -> X). 2 | Notation "X <=> Y" := (iffT X Y) (at level 95, no associativity). 3 | From Stdlib Require Import Arith Lia Eqdep_dec. 4 | Notation sig := sigT. 5 | Notation Sig := existT. 6 | Notation pi1 := projT1. 7 | Notation pi2 := projT2. 8 | Notation "'Sigma' x .. y , p" := 9 | (sig (fun x => .. (sig (fun y => p)) ..)) 10 | (at level 200, x binder, right associativity, 11 | format "'[' 'Sigma' '/ ' x .. y , '/ ' p ']'") 12 | : type_scope. 13 | 14 | Inductive L (x: nat) : nat -> Type := 15 | | L1: L x (S x) 16 | | L2: forall y z, L x y -> L y z -> L x z. 17 | 18 | Fact L_complete x y : 19 | x < y -> L x y. 20 | Proof. 21 | induction y as [|y IH]. 22 | - lia. 23 | - intros H. 24 | destruct (Nat.eq_dec x y) as [<-|H1]. 25 | + constructor. 26 | + eapply L2. 2:exact (L1 y). 27 | apply IH. lia. 28 | Qed. 29 | 30 | Definition L_elim (p: forall x y, L x y -> Type) 31 | : (forall x, p x (S x) (L1 x)) -> 32 | (forall x y z a b, p x y a -> p y z b -> p x z (L2 x y z a b)) -> 33 | forall x y a, p x y a 34 | := fun e1 e2 => fix f x _ a := 35 | match a with 36 | | L1 _ => e1 x 37 | | L2 _ y z a b => e2 x y z a b (f x y a) (f y z b) 38 | end. 39 | 40 | Fact L_sound {x y} : 41 | L x y -> x < y. 42 | Proof. 43 | induction 1; lia. 44 | Qed. 45 | 46 | Fixpoint depth_left {x y} (a: L x y) : nat := 47 | match a with 48 | | L1 _ => 1 49 | | L2 _ y z a b => S (depth_left a) 50 | end. 51 | 52 | Goal 53 | L2 3 4 6 (L1 3) (L2 4 5 6 (L1 4) (L1 5)) 54 | <> L2 3 5 6 (L2 3 4 5 (L1 3) (L1 4)) (L1 5) . 55 | Proof. 56 | intros [=] % (f_equal depth_left). 57 | Qed. 58 | 59 | Definition L_inv {x y} 60 | : L x y -> sum (y = S x) (Sigma z, prod (L x z) (L z y)) 61 | := fun a => match a with 62 | | L1 _ => inl eq_refl 63 | | L2 _ z y a b => inr (Sig _ z (a,b)) 64 | end. 65 | 66 | Compute fun x => L_inv (L1 x). 67 | Compute fun x y z a b => L_inv (L2 x y z a b). 68 | 69 | Fact inl_inr_disjoint X Y x y : 70 | inl Y x = inr X y -> False. 71 | Proof. 72 | intros H % (f_equal (fun a: X + Y => match a with 73 | | inl _ => False 74 | | inr _ => True 75 | end)). 76 | rewrite H. exact I. 77 | Qed. 78 | 79 | Goal forall x y a b, L1 x <> L2 x y (S x) a b. 80 | Proof. 81 | intros * H %(f_equal L_inv). cbn in H. 82 | eapply inl_inr_disjoint, H. 83 | Qed. 84 | 85 | (** Injectivity of L2 *) 86 | 87 | Fact inr_injective X Y (y y': Y) : 88 | inr X y = inr y' -> y = y'. 89 | Proof. 90 | intros H % (f_equal (fun a: X + Y => match a with 91 | | inl x => y 92 | | inr y => y 93 | end)). 94 | exact H. 95 | Qed. 96 | 97 | Section L2_injective. 98 | Variable DPI_nat: forall (p: nat -> Type) n a b, Sig p n a = Sig p n b -> a = b. 99 | 100 | Fact L2_injective x y z a b a' b': 101 | L2 x z y a b = L2 x z y a' b' -> (a,b) = (a',b'). 102 | Proof. 103 | intros H. 104 | apply (f_equal L_inv) in H. cbn in H. 105 | apply inr_injective in H. 106 | apply DPI_nat in H. 107 | exact H. 108 | Qed. 109 | End L2_injective. 110 | -------------------------------------------------------------------------------- /rocq/indeq.v: -------------------------------------------------------------------------------- 1 | (*** Inductive Equality *) 2 | From Stdlib Require Import Arith Bool. 3 | Notation sig := sigT. 4 | Notation Sig := existT. 5 | Notation pi1 := projT1. 6 | Notation pi2 := projT2. 7 | Notation "'if' x 'is' p 'then' A 'else' B" := 8 | (match x with p => A | _ => B end) 9 | (at level 200, p pattern,right associativity). 10 | Definition FE : Prop := 11 | forall X (p: X -> Type) (f g: forall x, p x), (forall x, f x = g x) -> f = g. 12 | Definition stable (P: Prop) := 13 | ~~P -> P. 14 | 15 | Definition J {X} {x: X} (p: forall y, x = y -> Type) 16 | : p x eq_refl -> forall y e, p y e 17 | := fun a y e => match e with eq_refl => a end. 18 | 19 | Definition cast {X} [p: X -> Type] {x: X} 20 | : forall {y}, x = y -> p x -> p y 21 | := fun y e a => match e with eq_refl => a end. 22 | 23 | Goal forall X (x y: X) p (e: x = y) a, 24 | cast e a = J (fun y e => p y) a y e. 25 | Proof. 26 | reflexivity. 27 | Qed. 28 | 29 | Goal forall X (x y: X) p (a: p x), 30 | cast eq_refl a = a. 31 | Proof. 32 | reflexivity. 33 | Qed. 34 | 35 | Definition id {X} (x: X) := x. 36 | Arguments id {X} x /. 37 | 38 | Section SigmaTau. 39 | Variable X: Type. 40 | Implicit Types x y z: X. 41 | Definition sigma {x y} 42 | : x = y -> y = x 43 | := fun e => cast (p:= fun y => y = x) e eq_refl. 44 | Definition tau {x y z} 45 | : x = y -> y = z -> x = z 46 | := fun e => cast (p:= fun y => y = z -> x = z) e id. 47 | Definition phi {A} (f: X -> A) {x y} 48 | : x = y -> f x = f y 49 | := fun e => cast (p:= fun y => f x = f y) e eq_refl. 50 | 51 | Fact sigma_involutive x y (e: x = y) : 52 | sigma (sigma e) = e. 53 | Proof. 54 | unfold sigma. destruct e. reflexivity. 55 | Qed. 56 | Fact tau_trans x y z z' (e1: x = y) (e2: y = z) (e3: z = z') : 57 | tau e1 (tau e2 e3) = tau (tau e1 e2) e3. 58 | Proof. 59 | unfold tau. destruct e1; cbn. destruct e2; cbn. reflexivity. 60 | Qed. 61 | Fact tau_sigma x y (e: x = y) : 62 | tau e (sigma e) = eq_refl. 63 | Proof. 64 | unfold sigma, tau. destruct e; cbn. reflexivity. 65 | Qed. 66 | End SigmaTau. 67 | Arguments sigma {X x y}. 68 | Arguments tau {X x y z}. 69 | 70 | Section UIP. 71 | Variable X: Type. 72 | Definition UIP := forall (x y: X) (e e': x = y), e = e'. 73 | Definition UIP' := forall (x : X) (e: x = x), e = eq_refl. 74 | Definition K_Streicher := forall (x: X) (p: x = x -> Prop), p (eq_refl x) -> forall e, p e. 75 | Definition CD := forall (p: X -> Type) (x: X) (a: p x) (e: x = x), cast e a = a. 76 | Definition DPI := forall (p: X -> Type) x u v, Sig p x u = Sig p x v -> u = v. 77 | 78 | Goal UIP -> UIP'. 79 | Proof. 80 | intros H x e. apply H. 81 | Qed. 82 | 83 | Goal UIP' -> UIP. 84 | Proof. 85 | intros H x y e e'. destruct e'. apply H. 86 | Qed. 87 | 88 | Goal UIP' -> K_Streicher. 89 | Proof. 90 | intros H x p a e. rewrite (H x e). exact a. 91 | Qed. 92 | 93 | Goal K_Streicher -> CD. 94 | Proof. 95 | intros H p x a. apply H. reflexivity. 96 | Qed. 97 | 98 | Goal CD -> DPI. 99 | Proof. 100 | intros H p x. 101 | enough (forall a b: sig p, a = b -> forall e: pi1 a = pi1 b, cast e (pi2 a) = pi2 b) as H'. 102 | - intros u v e'. apply (H' _ _ e' eq_refl). 103 | - intros a b <-. apply H. 104 | Qed. 105 | 106 | Goal DPI -> UIP'. 107 | Proof. 108 | intros H x e. 109 | apply (H (eq x)). revert e. 110 | enough (forall y (e: x = y), Sig (eq x) y e = Sig (eq x) x eq_refl) as H'. 111 | - apply H'. 112 | - intros y <-. reflexivity. 113 | Qed. 114 | 115 | Lemma cast_eq {x y: X} : 116 | forall e: x = y, cast e (eq_refl x) = e. 117 | Proof. 118 | destruct e. reflexivity. 119 | Qed. 120 | 121 | Goal CD -> UIP'. 122 | Proof. 123 | intros H x. 124 | enough (forall y (e: x = y), e = cast e (eq_refl x)) as H1. 125 | - intros e. rewrite (H1 x e). apply H. 126 | - destruct e. reflexivity. 127 | Qed. 128 | End UIP. 129 | 130 | Definition HF X (f: forall x y: X, x = y -> x = y) := 131 | forall x y (e e': x = y), f x y e = f x y e'. 132 | 133 | Lemma Hedberg' X : 134 | ex (HF X) -> UIP X. 135 | Proof. 136 | intros [f H] x y. 137 | enough (forall e: x = y, tau (f x y e) (sigma (f y y eq_refl)) = e) as H1. 138 | - intros e e'. specialize (H _ _ e e'). congruence. 139 | - destruct e. apply tau_sigma. 140 | Qed. 141 | 142 | Theorem Hedberg X : 143 | (forall x y: X, x = y \/ x <> y) -> UIP X. 144 | Proof. 145 | intros d. apply Hedberg'. 146 | exists (fun x y e => if d x y is or_introl e' then e' else e). 147 | intros x y e e'. 148 | destruct d as [e''|h]. 149 | - reflexivity. 150 | - exfalso. exact (h e). 151 | Qed. 152 | 153 | (* Contributed by Dominik Kirst, 18 Feb. 2021 *) 154 | Fact FE_stable_HF X : 155 | FE -> (forall x y: X, stable (x = y)) -> sig (HF X). 156 | Proof. 157 | intros F G. 158 | exists (fun x y e => G x y (fun f => match f e with end)). 159 | intros x y e e'. f_equal. apply F. intros f. 160 | exfalso. exact (f e). 161 | Qed. 162 | 163 | Fact FE_test_eq_stable : 164 | FE -> forall f g: nat -> bool, stable (f = g). 165 | Proof. 166 | intros F f g H. apply F. intros x. 167 | destruct (bool_dec (f x) (g x)) as [H1|H1]. 168 | - exact H1. 169 | - contradict H. contradict H1. f_equal. 170 | Qed. 171 | 172 | 173 | (*** Direct proof of UIP_nat *) 174 | (* Idea from Andrej Dudenhefner *) 175 | 176 | Module UIP_nat. 177 | Lemma nat_eqdec_eq x : 178 | Nat.eq_dec x x = left eq_refl. 179 | Proof. 180 | induction x. reflexivity. 181 | simpl. rewrite IHx. reflexivity. 182 | Qed. 183 | Lemma UIP_nat' (x y: nat) : 184 | forall e: x = y, 185 | match Nat.eq_dec x y with 186 | | left e' => cast e' eq_refl = e 187 | | _ => True 188 | end. 189 | Proof. 190 | destruct e. rewrite nat_eqdec_eq. reflexivity. 191 | Qed. 192 | Fact UIP_refl_nat (x: nat) : 193 | forall e: x = x, e = eq_refl. 194 | Proof. 195 | intros e. 196 | generalize (UIP_nat' x x e). 197 | rewrite nat_eqdec_eq. intros <-. reflexivity. 198 | Qed. 199 | End UIP_nat. 200 | 201 | (*** UIP propagates to identity at Type *) 202 | 203 | Inductive Eq X (x: X) : X -> Set := Q : Eq X x x. 204 | (* Must write Set, Type will be downgraded to Prop *) 205 | Arguments Eq {X}. 206 | Arguments Q {X}. 207 | 208 | Check Eq nat. 209 | Check Eq Type. 210 | 211 | Definition D {X} {x y: X} 212 | : Eq x y -> x = y 213 | := fun a => match a with Q _ => eq_refl x end. 214 | Definition U {X} {x y: X} 215 | : x = y -> Eq x y 216 | := fun e => match e with eq_refl _ => Q x end. 217 | Fact UD_eq {X} {x y: X} (a: Eq x y) : 218 | U (D a) = a. 219 | Proof. 220 | destruct a. reflexivity. 221 | Qed. 222 | Fact DU_eq {X} {x y: X} (e: x = y) : 223 | D (U e) = e. 224 | Proof. 225 | destruct e. reflexivity. 226 | Qed. 227 | 228 | Fact Eq_fun X Y (f: X -> Y) (x x': X) : 229 | Eq x x' -> Eq (f x) (f x'). 230 | Proof. 231 | destruct 1. apply Q. 232 | Qed. 233 | 234 | Definition UIP_Eq X := forall (x y: X) (a b: Eq x y), Eq a b. 235 | Fact UIP_up X : 236 | UIP X -> UIP_Eq X. 237 | Proof. 238 | intros H x y a b. 239 | rewrite <-(UD_eq a), <-(UD_eq b). 240 | apply Eq_fun. 241 | apply U. apply H. 242 | Qed. 243 | Fact UIP_down X : 244 | UIP_Eq X -> UIP X. 245 | Proof. 246 | intros H x y e e'. 247 | rewrite <-(DU_eq e), <-(DU_eq e'). 248 | apply f_equal. apply D. apply H. 249 | Qed. 250 | 251 | -------------------------------------------------------------------------------- /rocq/indind.v: -------------------------------------------------------------------------------- 1 | (*** MPCTT, Chapter Indexed Inductives *) 2 | From Stdlib Require Import Lia. 3 | 4 | (*** Inductive Equality *) 5 | Module IndEquality. 6 | 7 | Inductive eq X (x: X) : X -> Prop := 8 | | Q : eq X x x. 9 | 10 | Check eq. 11 | Check Q. 12 | 13 | 14 | Definition E 15 | : forall X (x: X) (p: X -> Type), p x -> forall y, eq X x y -> p y 16 | := fun X x p a _ e => 17 | match e with 18 | | Q _ _ => a 19 | end. 20 | 21 | 22 | Definition R 23 | : forall X (x y: X) (p: X -> Type), eq X x y -> p x -> p y 24 | := fun X x _ p e => match e with 25 | | Q _ _ => fun a => a 26 | end. 27 | End IndEquality. 28 | 29 | (*** Reflexive Transitive Closure *) 30 | 31 | Module Star. 32 | Section Star. 33 | Variable X : Type. 34 | Implicit Type R: X -> X -> Prop. 35 | 36 | Inductive star (R: X -> X -> Prop) (x: X) : X -> Prop := 37 | | Nil : star R x x 38 | | Cons y z : R x y -> star R y z -> star R x z. 39 | 40 | Definition elim R (p: X -> X -> Prop) 41 | : (forall x, p x x) -> 42 | (forall x y z, R x y -> p y z -> p x z) -> 43 | forall x y, star R x y -> p x y 44 | := fun f1 f2 => fix f x _ a := 45 | match a with 46 | | Nil _ _ => f1 x 47 | | Cons _ _ x' z r a => f2 x x' z r (f x' z a) 48 | end. 49 | 50 | Implicit Type p: X -> X -> Prop. 51 | Definition reflexive p := forall x, p x x. 52 | Definition transitive p := forall x y z, p x y -> p y z -> p x z. 53 | Definition incl p p' := forall x y, p x y -> p' x y. 54 | 55 | Fact star_incl R : 56 | incl R (star R). 57 | Proof. 58 | intros x y r. eapply Cons. exact r. apply Nil. 59 | Qed. 60 | 61 | Fact star_trans R : 62 | transitive (star R). 63 | Proof. 64 | intros x y z. 65 | revert x y. 66 | refine (elim _ _ _ _). 67 | - easy. 68 | - intros x x' y r IH. 69 | intros H%IH. revert r H. apply Cons. 70 | Qed. 71 | 72 | (** We may also use the automatically generated eliminator for star. *) 73 | 74 | Check star_ind. 75 | 76 | Goal 77 | forall R, transitive (star R). 78 | Proof. 79 | induction 1 as [|x x' y r _ IH]. 80 | - easy. 81 | - intros H%IH. econstructor. exact r. exact H. 82 | Qed. 83 | 84 | Fact least R p : 85 | reflexive p -> 86 | transitive p -> 87 | incl R p -> 88 | incl (star R) p. 89 | Proof. 90 | intros H1 H2 H3. hnf. 91 | apply elim. 92 | - exact H1. 93 | - intros x y z H%H3. apply H2, H. 94 | Qed. 95 | 96 | Goal forall R, incl (star (star R)) (star R). 97 | Proof. 98 | intros R. apply least. 99 | - hnf. apply Nil. 100 | - apply star_trans. 101 | - easy. 102 | Qed. 103 | 104 | Fixpoint path R n x y : Prop := 105 | match n with 106 | | 0 => x = y 107 | | S n => exists x', R x x' /\ path R n x' y 108 | end. 109 | 110 | Goal forall R x y, 111 | star R x y <-> exists n, path R n x y. 112 | Proof. 113 | split. 114 | - induction 1 as [|x x' y r _ IH]. 115 | + exists 0. easy. 116 | + destruct IH as [n IH]. 117 | exists (S n). exists x'. easy. 118 | - intros [n H]. revert x H. 119 | induction n as [|n IH]; cbn. 120 | + intros x <-. apply Nil. 121 | + intros x (x'&H1&H2). 122 | eapply Cons. exact H1. 123 | apply IH, H2. 124 | Qed. 125 | 126 | (*** Index Elimination *) 127 | 128 | (** We can define R* as an inductive predicate without indices. **) 129 | 130 | Inductive star' (R: X -> X -> Prop) (x: X) (y:X) : Prop := 131 | | Nil' : x = y -> star' R x y 132 | | Cons' x' : R x x' -> star' R x' y -> star' R x y. 133 | 134 | Definition elim' R (p: X -> Prop) (y: X) 135 | : (forall x, x = y -> p x) -> 136 | (forall x x', R x x' -> p x' -> p x) -> 137 | forall x, star' R x y -> p x 138 | := fun f1 f2 => fix f x a := 139 | match a with 140 | | Nil' _ _ _ e => f1 x e 141 | | Cons' _ _ _ x' r a => f2 x x' r (f x' a) 142 | end. 143 | 144 | Goal forall R x y, star' R x y <-> star R x y. 145 | Proof. 146 | intros *; split. 147 | - revert x. apply elim'. 148 | + intros x <-. apply Nil. 149 | + intros * r IH. eapply Cons; eassumption. 150 | - revert x y. apply elim. 151 | + intros *. apply Nil'. reflexivity. 152 | + intros * r IH. eapply Cons'; eassumption. 153 | Qed. 154 | End Star. 155 | End Star. 156 | 157 | (** Index eliminatiom will not work for equality 158 | since Leibniz equality doesnt give us rewriting at type. *) 159 | 160 | Module Eq_without_index. 161 | 162 | Inductive eq' X (x: X) (y: X) : Prop := 163 | | L : (forall p: X -> Type, p x -> p y) -> eq' X x y. 164 | 165 | Inductive eq'' X (x: X) (y: X) : Prop := 166 | | L' : (forall p: X -> Prop, p x -> p y) -> eq'' X x y. 167 | 168 | Goal forall X x y, eq' X x y <-> x = y. 169 | Proof. 170 | intros *; split. 171 | - intros [H]. pattern y. apply H. reflexivity. 172 | - intros <-. apply L. auto. 173 | Qed. 174 | 175 | Goal forall X (x y: X) (p: X -> Type), 176 | eq' X x y -> p x -> p y. 177 | Proof. 178 | Fail intros * [H]. (* PDR kills it *) 179 | Abort. 180 | 181 | Goal forall X x y, eq' X x y <-> eq'' X x y. 182 | Proof. 183 | intros *; split. 184 | - intros [H]. pattern y. apply H. apply L'. auto. 185 | - intros [H]. pattern y. apply H. apply L. auto. 186 | Qed. 187 | 188 | End Eq_without_index. 189 | 190 | (*** Inductive Comparisons *) 191 | 192 | Module LE. 193 | Inductive le (x: nat) : nat -> Prop := 194 | | leE : le x x 195 | | leS y : le x y -> le x (S y). 196 | 197 | Definition elim (x: nat) (p: nat -> Prop) 198 | : p x -> 199 | (forall y, p y -> p (S y)) -> 200 | forall y, le x y -> p y 201 | := fun e1 e2 => fix f _ a := 202 | match a with 203 | | leE _ => e1 204 | | leS _ y a => e2 y (f y a) 205 | end. 206 | 207 | Fact le_correct x y : 208 | le x y <-> x <= y. 209 | Proof. 210 | split. 211 | - revert y. apply elim. 212 | + lia. 213 | + lia. 214 | - induction y as [|y IH]; intros H. 215 | + assert (x = 0) as -> by lia. apply leE. 216 | + assert (x = S y \/ x <= y) as [->|H1] by lia. 217 | * apply leE. 218 | * apply leS, IH. lia. 219 | Qed. 220 | End LE. 221 | 222 | (*** Inductive Numeral Types *) 223 | 224 | Module Fin. 225 | Fixpoint num n : Type := 226 | match n with 227 | | 0 => False 228 | | S n => option (num n) 229 | end. 230 | 231 | Inductive fin : nat -> Type := 232 | | Old: forall n, fin n -> fin (S n) 233 | | New: forall n, fin (S n). 234 | 235 | Definition fin_num 236 | : forall n, fin n -> num n 237 | := fix f n a := 238 | match a with 239 | | Old n a => Some (f n a) 240 | | New n => None 241 | end. 242 | 243 | Definition num_fin 244 | : forall n, num n -> fin n 245 | := fix f n := 246 | match n with 247 | | 0 => fun a => match a with end 248 | | S n => fun a => match a with 249 | | Some a => Old n (f n a) 250 | | None => New n 251 | end 252 | end. 253 | 254 | Fact num_fin_inv n a : 255 | num_fin n (fin_num n a) = a. 256 | Proof. 257 | induction a as [n a IH | n]; cbn. 258 | - f_equal. apply IH. 259 | - reflexivity. 260 | Qed. 261 | 262 | Definition elim (p: forall n, fin n -> Type) 263 | : (forall n a, p n a -> p (S n) (Old n a)) -> 264 | (forall n, p (S n) (New n)) -> 265 | forall n a, p n a 266 | := fun e1 e2 => fix f n a := 267 | match a with 268 | | Old n a => e1 n a (f n a) 269 | | New n => e2 n 270 | end. 271 | 272 | Fact num_fin_inv' n a : 273 | num_fin n (fin_num n a) = a. 274 | Proof. 275 | revert n a. apply elim; cbn. 276 | - intros n a IH. f_equal. exact IH. 277 | - reflexivity. 278 | Qed. 279 | 280 | Fact fin_num_inv n a : 281 | fin_num n (num_fin n a) = a. 282 | Proof. 283 | induction n as [|n IH]. 284 | - destruct a. 285 | - destruct a as [a|]; cbn. 286 | + f_equal. apply IH. 287 | + reflexivity. 288 | Qed. 289 | 290 | Goal fin 0 -> False. 291 | Proof. 292 | enough (forall n, fin n -> n = 0 -> False) by eauto. 293 | refine (elim _ _ _). 294 | - lia. 295 | - lia. 296 | Qed. 297 | 298 | Goal fin 0 -> False. 299 | Proof. 300 | intros a. apply (fin_num 0 a). 301 | Qed. 302 | End Fin. 303 | 304 | (*** Inductive Vector Types *) 305 | 306 | Module Vector. 307 | Section Vector. 308 | Variable X: Type. 309 | 310 | Fixpoint vec n : Type := 311 | match n with 312 | | 0 => unit 313 | | S n => X * vec n 314 | end. 315 | 316 | Inductive V : nat -> Type := 317 | | Nil : V 0 318 | | Cons n : X -> V n -> V (S n). 319 | 320 | Definition V_vec 321 | : forall n, V n -> vec n 322 | := fix f n a := 323 | match a with 324 | | Nil => tt 325 | | Cons n x a => (x, f n a) 326 | end. 327 | 328 | Definition vec_V 329 | : forall n, vec n -> V n 330 | := fix f n := 331 | match n with 332 | | 0 => fun _ => Nil 333 | | S n => fun v => let (x,v) := v in Cons n x (f n v) 334 | end. 335 | 336 | Fact vec_V_inv n v : 337 | vec_V n (V_vec n v) = v. 338 | Proof. 339 | induction v as [|n x v IH]; cbn. 340 | - reflexivity. 341 | - f_equal. exact IH. 342 | Qed. 343 | 344 | Definition elim (p: forall n, V n -> Type) 345 | : p 0 Nil -> 346 | (forall n x v, p n v -> p (S n) (Cons n x v)) -> 347 | forall n v, p n v 348 | := fun e1 e2 => fix f n v := 349 | match v with 350 | | Nil => e1 351 | | Cons n x v => e2 n x v (f n v) 352 | end. 353 | 354 | Fact vec_V_inv' n v : 355 | vec_V n (V_vec n v) = v. 356 | Proof. 357 | revert n v. apply elim; cbn. 358 | - reflexivity. 359 | - intros n x v IH. f_equal. exact IH. 360 | Qed. 361 | 362 | Fact V_vec_inv n v : 363 | V_vec n (vec_V n v) = v. 364 | Proof. 365 | induction n as [|n IH]; cbn. 366 | - destruct v. reflexivity. 367 | - destruct v as [x v]; cbn. f_equal. apply IH. 368 | Qed. 369 | End Vector. 370 | End Vector. 371 | 372 | (*** Post Correspondence Problem *) 373 | 374 | From Stdlib Require Import List. 375 | Import ListNotations. 376 | Notation "x 'el' A" := (In x A) (at level 70). 377 | 378 | Module Post. 379 | Inductive char := a | b. 380 | Definition string : Type := list char. 381 | Definition card : Type := string * string. 382 | Definition deck : Type := list card. 383 | 384 | Inductive post (D: deck) | : string -> string -> Prop := 385 | | post1 : forall A B, (A,B) el D -> post A B 386 | | post2 : forall A B A' B', (A,B) el D -> post A' B' -> post (A ++ A') (B ++ B'). 387 | 388 | Definition D : deck := [([a], []); ([b],[a]); ([], [b;b])]. 389 | 390 | Goal post D [b;b;a;a] [b;b;a;a]. 391 | Proof. 392 | refine (post2 _ [] [b;b] [b;b;a;a] [a;a] _ _). 393 | { cbn; auto. } 394 | refine (post2 _ [b] [a] [b;a;a] [a] _ _). 395 | { cbn; auto. } 396 | refine (post2 _ [b] [a] [a;a] [] _ _). 397 | { cbn; auto. } 398 | refine (post2 _ [a] [] [a] [] _ _). 399 | { cbn; auto. } 400 | refine (post1 _ [a] [] _). 401 | { cbn; auto. } 402 | Qed. 403 | End Post. 404 | 405 | 406 | 407 | 408 | 409 | 410 | -------------------------------------------------------------------------------- /rocq/leibniz.v: -------------------------------------------------------------------------------- 1 | (*** MPCTT, Chapter 4 *) 2 | 3 | (* We study conversion and propositonal equality. *) 4 | 5 | Definition id X : X -> X := fun x => x. 6 | 7 | Section Conversion. 8 | Variable p: nat -> Prop. 9 | 10 | Check fun a : p 4 => a. 11 | Check fun a : p (2 + 2) => a. 12 | Check fun a : p 4 => id (p (2 + 2)) a. 13 | 14 | (* Rocq uses the notation $s:t$ for applications "id t s" 15 | of the polymorphic identity function *) 16 | Check fun a : p 4 => a : p (2 + 2). 17 | Check id (p 4 -> p (2 + 2)) (fun a : p 4 => a). 18 | Check (fun a : p 4 => a) : p 4 -> p (2 + 2). 19 | 20 | (* One can force cenversions with the change tactic *) 21 | Goal p 4 -> p (2 + 2). 22 | Proof. 23 | intros a. 24 | change (p 4). 25 | exact a. 26 | Show Proof. 27 | Qed. 28 | 29 | Goal p 4 -> p (2 + 2). 30 | Proof. 31 | intros a. exact a. 32 | Show Proof. 33 | Qed. 34 | 35 | Goal p 4 -> p (2 + 2). 36 | Proof. 37 | intros a. 38 | apply (id (p 4)). 39 | exact a. 40 | Show Proof. 41 | Qed. 42 | 43 | End Conversion. 44 | 45 | (** Rocq implements propositional negation 46 | and propositional equivalence with 47 | predicates "not" and "iff". *) 48 | 49 | Locate "~". 50 | Print not. 51 | Locate "<->". 52 | Print iff. 53 | 54 | Goal forall X, ~X -> X -> False. 55 | intros X a. 56 | change (~X). 57 | exact a. 58 | Show Proof. 59 | Qed. 60 | 61 | Goal forall X, ~X -> X -> False. 62 | intros X a. exact a. 63 | Show Proof. 64 | Qed. 65 | 66 | Goal forall X Y, (X <-> Y) -> (X -> Y) /\ (Y -> X). 67 | Proof. 68 | intros * a. 69 | change (X <-> Y). 70 | Show Proof. (* not visible in partial proof term *) 71 | exact a. 72 | Show Proof. 73 | Qed. 74 | 75 | (* The pattern tactic does beta expansions *) 76 | 77 | (** Exercise Leibniz symmetry *) 78 | 79 | Fact Leibniz_symmetry X (x y: X) : 80 | (forall p: X -> Prop, p x -> p y) -> (forall p: X -> Prop, p y -> p x). 81 | Proof. 82 | intros F p. 83 | pattern y. 84 | Show Proof. (* pattern puts in a type ascription *) 85 | apply F. 86 | intros a. exact a. 87 | Show Proof. 88 | Qed. 89 | 90 | Check fun X (x y: X) (F: forall p: X -> Prop, p x -> p y) (p: X -> Prop) => 91 | F (fun y => p y -> p x) (fun a => a). 92 | 93 | (*** Abstract Equality *) 94 | 95 | Section AbstractEquality. 96 | Variable eq : forall X, X -> X -> Prop. 97 | Variable Q : forall X x, eq X x x. 98 | Variable R : forall X x y (p: X -> Prop), eq X x y -> p x -> p y. 99 | 100 | Arguments eq {X}. 101 | Arguments Q {X}. 102 | Arguments R {X x y}. 103 | Notation "s = t" := (eq s t) : type_scope. 104 | Notation "s <> t" := (~ s = t) : type_scope. 105 | 106 | Goal 2 + 3 = 9 - 4. 107 | Proof. 108 | apply (Q 5). 109 | Show Proof. 110 | Qed. 111 | 112 | Goal forall X (x y z: X), 113 | x = y -> y = z -> x = z. 114 | Proof. 115 | intros * e. 116 | pattern y. 117 | apply (R _ e). 118 | (* beta reduction was done tacitly *) 119 | exact (fun e => e). 120 | Show Proof. 121 | Qed. 122 | 123 | Check fun X (x y z : X) (e : x = y) => 124 | R (fun y => y = z -> x = z) e (fun e => e). 125 | 126 | (*** Basic Equational Facts *) 127 | 128 | Goal True <> False. 129 | Proof. 130 | intros e. 131 | pattern False. 132 | apply (R _ e). 133 | exact I. 134 | Show Proof. 135 | Qed. 136 | 137 | Check fun e : True = False => R (fun X => X) e I. 138 | 139 | Goal true <> false. 140 | Proof. 141 | intros e. 142 | change ((fun x:bool => if x then True else False) false). 143 | apply (R _ e). 144 | exact I. 145 | Show Proof. 146 | Qed. 147 | 148 | Check (fun e : true = false => 149 | R (fun x : bool => if x then True else False) e I). 150 | 151 | Goal forall x y, S x = S y -> x = y. 152 | Proof. 153 | intros x y e. 154 | change ((fun z => x = match z with 0 => 0 | S z' => z' end) (S y)). 155 | apply (R _ e). 156 | exact (Q x). 157 | Show Proof. 158 | Qed. 159 | 160 | Check fun (x y : nat) (e : S x = S y) => 161 | R (fun z => x = match z with 162 | | 0 => 0 163 | | S z' => z' 164 | end) e (Q x). 165 | 166 | Goal forall X (x y: X), 167 | x = y -> y = x. 168 | Proof. 169 | intros * e. 170 | pattern y. 171 | apply (R _ e). 172 | exact (Q x). 173 | Show Proof. 174 | Qed. 175 | 176 | Check fun X (x y : X) (e : x = y) => 177 | R (fun y => y = x) e (Q x). 178 | 179 | Goal forall X (x y z: X), 180 | x = y -> y = z -> x = z. 181 | Proof. 182 | intros * e. 183 | pattern y. 184 | apply (R _ e). 185 | exact (fun e => e). 186 | Show Proof. 187 | Qed. 188 | 189 | Check fun X (x y z : X) (e : x = y) => 190 | R (fun y => y = z -> x = z) e (fun e => e). 191 | 192 | Goal forall X (x y z: X), 193 | x = y -> y = z -> x = z. 194 | Proof. 195 | intros * e1 e2. 196 | (* claim is "(eq x) z" up to notation *) 197 | exact (R _ e2 e1). 198 | Show Proof. 199 | Qed. 200 | 201 | Check fun X (x y z : X) (e1 : x = y) (e2 : y = z) => 202 | R (eq x) e2 e1. 203 | 204 | (** Applicative closure laws *) 205 | 206 | Goal forall X Y (f: X -> Y) x x', 207 | x = x' -> f x = f x'. 208 | Proof. 209 | intros * e. 210 | pattern x'. 211 | apply (R _ e). 212 | apply Q. 213 | Show Proof. 214 | Qed. 215 | 216 | Goal forall X Y (f g: X -> Y) x, 217 | f = g -> f x = g x. 218 | Proof. 219 | intros * e. 220 | pattern g. 221 | apply (R _ e). 222 | apply Q. 223 | Show Proof. 224 | Qed. 225 | 226 | (** Leibniz characterization *) 227 | 228 | Goal forall X x y, 229 | x = y <-> forall p: X -> Prop, p x -> p y. 230 | Proof. 231 | intros *. split. 232 | - intros e p H. 233 | apply (R _ e H). 234 | - intros H. 235 | change ((fun z => x = z) y). 236 | apply H. 237 | apply Q. 238 | Show Proof. 239 | Qed. 240 | 241 | Goal (forall X:Prop, X -> X) <> (forall X:Prop, X). 242 | Proof. 243 | intros e. 244 | enough (forall X:Prop, X) as H. 245 | - apply H. 246 | - apply (R (fun X:Prop => X) e). 247 | exact (fun X x => x). 248 | Qed. 249 | 250 | End AbstractEquality. 251 | 252 | (*** Definition of Leibniz Equality *) 253 | 254 | Module Type EQ. 255 | Parameter eq : forall X, X -> X -> Prop. 256 | Parameter Q : forall X x, eq X x x. 257 | Parameter R : forall X x y (p: X -> Prop), eq X x y -> p x -> p y. 258 | End EQ. 259 | 260 | Module EqLeibniz : EQ. 261 | Definition eq X x y := forall p: X -> Prop, p x -> p y. 262 | Definition Q X x : eq X x x := fun p a => a. 263 | Definition R X x y p (f: eq X x y) := f p. 264 | End EqLeibniz. 265 | 266 | Print EqLeibniz.R. 267 | 268 | (* We sandbox the import of EqLeibniz 269 | since we will use Rocq's predefined equality afterwards. 270 | *) 271 | 272 | Module Sandbox. 273 | Import EqLeibniz. 274 | Print R. 275 | Print eq. 276 | End Sandbox. 277 | 278 | (*** Rocq's Predefined Equality *) 279 | 280 | (* From now on we use Rocq's predefined equality 281 | with the tactics "reflexivity" and "rewrite". *) 282 | 283 | Locate "=". 284 | Check @eq. 285 | Check @eq_refl. (* Q, used by reflexivity tactic*) 286 | Check @eq_ind. (* R *) 287 | Check @eq_ind_r. (* used by rewrite tactic *) 288 | 289 | Fact feq {X Y} (f: X -> Y) {x x'} : 290 | x = x' -> f x = f x'. 291 | Proof. 292 | intros e. rewrite e. reflexivity. 293 | Show Proof. 294 | Qed. 295 | 296 | Check @feq. 297 | Print feq. 298 | Compute @feq. (* No delta reduction *) 299 | 300 | (* The Fact-Proof-Qed format defines an abstract constant (above "feq"). 301 | The command "Print" still shows the definition of abstract constants 302 | introduced with "Qed". 303 | *) 304 | 305 | (* The "Fact" keyword is synonnymous with the keywords 306 | "Theorem", "Lemma", "Corollary", "Example". 307 | *) 308 | 309 | Definition pred (x: nat) : nat := 310 | match x with 0 => 0 | S x' => x' end. 311 | 312 | Goal forall x y, S x = S y -> x = y. 313 | Proof. 314 | intros * e. 315 | exact (feq pred e). 316 | Show Proof. 317 | Qed. 318 | 319 | Definition zero (x:nat) : Prop := 320 | if x then True else False. 321 | 322 | Goal forall x, S x <> 0. 323 | Proof. 324 | intros x e. 325 | assert (H: False = True) by exact (feq zero e). 326 | rewrite H. exact I. 327 | Show Proof. 328 | Qed. 329 | 330 | Goal true <> false. 331 | Proof. 332 | intros e. 333 | rewrite (feq (fun b: bool => if b then False else True) e). 334 | exact I. 335 | Show Proof. 336 | Qed. 337 | 338 | (*** Automation tactic "congruence" *) 339 | 340 | Goal true <> false. 341 | Proof. 342 | congruence. 343 | Qed. 344 | 345 | Goal forall x, S x <> 0. 346 | Proof. 347 | congruence. 348 | Qed. 349 | 350 | Goal forall x y, S x = S y -> x = y. 351 | Proof. 352 | congruence. 353 | Qed. 354 | 355 | Goal forall x y, S x = S y -> x = y. 356 | Proof. 357 | intros * [= e]. exact e. 358 | Qed. 359 | 360 | Goal forall x y, S x = S y -> x = y. 361 | Proof. 362 | intros * [= <-]. reflexivity. 363 | Qed. 364 | 365 | Goal forall X Y (f: X -> Y) x x', 366 | x = x' -> f x = f x'. 367 | Proof. 368 | congruence. 369 | Qed. 370 | 371 | Goal forall X Y (f g: X -> Y) x, 372 | f = g -> f x = g x. 373 | Proof. 374 | congruence. 375 | Qed. 376 | 377 | Goal forall X Y (x x': X) (y y': Y), 378 | (x,y) = (x',y') -> x = x' /\ y = y'. 379 | Proof. 380 | intros * e. split; congruence. 381 | Qed. 382 | 383 | (*** Abstract Presentation of Propositional Connectives *) 384 | 385 | Section Abstract_Prop_Connectives. 386 | Variable bot : Prop. 387 | Variable elim_bot : bot -> forall X: Prop, X. 388 | 389 | Fact bot_char : 390 | bot <-> forall Z, Z. 391 | Proof. 392 | split. 393 | - exact elim_bot. 394 | - intros H. apply H. 395 | Qed. 396 | 397 | Variable and : Prop -> Prop -> Prop. 398 | Variable intro_and : forall X Y, X -> Y -> and X Y. 399 | Variable elim_and : forall X Y Z , and X Y -> (X -> Y -> Z) -> Z. 400 | 401 | Fact and_char X Y : 402 | and X Y <-> forall Z, (X -> Y -> Z) -> Z. 403 | Proof. 404 | split. 405 | - intros H Z. apply elim_and, H. 406 | - intros H. apply H. apply intro_and. 407 | Qed. 408 | 409 | Variable or : Prop -> Prop -> Prop. 410 | Variable intro_or_l : forall X Y, X -> or X Y. 411 | Variable intro_or_r : forall X Y, Y -> or X Y. 412 | Variable elim_or : forall X Y Z , or X Y -> (X -> Z) -> (Y -> Z) -> Z. 413 | 414 | Fact or_char X Y : 415 | or X Y <-> forall Z, (X -> Z) -> (Y -> Z) -> Z. 416 | Proof. 417 | split. 418 | - intros H Z. apply elim_or, H. 419 | - intros H. apply H. apply intro_or_l. apply intro_or_r. 420 | Qed. 421 | 422 | End Abstract_Prop_Connectives. 423 | 424 | (*** Eta Conversion *) 425 | 426 | Goal forall X Y (f: X -> Y), 427 | (fun x => f x) = f. 428 | Proof. 429 | intros X Y f. 430 | cbv. 431 | reflexivity. 432 | Abort. 433 | 434 | 435 | Goal 436 | (fun x => 3 + x - 2) = S. 437 | Proof. 438 | cbv. 439 | reflexivity. 440 | Abort. 441 | 442 | Section Currying. 443 | Variables X Y Z : Type. 444 | Definition C : (X * Y -> Z) -> X -> Y -> Z 445 | := fun f x y => f (x,y). 446 | Definition U : (X -> Y -> Z) -> X * Y -> Z 447 | := fun f '(x,y) => f x y. 448 | Goal forall f x y, U (C f) (x,y) = f (x,y). 449 | Proof. 450 | cbv. 451 | reflexivity. 452 | Qed. 453 | Goal forall f x y, C (U f) x y = f x y. 454 | Proof. 455 | cbv. reflexivity. 456 | Qed. 457 | Goal forall f, C (U f) = f. 458 | Proof. 459 | cbv. reflexivity. 460 | (* Note the use of eta equivalence *) 461 | Qed. 462 | Goal forall f, U (C f) = f. 463 | Proof. 464 | cbv. 465 | Fail reflexivity. 466 | Abort. 467 | End Currying. 468 | -------------------------------------------------------------------------------- /rocq/numeral.v: -------------------------------------------------------------------------------- 1 | From Stdlib Require Import Arith Lia List. 2 | Definition dec (X: Type) := sum X (X -> False). 3 | Notation sig := sigT. 4 | Notation Sig := existT. 5 | Notation pi1 := projT1. 6 | Notation pi2 := projT2. 7 | Notation "'Sigma' x .. y , p" := 8 | (sig (fun x => .. (sig (fun y => p)) ..)) 9 | (at level 200, x binder, right associativity, 10 | format "'[' 'Sigma' '/ ' x .. y , '/ ' p ']'") 11 | : type_scope. 12 | 13 | Import ListNotations. 14 | Notation "x 'el' A" := (In x A) (at level 70). 15 | Notation "x 'nel' A" := (~ In x A) (at level 70). 16 | Fixpoint nrep {X} (A: list X) : Prop := 17 | match A with 18 | | [] => True 19 | | x::A => x nel A /\ nrep A 20 | end. 21 | Definition injective {X Y} (f: X -> Y) := 22 | forall x x', f x = f x' -> x = x'. 23 | Fact nrep_map X Y (f: X -> Y) A : 24 | injective f -> nrep A -> nrep (map f A). 25 | Proof. 26 | intros H1. 27 | induction A as [|x A IH]; cbn. 28 | - auto. 29 | - intros [H2 H3]. 30 | split. 2:{ apply IH, H3. } 31 | intros (x'&H4&H5)%in_map_iff. 32 | apply H1 in H4 as ->. 33 | auto. 34 | Qed. 35 | 36 | (** Inductive definition *) 37 | 38 | Inductive num : nat -> Type := 39 | | Zero: forall n, num (S n) 40 | | Up: forall {n}, num n -> num (S n). 41 | 42 | Check Zero 5. 43 | Check Up (Zero 4). 44 | Check Up (Up (Zero 3)). 45 | 46 | (** Eliminator *) 47 | 48 | Definition num_elim (p: forall n, num n -> Type) 49 | : (forall n, p (S n) (Zero n)) -> 50 | (forall n a, p n a -> p (S n) (Up a)) -> 51 | forall n a, p n a 52 | := fun e1 e2 => fix f n a := 53 | match a with 54 | | Zero n => e1 n 55 | | @Up n a => e2 n a (f n a) 56 | end. 57 | 58 | (** Predecessor and constructor laws *) 59 | 60 | Definition pre' 61 | : forall n, num n -> 62 | match n return Type with 63 | | 0 => False 64 | | S n' => option (num n') 65 | end 66 | := fun n a => match a with 67 | | Zero _ => None 68 | | Up a' => Some a' 69 | end. 70 | 71 | Definition pre 72 | : forall n, num (S n) -> option (num n) 73 | := fun n => pre' (S n). 74 | 75 | Goal forall n a, pre n (Up a) = Some a. 76 | reflexivity. 77 | Qed. 78 | Goal forall n, pre n (Zero n) = None. 79 | reflexivity. 80 | Qed. 81 | 82 | Fact num_disjoint n a : 83 | Zero n = Up a -> False. 84 | Proof. 85 | intros H % (f_equal (pre _)). cbn in H. 86 | change (match Some a with Some _ => False | None => True end). 87 | rewrite <- H. exact I. 88 | Qed. 89 | 90 | Fact Up_injective n (a b: num n) : 91 | Up a = Up b -> a = b. 92 | Proof. 93 | intros H % (f_equal (pre _)). cbn in H. 94 | change (match Some a with Some a' => a' | None => a end = b). 95 | rewrite H. reflexivity. 96 | Qed. 97 | 98 | Definition pre_hat 99 | : forall n, num (S (S n)) -> num (S n) 100 | := fun n a => match pre' _ a with 101 | | None => Zero _ 102 | | Some a => a 103 | end. 104 | Goal forall n a, pre_hat n (Up a) = a. 105 | Proof. 106 | reflexivity. 107 | Qed. 108 | Goal forall n, pre_hat n (Zero (S n)) = Zero n. 109 | Proof. 110 | reflexivity. 111 | Qed. 112 | 113 | (** Listing *) 114 | 115 | Fixpoint num_listing n : list (num n) := 116 | match n with 117 | | 0 => [] 118 | | S n' => Zero n' :: map Up (num_listing n') 119 | end. 120 | 121 | Compute num_listing 0. 122 | Compute num_listing 1. 123 | Compute num_listing 2. 124 | 125 | Goal forall n (a: num n), a el num_listing n. 126 | Proof. 127 | induction a as [n|n a IH]; cbn. 128 | - left. reflexivity. 129 | - right. apply in_map, IH. 130 | Qed. 131 | 132 | Goal forall n, length (num_listing n) = n. 133 | Proof. 134 | induction n as [|n IH]; cbn. 135 | - reflexivity. 136 | - f_equal. rewrite map_length. exact IH. 137 | Qed. 138 | 139 | Goal forall n, nrep (num_listing n). 140 | Proof. 141 | induction n as [|n IH]; cbn. 142 | - exact I. 143 | - split. 144 | + intros (a&H&_) % in_map_iff. 145 | symmetry in H. apply num_disjoint in H. easy. 146 | + apply nrep_map. 2:exact IH. exact (Up_injective n). 147 | Qed. 148 | 149 | (*** Inversion operator *) 150 | 151 | Definition num_inv 152 | : forall {n} (a: num n), 153 | match n return num n -> Type with 154 | | 0 => fun a => False 155 | | S n' => fun a => sum (a = Zero n') (Sigma a', a = Up a') 156 | end a. 157 | Proof. 158 | destruct a as [n|n a]. 159 | - left. reflexivity. 160 | - right. exists a. reflexivity. 161 | Defined. 162 | 163 | Eval cbn in fun n => num_inv (Zero n). 164 | Eval cbn in fun n (a: num n) => num_inv (Up a). 165 | 166 | (** Equality decider *) 167 | 168 | Definition num_eqdec : 169 | forall n (a1 a2: num n), dec (a1 = a2). 170 | Proof. 171 | induction a1 as [n|n a1 IH]; intros a2. 172 | all: destruct (num_inv a2) as [->|[a2' ->]]. 173 | - left. reflexivity. 174 | - right. intros [] % num_disjoint. 175 | - right. intros H. symmetry in H. eapply num_disjoint, H. 176 | - specialize (IH a2') as [[]|H]. 177 | + left. reflexivity. 178 | + right. contradict H. apply Up_injective, H. 179 | Defined. 180 | 181 | (** Enumerations *) 182 | 183 | Fact num0 : 184 | num 0 -> False. 185 | Proof. 186 | exact num_inv. 187 | Qed. 188 | Fact num_sum {n} : 189 | forall a: num (S n), (a = Zero n) + Sigma a', a = Up a'. 190 | Proof. 191 | exact num_inv. 192 | Qed. 193 | Fact num1 : 194 | forall a: num 1, a = Zero 0. 195 | Proof. 196 | intros a. 197 | destruct (num_sum a) as [->|[a' ->]]. 198 | - reflexivity. 199 | - exfalso. apply num0, a'. 200 | Qed. 201 | Fact num2 : 202 | forall a: num 2, sum (a = Zero 1) (a = Up (Zero 0)). 203 | Proof. 204 | intros a. 205 | destruct (num_sum a) as [->|[a' ->]]. 206 | - left. reflexivity. 207 | - right. f_equal. apply num1. 208 | Qed. 209 | 210 | (** Predecessor *) 211 | 212 | Definition P 213 | : forall n, num (S n) -> option (num n) 214 | := fun n a => match num_inv a with 215 | | inl _ => None 216 | | inr (Sig _ a' _) => Some a' 217 | end. 218 | 219 | Goal forall n, P n (Zero n) = None. 220 | Proof. 221 | reflexivity. 222 | Qed. 223 | Goal forall n a, P n (Up a) = Some a. 224 | Proof. 225 | reflexivity. 226 | Qed. 227 | 228 | (** Embedding into numbers *) 229 | 230 | Fixpoint N {n} (a: num n) : nat := 231 | match a with 232 | | Zero n => 0 233 | | Up a => S (N a) 234 | end. 235 | 236 | Compute N (Zero 3). 237 | Compute N (Up (Up (Zero 3))). 238 | 239 | Fact N_lt {n} (a: num n) : 240 | N a < n. 241 | Proof. 242 | induction a as [n|n a IH]; cbn. 243 | - lia. 244 | - lia. 245 | Qed. 246 | 247 | Fact N_Up n (a: num n) : 248 | N (Up a) = S (N a). 249 | Proof. 250 | destruct a as [n| n a]; reflexivity. 251 | Qed. 252 | 253 | Fact N_injective n (a b: num n) : 254 | N a = N b -> a = b. 255 | Proof. 256 | revert n a b. 257 | induction a as [|n a IH]; 258 | intros b; 259 | destruct (num_inv b) as [->|[a' ->]]; 260 | cbn. 261 | - easy. 262 | - intros [=]. 263 | - intros [=]. 264 | - intros [= H]. f_equal. apply IH, H. 265 | Qed. 266 | 267 | (** Lift *) 268 | 269 | Fixpoint lift {n} (a: num n) : num (S n) := 270 | match a with 271 | | Zero n => Zero (S n) 272 | | Up a => Up (lift a) 273 | end. 274 | 275 | Fact N_lift n (a: num n) : 276 | N (lift a) = N a. 277 | Proof. 278 | induction a as [n|n a IH]; cbn. 279 | - reflexivity. 280 | - f_equal. exact IH. 281 | Qed. 282 | 283 | Fact lift_injective n (a b: num n) : 284 | lift a = lift b -> a = b. 285 | Proof. 286 | intros H % (f_equal N). revert H. 287 | rewrite !N_lift. apply N_injective. 288 | Qed. 289 | 290 | (** Mapping numbers into numerals *) 291 | 292 | Fixpoint B k n : num (S n) := 293 | match k, n with 294 | | 0, n => Zero n 295 | | S k, 0 => Zero 0 296 | | S k, S n => Up (B k n) 297 | end. 298 | 299 | Compute B 3 3. 300 | Compute B 4 3. 301 | Compute B 2 3. 302 | Compute N (B 2 3). 303 | Compute B (N (Zero 5)) 5. 304 | Compute B (N (B 3 5)) 5. 305 | 306 | Fact NB_eq k n : 307 | k <= n -> N (B k n) = k. 308 | Proof. 309 | induction k as [|k IH] in n |-*; cbn. 310 | - easy. 311 | - destruct n as [|n]; cbn. 312 | + intros H. exfalso. lia. 313 | + intros H. f_equal. apply IH. lia. 314 | Qed. 315 | 316 | Fact BN_eq {n} (a: num (S n)) : 317 | B (N a) n = a. 318 | Proof. 319 | induction n as [|n IH]; 320 | destruct (num_inv a) as [->|[a' ->]]; 321 | cbn. 322 | - reflexivity. 323 | - exfalso. contradict (num_inv a'). 324 | - reflexivity. 325 | - f_equal. apply IH. 326 | Qed. 327 | 328 | (*** Recursive numerals *) 329 | 330 | Fixpoint fin (n: nat) : Type := 331 | match n with 332 | | 0 => False 333 | | S n' => option (fin n') 334 | end. 335 | 336 | (** Bijection *) 337 | 338 | Fixpoint num_fin {n} (a: num n) : fin n := 339 | match a with 340 | | Zero _ => None 341 | | Up a => Some (num_fin a) 342 | end. 343 | 344 | Fixpoint fin_num {n} (c: fin n) : num n := 345 | match n, c with 346 | | 0, c => match c with end 347 | | S n', None => Zero n' 348 | | S n', Some c => Up (fin_num c) 349 | end. 350 | 351 | Goal forall n (a: num n), 352 | fin_num (num_fin a) = a. 353 | Proof. 354 | induction a as [n|n a IH]; cbn. 355 | - reflexivity. 356 | - f_equal. exact IH. 357 | Qed. 358 | 359 | Goal forall n (c: fin n), 360 | num_fin (fin_num c) = c. 361 | Proof. 362 | induction n as [|n IH]. 363 | - intros []. 364 | - intros [c|]; cbn. 365 | + f_equal. apply IH. 366 | + reflexivity. 367 | Qed. 368 | 369 | Compute num_fin (Up (Up (Zero 5))). 370 | Compute fin_num (num_fin (Up (Up (Zero 5)))). 371 | 372 | (** Direct proofs of transported theorems *) 373 | 374 | Goal fin 0 -> False. 375 | Proof. 376 | easy. 377 | Qed. 378 | Goal forall a: fin 1, a = None. 379 | Proof. 380 | intros [a|]. 381 | - exfalso. exact a. 382 | - reflexivity. 383 | Qed. 384 | 385 | Definition fin_num_elim (p: forall n, fin n -> Type) 386 | : (forall n, p (S n) None) -> 387 | (forall n a, p n a -> p (S n) (Some a)) -> 388 | forall n a, p n a 389 | := fun e1 e2 => fix f n := 390 | match n with 391 | | 0 => fun a => match a with end 392 | | S n' => fun a => match a with 393 | | None => e1 n' 394 | | Some a' => e2 n' a' (f n' a') 395 | end 396 | end. 397 | 398 | Definition fin_num_inv 399 | : forall {n} (a: fin n), 400 | match n return fin n -> Type with 401 | | 0 => fun a => False 402 | | S n' => fun a => sum (a = None) (Sigma a', a = Some a') 403 | end a. 404 | Proof. 405 | intros n. 406 | destruct n as [|n]. 407 | - intros a. exact a. 408 | - destruct a as [a|]. 409 | + right. exists a. reflexivity. 410 | + left. reflexivity. 411 | Defined. 412 | 413 | (** Automation sometimes scales to indexed inductive types *) 414 | 415 | Print num_rect. 416 | 417 | Goal forall n a, Zero n = Up a -> False. 418 | Proof. 419 | Fail intros n [=]. 420 | congruence. 421 | Qed. 422 | 423 | Goal forall n (a b: num n), Up a = Up b -> a = b. 424 | Proof. 425 | Fail intros n [= H]. 426 | Fail congruence. 427 | Abort. 428 | -------------------------------------------------------------------------------- /rocq/pairing.v: -------------------------------------------------------------------------------- 1 | (*** MPCTT, Chapter Arithmetic Pairing *) 2 | 3 | (* This development is a pearl both mathematically 4 | and regarding the Rocq mechanization. We will 5 | demonstrate several advanced tactic uses. *) 6 | 7 | From Stdlib Require Import Lia. 8 | 9 | Implicit Types (n x y: nat) (a: nat * nat). 10 | 11 | Definition next a : nat * nat := 12 | match a with 13 | | (0,y) => (S y, 0) 14 | | (S x, y) => (x, S y) 15 | end. 16 | 17 | Fixpoint decode n : nat * nat := 18 | match n with 19 | | 0 => (0,0) 20 | | S n' => next (decode n') 21 | end. 22 | 23 | Fixpoint sum n : nat := 24 | match n with 25 | | 0 => 0 26 | | S n' => S n' + sum n' 27 | end. 28 | 29 | Definition encode '(x, y) : nat := 30 | sum (x + y) + y. 31 | 32 | Fact encode_next a : 33 | encode (next a) = S (encode a). 34 | Proof. 35 | destruct a as [[|x] y]; cbn. 36 | - replace (y + 0) with y; lia. 37 | - replace (x + S y) with (S (x + y)); cbn; lia. 38 | Qed. 39 | 40 | Opaque encode. 41 | (* Disables simplification of encode; 42 | doesn't affect reflexivity and easy. *) 43 | 44 | Fact encode_decode n : 45 | encode (decode n) = n. 46 | Proof. 47 | induction n as [|n IH]; cbn. 48 | - reflexivity. 49 | - rewrite encode_next. congruence. 50 | Qed. 51 | 52 | 53 | Fact encode_zero a : 54 | encode a = 0 -> a = (0,0). 55 | Proof. 56 | destruct a as [[|x] [|y]]; easy. 57 | Qed. 58 | 59 | Fact encode_eq_zero x n : 60 | encode (S x, 0) = S n -> encode (0, x) = n. 61 | Proof. 62 | change (S x, 0) with (next (0,x)). 63 | rewrite encode_next. congruence. 64 | Qed. 65 | 66 | Fact encode_eq_S x y n : 67 | encode (x, S y) = S n -> encode (S x, y) = n. 68 | Proof. 69 | change (x, S y) with (next (S x, y)). 70 | rewrite encode_next. congruence. 71 | Qed. 72 | 73 | Fact decode_encode a : 74 | decode (encode a) = a. 75 | Proof. 76 | revert a. 77 | enough (forall n a, encode a = n -> decode n = a) by eauto. 78 | induction n as [|n IH]; cbn. 79 | - intros a H%encode_zero. easy. 80 | - destruct a as [x [|y]]. 81 | + destruct x. easy. 82 | intros <-%encode_eq_zero. 83 | rewrite (IH (0,x)); reflexivity. 84 | + intros <-%encode_eq_S. 85 | rewrite (IH (S x,y)); reflexivity. 86 | Qed. 87 | 88 | Fact Gauss n : 89 | 2 * sum n = n * S n. 90 | Proof. 91 | induction n as [|n IH]; cbn; lia. 92 | Qed. 93 | -------------------------------------------------------------------------------- /rocq/pat.v: -------------------------------------------------------------------------------- 1 | (*** MPCTT, Chapter 3 *) 2 | 3 | (* We show how Rocq accommodates propositions and proofs. *) 4 | 5 | Print False. 6 | Definition exfalso : forall X: Prop, False -> X := 7 | fun X a => match a with end. 8 | Check exfalso. 9 | Notation "~ X" := (X -> False) (at level 75, right associativity) : type_scope. 10 | 11 | Section Demo. 12 | Variables X Y Z: Prop. 13 | 14 | Goal X -> ~X -> False. 15 | Proof. 16 | exact (fun x f => f x). 17 | Show Proof. 18 | Qed. 19 | 20 | Goal X -> ~X -> False. 21 | Proof. 22 | refine (fun x f => _). 23 | exact (f x). 24 | Show Proof. 25 | Qed. 26 | 27 | Goal X -> ~X -> False. 28 | Proof. 29 | intros x f. apply f. exact x. 30 | Show Proof. 31 | Qed. 32 | 33 | Goal X -> ~X -> False. 34 | Proof. 35 | auto. 36 | Show Proof. 37 | Qed. 38 | 39 | Goal X -> ~X -> Y. 40 | Proof. 41 | intros x f. 42 | apply exfalso. 43 | exact (f x). 44 | Show Proof. 45 | Qed. 46 | 47 | Goal X -> ~X -> Y. 48 | Proof. 49 | intros x f. 50 | exfalso. (* Rocq's exfalso tactic *) 51 | exact (f x). 52 | Show Proof. 53 | Qed. 54 | 55 | Goal (X -> ~X) -> (~X -> X) -> False. 56 | Proof. 57 | refine (fun f g => _). 58 | refine (f _ _). 59 | - refine (g (fun x => _)). exact (f x x). 60 | - refine (g (fun x => _)). exact (f x x). 61 | Show Proof. 62 | Qed. 63 | 64 | Goal (X -> ~X) -> (~X -> X) -> False. 65 | Proof. 66 | refine (fun f g => _). 67 | unshelve refine (let x:X := _ in _). 68 | - refine (g (fun x => _)). exact (f x x). 69 | - exact (f x x). 70 | Show Proof. 71 | Qed. 72 | 73 | Goal (X -> ~X) -> (~X -> X) -> False. 74 | Proof. 75 | intros f g. 76 | assert (x: X). 77 | - apply g. intros x. exact (f x x). 78 | - exact (f x x). 79 | Show Proof. 80 | Qed. 81 | 82 | 83 | Goal ~ (X <-> ~X). 84 | Proof. 85 | refine (fun a => match a with conj f g => _ end). 86 | unshelve refine (let x:X := _ in _). 87 | - refine (g (fun x => _)). exact (f x x). 88 | - exact (f x x). 89 | Show Proof. 90 | Qed. 91 | 92 | Goal X -> ~ ~ X. 93 | Proof. 94 | intros x f. exact (f x). 95 | Qed. 96 | 97 | Goal ~ ~X -> (X -> ~X) -> False. 98 | Proof. 99 | intros f g. 100 | apply f. intros x. 101 | exact (g x x). 102 | Show Proof. 103 | Qed. 104 | 105 | Goal ~ ~ (~ ~ X -> X). 106 | Proof. 107 | intros f. apply f. intros g. 108 | exfalso. apply g. intros x. 109 | apply f. intros _. exact x. 110 | Show Proof. 111 | Qed. 112 | 113 | Goal ~ ~ (~ ~ X -> X). 114 | Proof. 115 | tauto. 116 | Qed. 117 | 118 | End Demo. 119 | 120 | Fact mobility1 X (P: Prop) (p: X -> Prop) : 121 | (forall x, P -> p x) -> (P -> forall x, p x). 122 | Proof. 123 | intros f H x. apply f. apply H. 124 | Show Proof. 125 | Qed. 126 | 127 | Fact mobility2 X (P: Prop) (p: X -> Prop) : 128 | (P -> forall x, p x) -> (forall x, P -> p x). 129 | Proof. 130 | intros f x H. apply f. apply H. 131 | Show Proof. 132 | Qed. 133 | 134 | Fact exchange X Y (p: X -> Y -> Prop) : 135 | (forall x y, p x y) -> (forall y x, p x y). 136 | Proof. 137 | intros f y x. exact (f x y). 138 | Qed. 139 | 140 | Fact exchange_alpha X Y (p: X -> Y -> Prop) : 141 | (forall x y, p x y) -> (forall x' y', p x' y'). 142 | Proof. 143 | exact (fun a => a). 144 | Qed. 145 | 146 | (** We use Rocq's conjunctions, and disjunctions. *) 147 | 148 | Locate "/\". 149 | Print and. 150 | Locate "\/". 151 | Print or. 152 | (* proof constructors of or have overloaded implicit arguments *) 153 | 154 | Definition match_and {X Y Z: Prop} 155 | : X /\ Y -> (X -> Y -> Z) -> Z 156 | := fun a e => match a with conj x y => e x y end. 157 | 158 | Definition match_or {X Y Z: Prop} 159 | : X \/ Y -> (X -> Z) -> (Y -> Z) -> Z 160 | := fun a e1 e2 => match a with or_introl x => e1 x | or_intror y => e2 y end. 161 | 162 | Section Demo. 163 | Variables X Y Z: Prop. 164 | 165 | Goal X /\ Y -> Y /\ X. 166 | Proof. 167 | intros H. 168 | apply (match_and H). 169 | intros x y. exact (conj y x). 170 | Show Proof. 171 | Qed. 172 | 173 | Goal X /\ Y -> Y /\ X. 174 | Proof. 175 | intros [x y]. 176 | split. exact y. exact x. 177 | Show Proof. 178 | Qed. 179 | 180 | Goal X \/ Y -> Y \/ X. 181 | Proof. 182 | intros H. 183 | apply (match_or H). 184 | - intros x. exact (or_intror x). 185 | - intros y. exact (or_introl y). 186 | Show Proof. 187 | Qed. 188 | 189 | Goal X \/ Y -> Y \/ X. 190 | Proof. 191 | intros [x|y]. 192 | - right. exact x. 193 | - left. exact y. 194 | Show Proof. 195 | Qed. 196 | 197 | Goal X \/ (Y /\ Z) -> (X \/ Y) /\ (X \/ Z). 198 | Proof. 199 | intros H. 200 | apply (match_or H). 201 | - intros x. apply conj. 202 | + exact (or_introl x). 203 | + exact (or_introl x). 204 | - intros H1. 205 | apply (match_and H1). 206 | intros y z. apply conj. 207 | + exact (or_intror y). 208 | + exact (or_intror z). 209 | Qed. 210 | 211 | Goal X \/ (Y /\ Z) -> (X \/ Y) /\ (X \/ Z). 212 | Proof. 213 | intros [x | [y z]]. 214 | - split. 215 | + left. exact x. 216 | + left. exact x. 217 | - split. 218 | + right. exact y. 219 | + right. exact z. 220 | Qed. 221 | 222 | Goal X \/ (Y /\ Z) -> (X \/ Y) /\ (X \/ Z). 223 | Proof. 224 | tauto. 225 | Qed. 226 | End Demo. 227 | 228 | Notation "A <-> B" := ((A -> B) /\ (B -> A)) 229 | (at level 95, no associativity) : type_scope. 230 | 231 | Section Demo. 232 | Variables X Y Z: Prop. 233 | 234 | Goal ~(X <-> ~X). 235 | Proof. 236 | intros H. apply (match_and H). 237 | intros f g. 238 | assert (x: X). 239 | - apply g. intros x. exact (f x x). 240 | - exact (f x x). 241 | Show Proof. 242 | Qed. 243 | 244 | Goal ~(X <-> ~X). 245 | Proof. 246 | intros [f g]. 247 | Show Proof. (* unnecessary beta redex *) 248 | assert (x:X). 249 | - apply g. intros x. exact (f x x). 250 | - exact (f x x). 251 | Qed. 252 | 253 | Fact Russell : 254 | ~(X <-> ~X). 255 | Proof. 256 | tauto. 257 | Qed. 258 | 259 | Fact deMorgan_or : 260 | ~ (X \/ Y) <-> ~ X /\ ~ Y. 261 | Proof. 262 | split. 263 | - intros f. split. 264 | + contradict f. left. exact f. 265 | + contradict f. right. exact f. 266 | - intros [f g] [x|y]. 267 | + auto. 268 | + auto. 269 | Qed. 270 | 271 | Goal ~ ~(X -> Y) <-> (~ ~X -> ~ ~Y). 272 | Proof. 273 | split. 274 | - intros f g h. 275 | apply g. intros x. apply f. intros f'. 276 | exact (h (f' x)). 277 | - intros f g. apply g. intros x. exfalso. 278 | apply f. 279 | + intros h. auto. 280 | + intros y. apply g. auto. 281 | Qed. 282 | End Demo. 283 | 284 | (** Assumed variables now appear as leading arguments *) 285 | 286 | Check Russell. 287 | Check deMorgan_or. 288 | 289 | Goal forall X (p q: X -> Prop), 290 | (forall x, p x <-> q x) -> (forall x, q x) -> forall x, p x. 291 | Proof. 292 | intros X p q f g x. 293 | destruct (f x) as [_ h]. 294 | exact (h (g x)). 295 | Qed. 296 | 297 | Goal forall X (p q: X -> Prop), 298 | (forall x, p x <-> q x) -> (forall x, q x) -> forall x, p x. 299 | Proof. 300 | intros X p q f g x. 301 | apply f. (* matchin direction is applied *) 302 | exact (g x). 303 | Qed. 304 | 305 | Goal forall X (p q: X -> Prop), 306 | (forall x, p x <-> q x) -> (forall x, q x) -> forall x, p x. 307 | Proof. 308 | intros X p q f g x. 309 | refine (match f x with conj _ h => _ end). 310 | exact (h (g x)). 311 | Show Proof. 312 | Qed. 313 | 314 | (** Impredicative characterizations *) 315 | 316 | Goal False <-> forall Z: Prop, Z. 317 | Proof. 318 | split. 319 | - intros []. 320 | - intros f. exact (f False). 321 | Show Proof. 322 | Qed. 323 | 324 | Goal forall X Y: Prop, 325 | X /\ Y <-> forall Z: Prop, (X -> Y -> Z) -> Z. 326 | Proof. 327 | intros X Y. 328 | split. 329 | - intros [x y] Z f. exact (f x y). 330 | - intros f. eapply f. apply conj. 331 | (* eapply is smarter than apply *) 332 | Show Proof. 333 | Qed. 334 | 335 | Goal forall X Y: Prop, 336 | X \/ Y <-> forall Z: Prop, (X -> Z) -> (Y -> Z) -> Z. 337 | Proof. 338 | intros X Y. 339 | split. 340 | - intros [x|y] Z f g. 341 | + exact (f x). 342 | + exact (g y). 343 | - intros f. eapply f. 344 | + eapply or_introl. 345 | + eapply or_intror. 346 | Show Proof. 347 | Qed. 348 | 349 | (** Excluded Middle *) 350 | 351 | Fact XM_DN : 352 | (forall X: Prop, X \/ ~X) <-> (forall X: Prop, ~ ~X -> X). 353 | Proof. 354 | split; intros H X. 355 | - destruct (H X) as [x|H1]. 356 | + auto. 357 | + intros H2. exfalso. apply H2. exact H1. 358 | - apply H. intros H1. apply H1. 359 | right. contradict H1. auto. 360 | Qed. 361 | 362 | 363 | Fact XM_CDM : 364 | (forall X: Prop, X \/ ~X) -> forall X Y : Prop, ~(X /\ Y) -> ~X \/ ~Y. 365 | Proof. 366 | intros H X Y f. 367 | specialize (H X) as [x|H]. 368 | - right. intros y. exact (f (conj x y)). 369 | - left. exact H. 370 | Show Proof. 371 | Qed. 372 | -------------------------------------------------------------------------------- /rocq/post.v: -------------------------------------------------------------------------------- 1 | (** Post Correspondence Problem *) 2 | 3 | From Stdlib Require Import List. 4 | Import ListNotations. 5 | Notation "x 'el' A" := (In x A) (at level 70). 6 | 7 | Inductive char := a | b. 8 | Definition string : Type := list char. 9 | Definition card : Type := string * string. 10 | Definition deck : Type := list card. 11 | 12 | Inductive post (D: deck) : string -> string -> Prop := 13 | | post1 : forall A B, (A,B) el D -> post D A B 14 | | post2 : forall A B A' B', (A,B) el D -> post D A' B' -> post D (A ++ A') (B ++ B'). 15 | 16 | Definition D : deck := [([a], []); ([b],[a]); ([], [b;b])]. 17 | 18 | Goal post D [b;b;a;a] [b;b;a;a]. 19 | Proof. 20 | refine (post2 _ [] [b;b] [b;b;a;a] [a;a] _ _). 21 | { cbn; auto. } 22 | refine (post2 _ [b] [a] [b;a;a] [a] _ _). 23 | { cbn; auto. } 24 | refine (post2 _ [b] [a] [a;a] [] _ _). 25 | { cbn; auto. } 26 | refine (post2 _ [a] [] [a] [] _ _). 27 | { cbn; auto. } 28 | refine (post1 _ [a] [] _). 29 | { cbn; auto. } 30 | Qed. 31 | 32 | -------------------------------------------------------------------------------- /rocq/provability.v: -------------------------------------------------------------------------------- 1 | Section Provable. 2 | Implicit Types X Y Z: Prop. 3 | Variable provable : Prop -> Prop. 4 | Definition disprovable X := provable(~X). 5 | Definition consistent X := ~provable(~X). 6 | Definition independent X := ~provable X /\ ~provable (~X). 7 | 8 | Variable MP : forall X Y, provable (X -> Y) -> provable X -> provable Y. 9 | Variable PK : forall X Y, provable (Y -> X -> Y). 10 | Variable PS : forall X Y Z, provable ((X -> Y -> Z) -> (X -> Y) -> X -> Z). 11 | 12 | Fact LK X Y : 13 | provable X -> provable (Y -> X). 14 | Proof. 15 | intros H. eapply MP. apply PK. exact H. 16 | Qed. 17 | 18 | Fact LS X Y Z : 19 | provable (X -> Y -> Z) -> provable (X -> Y) -> provable (X -> Z). 20 | Proof. 21 | intros H1 H2. eapply MP. 2:exact H2. 22 | eapply MP. 2:exact H1. apply PS. 23 | Qed. 24 | 25 | Fact LI X : 26 | provable (X -> X). 27 | Proof. 28 | apply LS with (X -> X); apply PK. 29 | Qed. 30 | 31 | Fact LC X Y Z : 32 | provable (X -> Y) -> provable ((Y -> Z) -> X -> Z). 33 | Proof. 34 | intros H. 35 | eapply LS. 36 | - eapply LS. 37 | + apply LK, PS. 38 | + eapply LS. 39 | * apply LK, PK. 40 | * apply LI. 41 | - apply LK, H. 42 | Qed. 43 | 44 | Fact transport1 X Y : 45 | provable (X -> Y) -> ~provable Y -> ~provable X. 46 | Proof. 47 | intros H H1. contradict H1. revert H H1. apply MP. 48 | Qed. 49 | 50 | Fact transport2 X Y : 51 | provable (X -> Y) -> consistent X -> consistent Y. 52 | Proof. 53 | intros H. apply transport1. apply LC, H. 54 | Qed. 55 | 56 | Fact transport3 X Y : 57 | provable (X -> Y) -> provable (Y -> X) -> 58 | independent X -> independent Y. 59 | Proof. 60 | intros H1 H2 [H3 H4]. split. 61 | - revert H3. apply transport1. exact H2. 62 | - revert H4. apply transport2. exact H1. 63 | Qed. 64 | 65 | Fact sandwich X Y Z : 66 | provable (X -> Y) -> provable (Y -> Z) -> 67 | consistent X -> ~provable Z -> 68 | independent Y. 69 | Proof. 70 | intros H1 H2 H3 H4. split. 71 | - revert H2 H4. apply transport1. 72 | - revert H1 H3. apply transport2. 73 | Qed. 74 | 75 | Fact sandwich' Y : 76 | independent Y <-> exists X Z, 77 | provable (X -> Y) /\ provable (Y -> Z) /\ 78 | consistent X /\ ~provable Z. 79 | Proof. 80 | split. 81 | - intros H. exists Y, Y. 82 | split. apply LI. split. apply LI. split; apply H. 83 | - intros (X&Z&H1&H2&H3&H4). eapply sandwich; eassumption. 84 | Qed. 85 | 86 | Fact cons1 : 87 | ~provable False -> consistent (~False). 88 | Proof. 89 | cbv. intros H. contradict H. 90 | eapply MP. exact H. apply LI. 91 | Qed. 92 | 93 | Fact cons2 X Y : 94 | consistent X -> ~ provable False. 95 | Proof. 96 | cbv. intros H. contradict H. apply LK, H. 97 | Qed. 98 | 99 | Fact cons3 : 100 | ~provable False <-> forall X, provable X -> consistent X. 101 | Proof. 102 | cbv. split. 103 | - intros H X H1. contradict H. 104 | eapply MP. exact H. exact H1. 105 | - intros H H1. eapply H. exact H1. apply LI. 106 | Qed. 107 | 108 | Fact cons4 : 109 | ~provable False <-> forall P, disprovable P -> ~provable P. 110 | Proof. 111 | unfold disprovable. split. 112 | - intros H1 P H2. contradict H1. apply (MP _ _ H2 H1). 113 | - intros H. specialize (H False). apply H. apply LI. 114 | Qed. 115 | 116 | End Provable. 117 | -------------------------------------------------------------------------------- /rocq/regexp.v: -------------------------------------------------------------------------------- 1 | (*** Regular Expressions *) 2 | From Stdlib Require Import Arith Lia List. 3 | Import ListNotations. 4 | Definition dec X := sum X (X -> False). 5 | Definition iffT (X Y: Type) : Type := (X -> Y) * (Y -> X). 6 | Notation "X <=> Y" := (iffT X Y) (at level 95, no associativity). 7 | Notation "'Sigma' x .. y , p" := 8 | (sigT (fun x => .. (sigT (fun y => p)) ..)) 9 | (at level 200, x binder, right associativity, 10 | format "'[' 'Sigma' '/ ' x .. y , '/ ' p ']'") 11 | : type_scope. 12 | 13 | Inductive exp: Type := 14 | | Sym (x: nat) 15 | | Nul 16 | | One 17 | | Add (s t: exp) 18 | | Mul (s t: exp) 19 | | Star (s: exp). 20 | 21 | Implicit Types s t : exp. 22 | Implicit Types A B : list nat. 23 | 24 | Reserved Notation "A |- s" (at level 70). 25 | Inductive sat : list nat -> exp -> Type := 26 | | SSym x : [x] |- Sym x 27 | | SOne : [] |- One 28 | | SAddL A s t : A |- s -> A |- Add s t 29 | | SAddR A s t : A |- t -> A |- Add s t 30 | | SMul A B s t : A |- s -> B |- t -> A ++ B |- Mul s t 31 | | SStar1 s : [] |- Star s 32 | | SStar2 A B s : A |- s -> B |- Star s -> A ++ B |- Star s 33 | where "A |- s" := (sat A s). 34 | 35 | Definition sat_inv 36 | : forall {A s}, A |- s -> 37 | match s return Type with 38 | | Sym x => A = [x] 39 | | Nul => False 40 | | One => A = [] 41 | | Add s t => sum (A |- s) (A |- t) 42 | | Mul s t => Sigma A1 A2, prod (A = A1++A2) (prod (A1 |- s ) (A2 |- t)) 43 | | Star s => sum (A = []) 44 | (Sigma A1 A2, prod (A = A1++A2) 45 | (prod (A1 |- s) (A2 |- Star s))) 46 | end. 47 | Proof. 48 | intros A s a. 49 | destruct a as [x| |A s t H |A s t H |A B s t Hs Ht|s|A B s Hs Ht]. 50 | - reflexivity. 51 | - reflexivity. 52 | - left. exact H. 53 | - right. exact H. 54 | - exists A, B. easy. 55 | - left. reflexivity. 56 | - right. exists A, B. easy. 57 | Defined. 58 | 59 | (*** Facts about Star *) 60 | 61 | Fact star1 A s : 62 | A |- s -> A |- Star s. 63 | Proof. 64 | intros H. rewrite <-(app_nil_r A). 65 | constructor. exact H. constructor. 66 | Qed. 67 | 68 | Fact star2 A B s : 69 | A |- Star s -> B |- Star s -> A ++ B |- Star s. 70 | Proof. 71 | revert A s. 72 | enough (forall A u, A |- u -> 73 | match u return Type with 74 | | Star s => B |- u -> A ++ B |- u 75 | | _ => True 76 | end) as H. 77 | { intros A s H1. apply (H A (Star s) H1). } 78 | induction 1 as [| | | | | |A1 A2 s Hs _ _ IHt]; 79 | try exact I. 80 | - easy. 81 | - intros H. 82 | rewrite <- app_assoc. constructor; auto. 83 | Qed. 84 | 85 | Fact equiv1 A s t : 86 | A |- Mul (Star s) (Star s) <=> A |- Star s. 87 | Proof. 88 | split. 89 | - intros (A1&A2&->&H1&H2) % sat_inv. 90 | apply star2; assumption. 91 | - intros H. change A with ([]++A). 92 | constructor. constructor. exact H. 93 | Qed. 94 | 95 | Fact star3 A s : 96 | A |- Star (Star s) -> A |- Star s. 97 | Proof. 98 | revert A s. 99 | enough (forall A u, A |- u -> 100 | match u return Type with 101 | | Star (Star s) => A |- Star s 102 | | _ => True 103 | end) as H. 104 | { intros A s H1. apply (H A (Star (Star s)) H1). } 105 | induction 1 as [ | | | | | |A1 A2 s Hs _ _ IHt]; 106 | try exact I. 107 | all: destruct s as [| | | | |s]; try exact I. 108 | - constructor. 109 | - apply star2. exact Hs. exact IHt. 110 | Qed. 111 | 112 | Fact equiv2 A s : 113 | A |- Star (Star s) <=> A |- Star s. 114 | Proof. 115 | split. 116 | - apply star3. 117 | - apply star1. 118 | Qed. 119 | 120 | (*** Derivatives *) 121 | 122 | Definition eps_dec s : 123 | dec ([] |- s). 124 | Proof. 125 | induction s as [x| | |s IHs t IHt|s IHs t IHt|s IH]. 126 | - right. intros H. discriminate (sat_inv H). 127 | - right. intros H. apply (sat_inv H). 128 | - left. constructor. 129 | - destruct IHs as [IHs|IHs]. 130 | + left. apply SAddL, IHs. 131 | + destruct IHt as [IHt|IHt]. 132 | * left. apply SAddR, IHt. 133 | * right. intros H. destruct (sat_inv H); easy. 134 | - destruct IHs as [IHs|IHs]. 135 | + destruct IHt as [IHt|IHt]. 136 | * left. change ([] ++ [] |- Mul s t). 137 | constructor; assumption. 138 | * right. intros H. 139 | destruct (sat_inv H) as (A1&A2&H1&_&H2). 140 | destruct A1, A2; try discriminate H1. easy. 141 | + right. intros H. 142 | destruct (sat_inv H) as (A1&A2&H1&H2&_). 143 | destruct A1, A2; try discriminate H1. easy. 144 | - left. constructor. 145 | Defined. 146 | 147 | Fixpoint deriv x s : exp := 148 | match s with 149 | | Sym y => if Nat.eq_dec x y then One else Nul 150 | | Nul => Nul 151 | | One => Nul 152 | | Add s t => Add (deriv x s) (deriv x t) 153 | | Mul s t => if eps_dec s 154 | then Add (Mul (deriv x s) t) (deriv x t) 155 | else Mul (deriv x s) t 156 | | Star s => Mul (deriv x s) (Star s) 157 | end. 158 | 159 | Lemma star4 x A s : 160 | x::A |- Star s -> Sigma A1 A2, prod (A = A1++A2) (prod (x::A1 |- s) (A2 |- Star s)). 161 | Proof. 162 | revert x A s. 163 | enough (forall A s, A |- s -> match A, s return Type with 164 | | x::A, Star s => Sigma A1 A2, prod (A = A1++A2) 165 | (prod (x::A1 |- s) 166 | (A2 |- Star s)) 167 | | _, _ => True 168 | end) as H. 169 | { intros x A s H1. apply (H _ _ H1). } 170 | induction 1 as [ | | | | | s |A B s Hs _ Ht IHt]; try exact I. 171 | 1-2: destruct A; exact I. 172 | - destruct (A ++ B); exact I. 173 | - destruct A as [|x A]; cbn. 174 | + exact IHt. 175 | + exists A, B. easy. 176 | Qed. 177 | 178 | Fact deriv_correct x A s : 179 | x::A |- s <=> A |- deriv x s. 180 | Proof. 181 | induction s as 182 | [y| | |s IHs t IHt|s IHs t IHt|s IH] in x, A |-*; 183 | cbn; split. 184 | - intros [= <- ->]%sat_inv. 185 | destruct Nat.eq_dec as [_|H1]. constructor. easy. 186 | - destruct Nat.eq_dec as [<-|H1]; intros H%sat_inv; revert H; cbn. 187 | + intros ->. constructor. 188 | + easy. 189 | - intros []%sat_inv. 190 | - intros []%sat_inv. 191 | - intros [=]%sat_inv. 192 | - intros []%sat_inv. 193 | - intros [H|H]%sat_inv. 194 | + apply SAddL, IHs, H. 195 | + apply SAddR, IHt, H. 196 | - intros [H|H]%sat_inv. 197 | + apply SAddL, IHs, H. 198 | + apply SAddR, IHt, H. 199 | - intros (A1&A2&H1&H2&H3)%sat_inv. 200 | revert H1. destruct A1 as [|y A1]; cbn. 201 | + intros <-. destruct eps_dec as [H|H]. 202 | * apply SAddR, IHt, H3. 203 | * easy. 204 | + intros [= <- ->]. destruct eps_dec as [H|H]. 205 | * apply SAddL. constructor. apply IHs, H2. exact H3. 206 | * constructor. apply IHs, H2. exact H3. 207 | - destruct eps_dec as [H|H]. 208 | + intros [H1|H1]%sat_inv. 209 | * revert H1. intros (A1&A2&->&H2&H3)%sat_inv. 210 | apply (SMul (x::A1)). apply IHs, H2. exact H3. 211 | * apply (SMul []). exact H. apply IHt, H1. 212 | + intros (A1&A2&->&H2&H3)%sat_inv. 213 | apply (SMul (x::A1)). apply IHs, H2. exact H3. 214 | - intros (A1&A2&->&H2&H3)%star4. 215 | apply SMul. apply IH, H2. exact H3. 216 | - intros (A1&A2&->&H2&H3)%sat_inv. 217 | apply (SStar2 (x::A1)). apply IH, H2. exact H3. 218 | Qed. 219 | 220 | Definition sat_dec A s : 221 | dec (A |- s). 222 | Proof. 223 | induction A as [|x A IH] in s |-*. 224 | - apply eps_dec. 225 | - specialize (IH (deriv x s)) as [H1|H1]. 226 | + left. apply deriv_correct, H1. 227 | + right. contradict H1. apply deriv_correct, H1. 228 | Defined. 229 | 230 | 231 | (*** Certifying Solver *) 232 | 233 | Definition solve s : 234 | (Sigma A, A |- s) + (forall A, A |- s -> False). 235 | Proof. 236 | induction s as [x| | |s IHs t IHt|s IHs t IHt|s _]. 237 | - left. exists [x]. constructor. 238 | - right. intros A H. apply (sat_inv H). 239 | - left. exists []. constructor. 240 | - destruct IHs as [[A IHs]|IHs]. 241 | + left. exists A. apply SAddL, IHs. 242 | + destruct IHt as [[B IHt]|IHt]. 243 | * left. exists B. apply SAddR, IHt. 244 | * right. intros A [H|H] % sat_inv. 245 | -- eapply IHs, H. 246 | -- eapply IHt, H. 247 | - destruct IHs as [[A IHs]|IHs]. 248 | + destruct IHt as [[B IHt]|IHt]. 249 | * left. exists (A ++ B). constructor; assumption. 250 | * right. intros C H. 251 | destruct (sat_inv H) as (_&B&_&_&H3). 252 | eapply IHt, H3. 253 | + right. intros C H. 254 | destruct (sat_inv H) as (A'&_&_&H2&_). 255 | eapply IHs, H2. 256 | - left. exists []. constructor. 257 | Defined. 258 | 259 | -------------------------------------------------------------------------------- /rocq/sdec.v: -------------------------------------------------------------------------------- 1 | (** Preliminaries, have been done before *) 2 | 3 | From Stdlib Require Import Bool. 4 | 5 | Definition iffT (X Y: Type) : Type := 6 | (X -> Y) * (Y -> X). 7 | 8 | Notation "X <=> Y" := (iffT X Y) (at level 95, no associativity). 9 | 10 | Notation dec X := ({ X } + { ~X }). 11 | 12 | Definition bdec {X} (p: X -> Prop) (f: X -> bool) : Prop := 13 | forall x, p x <-> f x = true. 14 | 15 | Fact dec_transport X Y : 16 | (X <-> Y) -> dec X -> dec Y. 17 | Proof. 18 | intros H1 [H2|H2]. 19 | - left. apply H1, H2. 20 | - right. contradict H2. apply H1, H2. 21 | Qed. 22 | 23 | Lemma cdec_bdec X (p: X -> Prop) : 24 | sig (bdec p) <=> forall x, dec (p x). 25 | Proof. 26 | split. 27 | - intros [f H] x. destruct (f x) eqn:H1. 28 | + left. apply H, H1. 29 | + right. intros H2%H. congruence. 30 | - intros H. 31 | exists (fun x => if H x then true else false). 32 | intros x. destruct (H x) as [H1|H1]. 33 | tauto. intuition; discriminate. 34 | Qed. 35 | 36 | (** Witness Operator *) 37 | 38 | Section WO. 39 | Variable f: nat -> bool. 40 | 41 | Inductive G (n: nat) : Prop := 42 | | GI : (f n = false -> G (S n)) -> G n. 43 | 44 | Lemma G_sig n : 45 | G n -> { k | f k = true }. 46 | Proof. 47 | induction 1 as [n _ IH]. 48 | destruct (f n) eqn:H. 49 | - exists n. exact H. 50 | - apply IH. reflexivity. 51 | Qed. 52 | 53 | Lemma G_zero n : 54 | G n -> G 0. 55 | Proof. 56 | induction n as [|n IH]. 57 | - intros H. exact H. 58 | - intros H. apply IH. constructor. intros _. exact H. 59 | Qed. 60 | 61 | Theorem wo : 62 | (exists n, f n = true) -> { n | f n = true }. 63 | Proof. 64 | intros H. apply (G_sig 0). 65 | destruct H as [n H]. (* witness needed *) 66 | apply (G_zero n). 67 | constructor. rewrite H. discriminate. 68 | Qed. 69 | End WO. 70 | 71 | (** Now we start with semi-decidability *) 72 | 73 | Definition tsat (f: nat -> bool) : Prop := 74 | exists n, f n = true. 75 | 76 | Definition sdec (X: Prop) : Type := 77 | { f | X <-> tsat f }. 78 | 79 | Goal forall X, dec X -> sdec X. 80 | Proof. 81 | intros X [H|H]. 82 | - exists (fun n => true). intuition. exists 0. reflexivity. 83 | - exists (fun n => false). split. tauto. intros [_ H1]. discriminate. 84 | Qed. 85 | 86 | Fact sdec_transport X Y : 87 | (X <-> Y) -> sdec X -> sdec Y. 88 | Proof. 89 | intros H1 [f H2]. exists f. tauto. 90 | Qed. 91 | 92 | Goal forall X (p: X -> Prop), 93 | (forall x, sdec (p x)) <=> { F | forall x, p x <-> tsat (F x) }. 94 | Proof. 95 | split. 96 | - intros H. 97 | exists (fun x => proj1_sig (H x)). 98 | intros x. destruct (H x) as [f H1]. cbn. exact H1. 99 | - intros [F H] x. exists (F x). apply H. 100 | Qed. 101 | 102 | (** Semi-decisions closed under conjunction and disjunction *) 103 | 104 | Fact sdec_or X Y : 105 | sdec X -> sdec Y -> sdec (X \/ Y). 106 | Proof. 107 | intros [f Hf] [g Hg]. exists (fun n => f n || g n). 108 | split. 109 | - intros [[n H1]%Hf|[n H1]%Hg]; exists n; rewrite H1. 110 | + reflexivity. 111 | + apply orb_true_r. 112 | - intros [n [H|H]%orb_true_elim]. 113 | + left. apply Hf. exists n. exact H. 114 | + right. apply Hg. exists n. exact H. 115 | Qed. 116 | 117 | Section Sdec_and. 118 | Variables (F: nat -> nat -> nat) (pi1: nat -> nat) (pi2: nat -> nat). 119 | Variable pi1_eq : forall x y, pi1 (F x y) = x. 120 | Variable pi2_eq : forall x y, pi2 (F x y) = y. 121 | 122 | Fact sdec_and X Y : 123 | sdec X -> sdec Y -> sdec (X /\ Y). 124 | Proof. 125 | intros [f Hf] [g Hg]. exists (fun n => f (pi1 n) && g (pi2 n)). 126 | split. 127 | - intros [x y]. 128 | apply Hf in x as [nx Hx]. 129 | apply Hg in y as [ny Hy]. 130 | exists (F nx ny). rewrite pi1_eq, pi2_eq. rewrite Hx. exact Hy. 131 | - intros [n [Hx Hy]%andb_true_iff]. split. 132 | + apply Hf. exists (pi1 n). exact Hx. 133 | + apply Hg. exists (pi2 n). exact Hy. 134 | Qed. 135 | End Sdec_and. 136 | 137 | (** [CO <=> forall X, sdec X -> dec X] *) 138 | 139 | Goal (forall X, sdec X -> dec X) <=> (forall f, dec (tsat f)). 140 | Proof. 141 | split. 142 | - intros H f. apply H. exists f. tauto. 143 | - intros H X [f H1]. specialize (H f) as [H|H]. 144 | + left. apply H1, H. 145 | + right. tauto. 146 | Qed. 147 | 148 | (** Markov <=> Post *) 149 | 150 | Definition Markov: Prop := 151 | forall f: nat -> bool, ~(forall x, f x = false) -> tsat f. 152 | 153 | Definition Post: Type := 154 | forall X: Prop, sdec X -> sdec (~X) -> dec X. 155 | 156 | Definition stable (X: Prop) : Prop := 157 | ~ ~X -> X. 158 | 159 | Lemma dma (f: nat -> bool) : 160 | ~ (forall n, f n = false) <-> ~ ~tsat f. 161 | Proof. 162 | split. 163 | - intros H. contradict H. intros n. 164 | destruct (f n) eqn:H1. 2:reflexivity. 165 | contradict H. exists n. exact H1. 166 | - intros H H1. contradict H. intros [n H]. 167 | revert H. rewrite H1. discriminate. 168 | Qed. 169 | 170 | Fact Markov_stable : 171 | Markov <-> forall f, stable (tsat f). 172 | Proof. 173 | split. 174 | - intros M f H. apply M, dma, H. 175 | - intros H f H1. apply H, dma, H1. 176 | Qed. 177 | 178 | Lemma Markov2Post : 179 | Markov -> Post. 180 | Proof. 181 | intros M X [f H1] [g H2]. 182 | pose (h n := f n || g n). 183 | enough ({ n | h n = true }) as [n [H3|H3]%orb_true_elim]. 184 | - left. apply H1. exists n; exact H3. 185 | - right. apply H2. exists n; exact H3. 186 | - apply wo, M. subst h. cbn. intros H3. 187 | enough (~X /\ ~ ~X) by tauto. split. 188 | + intros [n H4]%H1. 189 | specialize (H3 n) as [H3 _]%orb_false_iff. congruence. 190 | + intros [n H4]%H2. 191 | specialize (H3 n) as [_ H3]%orb_false_iff. congruence. 192 | Qed. 193 | 194 | Goal 195 | Post -> Markov. 196 | Proof. 197 | intros P f H%dma. 198 | specialize (P (tsat f)) as [H1|H1]. 199 | - exists f. tauto. 200 | - exists (fun _ => false). unfold tsat at 2. firstorder congruence. 201 | - exact H1. 202 | - contradict (H H1). 203 | Qed. 204 | 205 | (** tsat not co-semi-decidable *) 206 | 207 | Definition comp {X} (p: X -> Prop) x := ~p x. 208 | Definition CO := ex (bdec tsat). 209 | 210 | Goal Markov -> (forall f, sdec (comp tsat f)) -> CO. 211 | Proof. 212 | intros M H. hnf. 213 | enough (forall f, dec (tsat f)) as [phi H1]%cdec_bdec by eauto. 214 | intros f. apply (Markov2Post M). 215 | - exists f. tauto. 216 | - apply H. 217 | Qed. 218 | 219 | (** Reductions *) 220 | 221 | Definition red {X Y} (p: X -> Prop) (q: Y -> Prop) (f : X -> Y) : Prop 222 | := forall x, p x <-> q (f x). 223 | 224 | Goal forall X Y p q (F: X -> Y), 225 | red p q F -> (forall y, dec (q y)) -> (forall x, dec (p x)). 226 | Proof. 227 | intros * H1 H2 x. 228 | specialize (H2 (F x)). 229 | specialize (H1 x). 230 | revert H2. apply dec_transport. tauto. 231 | Qed. 232 | 233 | Goal forall X Y p q (F: X -> Y), 234 | red p q F -> (forall y, sdec (q y)) -> (forall x, sdec (p x)). 235 | Proof. 236 | intros * H1 H2 x. 237 | specialize (H2 (F x)) as [f H2]. 238 | specialize (H1 x). 239 | exists f. tauto. 240 | Qed. 241 | 242 | Goal forall X (p q: X -> Prop), 243 | (forall x, p x <-> q x) -> red p q (fun x => x). 244 | Proof. 245 | intros * H x. apply H. 246 | Qed. 247 | 248 | (** Predicate is semi-decidable iff it reduces to tsat *) 249 | 250 | Goal forall X (p: X -> Prop), 251 | (forall x, sdec (p x)) <=> sig (red p tsat). 252 | Proof. 253 | split. 254 | - intros H. 255 | exists (fun x => let (f, _) := H x in f). 256 | intros x. cbn. destruct (H x) as [f H1]. exact H1. 257 | - intros [F H] x. specialize (H x). exists (F x). exact H. 258 | Qed. 259 | -------------------------------------------------------------------------------- /rocq/syntax.v: -------------------------------------------------------------------------------- 1 | (*** MPCTT, Chapter Abstract Syntax *) 2 | 3 | From Stdlib Require Import List. 4 | Import ListNotations. 5 | 6 | (*** Expressions and Codes *) 7 | 8 | (* First time we use Rocq's BNF format for an inductive type *) 9 | Inductive exp := 10 | con (x: nat) 11 | | add (e: exp) (e: exp) 12 | | sub (e: exp) (e: exp). 13 | 14 | Fixpoint eval e : nat := 15 | match e with 16 | | con x => x 17 | | add e1 e2 => eval e1 + eval e2 18 | | sub e1 e2 => eval e1 - eval e2 19 | end. 20 | 21 | Compute eval (sub (add (con 3) (con 5)) (con 2)). 22 | 23 | Fixpoint run C A : list nat := 24 | match C, A with 25 | | nil, A => A 26 | | 0::C', x1::x2::A' => run C' (x1+x2::A') 27 | | 1::C', x1::x2::A' => run C' (x1-x2::A') 28 | | S(S x)::C', A => run C' (x::A) 29 | | _, _ => nil 30 | end. 31 | 32 | Compute run [4+2;7+2;1] []. 33 | 34 | (* Rocq compiles all matches into basic matches *) 35 | Set Printing All. 36 | Print run. 37 | Unset Printing All. 38 | 39 | (*** Compiler *) 40 | 41 | Fixpoint com e : list nat := 42 | match e with 43 | | con x => [S(S x)] 44 | | add e1 e2 => com e2 ++ com e1 ++ [0] 45 | | sub e1 e2 => com e2 ++ com e1 ++ [1] 46 | end. 47 | 48 | Compute com (sub (add (con 3) (con 5)) (con 2)). 49 | Compute run (com (sub (add (con 3) (con 5)) (con 2))). 50 | 51 | Theorem correctness e C A : 52 | run (com e ++ C) A = run C (eval e :: A). 53 | Proof. 54 | revert C A. 55 | induction e as [x | e1 IH1 e2 IH2 | e1 IH1 e2 IH2]; 56 | intros C A. 57 | - cbn. reflexivity. 58 | - cbn. rewrite <-app_assoc, IH2. rewrite <-app_assoc, IH1. reflexivity. 59 | - cbn. rewrite <-app_assoc, IH2. rewrite <-app_assoc, IH1. reflexivity. 60 | Qed. 61 | 62 | Corollary correctness' e : 63 | run (com e) nil = [eval e]. 64 | Proof. 65 | rewrite <-(app_nil_r (com e)). 66 | rewrite correctness. 67 | reflexivity. 68 | Qed. 69 | 70 | (*** Decompiler *) 71 | 72 | (* Implicit Type L: list exp. *) 73 | 74 | Fixpoint dcom C L : list exp := 75 | match C, L with 76 | | nil, L => L 77 | | 0::C', e1::e2::L' => dcom C' (add e1 e2 :: L') 78 | | 1::C', e1::e2::L' => dcom C' (sub e1 e2 :: L') 79 | | S(S x)::C', L => dcom C' (con x::L) 80 | | _, _ => nil 81 | end. 82 | 83 | Compute dcom (com (add (con 3) (sub (con 7) (con 5)))) []. 84 | Compute dcom [0;3;0;5;2] []. 85 | 86 | Theorem dcom_correct e C L : 87 | dcom (com e ++ C) L = dcom C (e::L). 88 | Proof. 89 | revert C L. 90 | induction e as [x | e1 IH1 e2 IH2 | e1 IH1 e2 IH2 ]; 91 | intros C L; cbn. 92 | - reflexivity. 93 | - rewrite <-app_assoc, IH2. rewrite <-app_assoc, IH1. reflexivity. 94 | - rewrite <-app_assoc, IH2. rewrite <-app_assoc, IH1. reflexivity. 95 | Qed. 96 | 97 | Corollary dcom_correct' e : 98 | dcom (com e) nil = [e]. 99 | Proof. 100 | rewrite <-(app_nil_r (com e)). 101 | rewrite dcom_correct. 102 | reflexivity. 103 | Qed. 104 | 105 | (*** List Basics *) 106 | 107 | Module List. 108 | Inductive list (X : Type) : Type := 109 | | nil : list X 110 | | cons : X -> list X -> list X. 111 | 112 | Arguments nil {X}. 113 | Arguments cons {X}. 114 | 115 | 116 | Goal forall X (x:X) A, 117 | cons x A <> nil. 118 | Proof. 119 | congruence. 120 | Qed. 121 | 122 | Goal forall X (x:X) A, 123 | cons x A <> nil. 124 | Proof. 125 | intros * H. 126 | change match cons x A with nil => True | cons _ _ => False end. 127 | rewrite H. easy. 128 | Qed. 129 | 130 | Goal forall X (x:X) A x' A', 131 | cons x A = cons x' A' -> x = x' /\ A = A'. 132 | Proof. 133 | intros * H. 134 | change match cons x A with nil => True | cons x A => x = x' /\ A = A' end. 135 | rewrite H. easy. 136 | Qed. 137 | 138 | Fixpoint concat {X} (A B : list X) : list X := 139 | match A with 140 | | nil => B 141 | | cons x A => cons x (concat A B) 142 | end. 143 | 144 | Definition elim_list {X} (p: list X -> Type) 145 | : p nil -> 146 | (forall x A, p A -> p (cons x A)) -> 147 | forall A, p A 148 | := fun f1 f2 => fix f A := 149 | match A with 150 | | nil => f1 151 | | cons x A => f2 x A (f A) 152 | end. 153 | 154 | Fact concat_nil X (A : list X) : 155 | concat A nil = A. 156 | Proof. 157 | pattern A. revert A. apply elim_list. 158 | - reflexivity. 159 | - intros x A IH. 160 | cbn. congruence. 161 | Qed. 162 | 163 | Fact concat_asso X (A B C : list X) : 164 | concat (concat A B) C = concat A (concat B C). 165 | Proof. 166 | pattern A. revert A. apply elim_list. 167 | - reflexivity. 168 | - intros x A IH. 169 | cbn. congruence. 170 | Qed. 171 | 172 | Fixpoint rev {X} (A : list X) : list X := 173 | match A with 174 | | nil => nil 175 | | cons x A => concat (rev A) (cons x nil) 176 | end. 177 | 178 | Goal forall X (A B: list X), 179 | rev (concat A B) = concat (rev B) (rev A). 180 | Proof. 181 | intros *. 182 | induction A as [|x A IH]. 183 | - cbn. rewrite concat_nil. reflexivity. 184 | - cbn. rewrite IH. apply concat_asso. 185 | Qed. 186 | End List. 187 | 188 | (*** Expression Basics *) 189 | 190 | From Stdlib Require Import Lia. 191 | 192 | Module Exp. 193 | Inductive exp := 194 | | con : nat -> exp 195 | | add : exp -> exp -> exp 196 | | sub : exp -> exp -> exp. 197 | 198 | Goal forall e1 e2 e1' e2', 199 | add e1 e2 <> sub e1' e2'. 200 | Proof. 201 | congruence. 202 | Qed. 203 | 204 | Goal forall e1 e2 e1' e2', 205 | add e1 e2 <> sub e1' e2'. 206 | Proof. 207 | intros * H. 208 | change match add e1 e2 with add _ _ => False | _ => True end. 209 | rewrite H. easy. 210 | Qed. 211 | 212 | Goal forall e1 e2 e1' e2', 213 | add e1 e2 = add e1' e2' -> e1 = e1' /\ e2 = e2'. 214 | Proof. 215 | intros * H. split; congruence. 216 | Qed. 217 | 218 | Goal forall e1 e2 e1' e2', 219 | add e1 e2 = add e1' e2' -> e1 = e1' /\ e2 = e2'. 220 | Proof. 221 | intros * H. 222 | change match add e1 e2 with add e1 e2 => e1 = e1' /\ e2 = e2' | _ => True end. 223 | rewrite H. easy. 224 | Qed. 225 | 226 | Fixpoint eval e : nat := 227 | match e with 228 | | con x => x 229 | | add e1 e2 => eval e1 + eval e2 230 | | sub e1 e2 => eval e1 - eval e2 231 | end. 232 | 233 | Fixpoint swap e : exp := 234 | match e with 235 | | con x => con x 236 | | add e1 e2 => add (swap e2) (swap e1) 237 | | sub e1 e2 => sub (swap e1) (swap e2) 238 | end. 239 | 240 | Definition elim_exp (p: exp -> Type) 241 | : (forall x, p (con x)) -> 242 | (forall e1 e2, p e1 -> p e2 -> p (add e1 e2)) -> 243 | (forall e1 e2, p e1 -> p e2 -> p (sub e1 e2)) -> 244 | forall e, p e 245 | := fun f1 f2 f3 => fix f e := 246 | match e with 247 | | con x => f1 x 248 | | add e1 e2 => f2 e1 e2 (f e1) (f e2) 249 | | sub e1 e2 => f3 e1 e2 (f e1) (f e2) 250 | end. 251 | 252 | Goal forall e, eval e = eval (swap e). 253 | Proof. 254 | apply elim_exp. 255 | - intros x. reflexivity. 256 | - intros e1 e2 IH1 IH2. 257 | cbn. lia. 258 | - intros e1 e2 IH1 IH2. 259 | cbn. congruence. 260 | Qed. 261 | End Exp. 262 | 263 | 264 | 265 | -------------------------------------------------------------------------------- /rocq/teaser-bctt.v: -------------------------------------------------------------------------------- 1 | Definition P : nat -> nat := 2 | fun x => match x with 0 => 0 | S x' => x' end. 3 | 4 | Goal 5 | P 2 = 1. 6 | Proof. 7 | cbv delta [P]. 8 | cbv beta. 9 | cbv match. 10 | Abort. 11 | 12 | Definition D: nat -> nat := 13 | fix f x := match x with 0 => 0 | S x' => S (S (f x')) end. 14 | 15 | Print D. 16 | 17 | Goal 18 | D 1 = 2. 19 | Proof. 20 | cbv delta [D]. 21 | cbv fix. fold D. 22 | cbv beta. 23 | cbv match. 24 | cbv delta [D]. 25 | cbv fix. fold D. 26 | cbv beta. 27 | cbv match. 28 | Abort. 29 | 30 | Example demo n : 31 | D (S n) = S (S (D n)). 32 | Proof. 33 | set (a:= S (S (D n))). 34 | cbv delta [D]. cbv fix. fold D. cbv beta. 35 | cbv match. 36 | subst a. 37 | Abort. 38 | 39 | -------------------------------------------------------------------------------- /rocq/teaser-cert.v: -------------------------------------------------------------------------------- 1 | Notation "~ X" := (X -> False) (at level 75, right associativity) : type_scope. 2 | 3 | Inductive void : Type := . 4 | Inductive unit : Type := U. 5 | 6 | Definition inv {X Y: Type} (g: Y -> X) (f: X -> Y) := 7 | forall x, g (f x) = x. 8 | 9 | Definition injective {X Y} (f: X -> Y) := 10 | forall x x', f x = f x' -> x = x'. 11 | 12 | Definition surjective {X Y} (f: X -> Y) := 13 | forall y, exists x, f x = y. 14 | 15 | Fact inv_injective X Y (f: X -> Y) (g: Y -> X) : 16 | inv g f -> injective f. 17 | Proof. 18 | (* congruence. *) 19 | intros H x x' H1. 20 | cut (g (f x) = g (f x')). 21 | - rewrite H. rewrite H. easy. 22 | - f_equal. exact H1. 23 | Qed. 24 | 25 | Fact inv_surjective X Y (f: X -> Y) (g: Y -> X) : 26 | inv g f -> surjective g. 27 | Proof. 28 | intros H x. exists (f x). apply H. 29 | Qed. 30 | 31 | Fact inv_injective_inv X Y (f: X -> Y) (g: Y -> X) : 32 | inv g f -> injective g -> inv f g. 33 | Proof. 34 | intros H1 H2 y. apply H2. apply H1. 35 | Qed. 36 | 37 | Fact inv_surjective_inv X Y (f: X -> Y) (g: Y -> X) : 38 | inv g f -> surjective f -> inv f g. 39 | Proof. 40 | intros H1 H2 y. 41 | specialize (H2 y) as [x H2]. subst y. f_equal. apply H1. 42 | Qed. 43 | 44 | Inductive injection (X Y: Type) : Type := 45 | | Injection {f: X -> Y} {g: Y -> X} (H: inv g f). 46 | 47 | Fact injection_refl X : 48 | injection X X. 49 | Proof. 50 | exists (fun x => x) (fun x => x). hnf. reflexivity. 51 | Qed. 52 | 53 | Fact injection_trans X Y Z : 54 | injection X Y -> injection Y Z -> injection X Z. 55 | Proof. 56 | intros [f g H] [f' g' H']. 57 | exists (fun x => f' (f x)) (fun z => g (g' z)). 58 | intros x. congruence. 59 | Qed. 60 | 61 | Fact injection_Cantor X : 62 | ~ injection (X -> bool) X. 63 | Proof. 64 | intros [f g H]. 65 | pose (h x := negb (g x x)). 66 | enough (g (f h) (f h) = h (f h)) as H1. 67 | { revert H1. unfold h at 3. destruct g; easy. } 68 | rewrite H. reflexivity. 69 | Qed. 70 | 71 | Inductive bijection (X Y: Type) : Type := 72 | | Bijection: forall (f: X -> Y) (g: Y -> X), inv g f -> inv f g -> bijection X Y. 73 | 74 | Goal bijection bool (unit + unit). 75 | Proof. 76 | exists (fun b => if b:bool then inl U else inr U) 77 | (fun a => match a with inl _ => true | inr _ => false end). 78 | - intros b. destruct b; reflexivity. 79 | - intros a. destruct a; destruct u; reflexivity. 80 | (* needs reduction for matches *) 81 | Qed. 82 | 83 | Goal forall X, bijection (X + void) X. 84 | Proof. 85 | intros X. 86 | exists (fun a: X + void => match a with inl x => x | inr v => match v with end end) 87 | inl. 88 | - intros [x|[]]. reflexivity. 89 | - intros x. reflexivity. 90 | Qed. 91 | 92 | Goal bijection (nat + nat) nat. 93 | (* Takes too much arithmetic for now, will do later *) 94 | Abort. 95 | 96 | Goal bijection (nat + unit) nat. 97 | Proof. 98 | exists (fun a: nat + unit => match a with inl x => S x | inr _ => 0 end) 99 | (fun x: nat => match x with 0 => inr U | S x => inl x end). 100 | - intros [x|[]]; reflexivity. 101 | - intros [|x]; reflexivity. 102 | Qed. 103 | 104 | Goal ~ injection unit void. 105 | Proof. 106 | intros [f g H]. 107 | destruct (f U). 108 | Qed. 109 | 110 | (* All empty types are in bijection *) 111 | Goal forall X Y, ~ X -> ~ Y -> bijection X Y. 112 | Proof. 113 | intros * FX FY. 114 | exists (fun x:X => match FX x with end) 115 | (fun y:Y => match FY y with end). 116 | - intros x. destruct (FX x). 117 | - intros y. destruct (FY y). 118 | (* Computational falsity elimination *) 119 | Qed. 120 | 121 | 122 | Definition dec (X: Type) : Type := X + (X -> False). 123 | 124 | Goal forall X Y, 125 | dec X -> dec Y -> dec (X + Y). 126 | Proof. 127 | unfold dec. tauto. 128 | Qed. 129 | 130 | Definition eqdec X := forall x y: X, dec (x = y). 131 | 132 | Fact eqdec_void : eqdec void. 133 | Proof. 134 | intros []. 135 | Qed. 136 | 137 | Fact eqdec_bot : eqdec False. 138 | Proof. 139 | intros []. 140 | (* Computational falsity elimination *) 141 | Qed. 142 | 143 | Fact eqdec_nat : eqdec nat. 144 | Proof. 145 | hnf. 146 | induction x; destruct y. 147 | all: unfold dec. 148 | 1-3: intuition congruence. 149 | destruct (IHx y) as [H|H]; intuition congruence. 150 | Qed. 151 | 152 | Fact eqdec_prod X Y : 153 | eqdec X -> eqdec Y -> eqdec (X*Y). 154 | Proof. 155 | intros dX dY [x y] [x' y']. 156 | destruct (dX x x') as [H1|H1]. 157 | - destruct (dY y y') as [H2|H2]. 158 | + left. congruence. 159 | + right. congruence. 160 | - right. congruence. 161 | Qed. 162 | 163 | Fact eqdec_injective {X Y f} : 164 | @injective X Y f -> eqdec Y -> eqdec X. 165 | Proof. 166 | intros H d x x'. specialize (H x x'). 167 | destruct (d (f x) (f x')) as [H1|H1]; 168 | unfold dec in *; intuition congruence. 169 | Qed. 170 | 171 | Fact eqdec_injection X Y : 172 | injection X Y -> eqdec Y -> eqdec X. 173 | Proof. 174 | intros [f g H] H1. 175 | apply inv_injective in H. 176 | revert H H1. apply eqdec_injective. 177 | Qed. 178 | 179 | 180 | 181 | 182 | 183 | 184 | -------------------------------------------------------------------------------- /rocq/teaser-elim.v: -------------------------------------------------------------------------------- 1 | Lemma elim_bool (p: bool -> Prop) : 2 | p true -> p false -> forall x, p x. 3 | Proof. 4 | intros H1 H2. destruct x. exact H1. exact H2. 5 | Qed. 6 | 7 | Check elim_bool. 8 | 9 | Lemma enum_bool : 10 | forall x:bool, x = true \/ x = false. 11 | Proof. 12 | apply elim_bool. 13 | - tauto. 14 | - tauto. 15 | Qed. 16 | 17 | Lemma Kaminski (f: bool -> bool) : 18 | forall x, f (f (f x )) = f x. 19 | Proof. 20 | apply elim_bool. 21 | - destruct (enum_bool (f true)) as [H|H]. 22 | + congruence. 23 | + destruct (enum_bool (f false)) as [H1|H1]. 24 | * congruence. 25 | * congruence. 26 | - destruct (enum_bool (f false)) as [H1|H1]. 27 | + destruct (enum_bool (f true)) as [H|H]. 28 | * congruence. 29 | * congruence. 30 | + congruence. 31 | Qed. 32 | 33 | Lemma elim_nat (p: nat -> Prop) : 34 | p 0 -> (forall x, p x -> p (S x)) -> forall x, p x. 35 | Proof. 36 | intros H F. 37 | induction x. 38 | - exact H. 39 | - exact (F x IHx). 40 | Qed. 41 | Lemma CA_nat (p: nat -> Prop) : 42 | p 0 -> (forall x, p (S x)) -> forall x, p x. 43 | Proof. 44 | intros H F. 45 | apply elim_nat. 46 | - exact H. 47 | - intros x _. apply F. 48 | Qed. 49 | 50 | Lemma nat_eq : 51 | forall x y:nat, x = y \/ x <> y. 52 | Proof. 53 | refine (elim_nat _ _ _). 54 | - apply CA_nat. 55 | + tauto. 56 | + intros x. right. congruence. 57 | - intros x IHx. 58 | apply CA_nat. 59 | + right. congruence. 60 | + intros y. 61 | destruct (IHx y) as [H|H]. 62 | * left. congruence. 63 | * right. contradict H. congruence. 64 | Qed. 65 | 66 | Lemma elim_pair X Y (p: X * Y -> Prop) : 67 | (forall x y, p (x,y)) -> forall a, p a. 68 | Proof. 69 | intros F. 70 | destruct a as [x y]. 71 | apply F. 72 | Qed. 73 | 74 | Lemma eta_pair X Y : 75 | forall a: X * Y, a = (fst a, snd a). 76 | Proof. 77 | apply elim_pair. 78 | intros x y. 79 | reflexivity. 80 | Qed. 81 | 82 | Definition card_le_two X := 83 | forall x y z : X, x = y \/ x = z \/ y = z. 84 | 85 | Fact bool_card : 86 | card_le_two bool. 87 | Proof. 88 | hnf. 89 | refine (elim_bool _ _ _). 90 | - refine (elim_bool _ _ _). 91 | + auto. 92 | + refine (elim_bool _ _ _); auto. 93 | - refine (elim_bool _ _ _). 94 | + refine (elim_bool _ _ _); auto. 95 | + auto. 96 | Qed. 97 | 98 | Fact nat_card : 99 | ~ card_le_two nat. 100 | Proof. 101 | intros F. 102 | specialize (F 0 1 2). 103 | destruct F as [H|[H|H]]. 104 | - congruence. 105 | - congruence. 106 | - congruence. 107 | Qed. 108 | 109 | Goal nat <> bool. 110 | Proof. 111 | intros H. 112 | apply nat_card. 113 | rewrite H. 114 | apply bool_card. 115 | Qed. 116 | 117 | Inductive void: Type := . 118 | Inductive unit: Type := U. 119 | 120 | Lemma elim_void (Z: Prop): 121 | void -> Z. 122 | Proof. 123 | intros a. 124 | destruct a. 125 | Qed. 126 | 127 | Lemma elim_unit (p: unit -> Prop) : 128 | p U -> forall x, p x. 129 | Proof. 130 | intros H. 131 | destruct x. 132 | exact H. 133 | Qed. 134 | 135 | (** Unit and void *) 136 | 137 | Inductive unit : Type := U. 138 | Inductive void : Type := . 139 | 140 | Lemma elim_void (Z: Prop) : 141 | void -> Z. 142 | intros x. destruct x. 143 | Qed. 144 | 145 | Lemma elim_unit (p: unit -> Prop) : 146 | p U -> forall x, p x. 147 | Proof. 148 | intros H. destruct x. exact H. 149 | Qed. 150 | 151 | -------------------------------------------------------------------------------- /rocq/teaser-ewo.v: -------------------------------------------------------------------------------- 1 | From Stdlib Require Import Lia. 2 | Definition dec (X: Type) : Type := X + (X -> False). 3 | Definition eqdec X := forall x y: X, dec (x = y). 4 | Notation decidable p := (forall x, dec (p x)). 5 | Notation sig := sigT. 6 | Notation Sig := existT. 7 | 8 | Implicit Types n k: nat. 9 | 10 | Section EWO. 11 | Variable p: nat -> Prop. 12 | Variable p_dec: decidable p. 13 | 14 | Inductive T (n: nat) : Prop := C (phi: ~p n -> T (S n)). 15 | 16 | Definition W' 17 | : forall n, T n -> sig p 18 | := fix f n a := let (phi) := a in 19 | match p_dec n with 20 | | inl h => (Sig p n h) 21 | | inr h => f (S n) (phi h) 22 | end. 23 | 24 | (* Eliminator generalizing W' *) 25 | 26 | Definition elim_T (q: nat -> Type) 27 | : (forall n, (~p n -> q (S n)) -> q n) -> 28 | forall n, T n -> q n 29 | := fun e => fix f n a := let (phi) := a in e n (fun h => f (S n) (phi h)). 30 | 31 | 32 | Fact W'_elim_T_agree : 33 | W' = elim_T (fun _ => sig p) 34 | (fun n f => match p_dec n with 35 | | inl h => (Sig p n h) 36 | | inr h => f h 37 | end). 38 | Proof. 39 | reflexivity. 40 | Qed. 41 | 42 | Fact elim_T_unfold q e n phi : 43 | elim_T q e n (C n phi) = e n (fun h => elim_T q e (S n) (phi h)). 44 | Proof. 45 | reflexivity. 46 | Qed. 47 | 48 | Lemma T_lower n : 49 | T n -> T 0. 50 | Proof. 51 | induction n as [|n IH]. easy. 52 | intros H. apply IH. constructor. auto. 53 | Qed. 54 | 55 | Definition W : ex p -> sig p. 56 | Proof. 57 | intros H. eapply W' with (n:=0). 58 | destruct H as [n H]. 59 | enough (T n) as H1%T_lower by easy. 60 | constructor. easy. 61 | Qed. 62 | 63 | End EWO. 64 | 65 | Definition inv {X Y: Type} (g: Y -> X) (f: X -> Y) := forall x, g (f x) = x. 66 | 67 | Section Countable_EWO. 68 | Variable X: Type. 69 | Variable f: X -> nat. 70 | Variable g: nat -> X. 71 | Variable gf: inv g f. 72 | Variable p: X -> Prop. 73 | Variable p_dec: decidable p. 74 | 75 | Definition cewo : ex p -> sig p. 76 | Proof. 77 | intros H. 78 | pose (q n := p (g n)). 79 | assert (q_dec: decidable q). 80 | { intros n. apply p_dec. } 81 | assert (q_ex: ex q). 82 | { destruct H as [x H]. exists (f x). hnf. congruence. } 83 | enough (sig q) as [n H1]. 84 | { exists (g n). exact H1. } 85 | apply W; assumption. 86 | Qed. 87 | End Countable_EWO. 88 | 89 | 90 | -------------------------------------------------------------------------------- /rocq/teaser-exquant.v: -------------------------------------------------------------------------------- 1 | Module Demo. 2 | 3 | Inductive ex {X: Type} (p: X -> Prop) : Prop := E (x:X) (a: p x). 4 | 5 | (* X implicit for ex and E *) 6 | 7 | Definition match_ex {X: Type} (p: X -> Prop) (Z: Prop) 8 | : ex p -> (forall x, p x -> Z) -> Z 9 | := fun a e => match a with E _ x b => e x b end. 10 | 11 | Lemma deMorgan_exists X (p: X -> Prop) : 12 | ~ ex (fun x => p x) <-> forall x, ~ p x. 13 | Proof. 14 | split. 15 | - intros f x a. 16 | apply f. 17 | exact (E p x a). (* note eta conversion *) 18 | - intros f a. 19 | apply (match_ex p False a). 20 | exact f. 21 | Qed. 22 | End Demo. 23 | 24 | Lemma deMorgan_exists X (p: X -> Prop) : 25 | ~ (exists x, p x) <-> forall x, ~ p x. 26 | Proof. 27 | split. 28 | - intros H x H1. apply H. exists x. exact H1. 29 | - intros H [x Hx]. apply (H x Hx). 30 | Qed. 31 | 32 | Lemma forall_exists X (p: X -> Prop) (Z: Prop) : 33 | (forall x, p x -> Z) <-> ((exists x, p x) -> Z). 34 | Proof. 35 | split. 36 | - intros H [x H1]. apply (H x H1). 37 | - intros H x H1. apply H. exists x. exact H1. 38 | Qed. 39 | 40 | 41 | Fact Barber X (p: X -> X -> Prop) : 42 | ~ (exists x, forall y, p x y <-> ~ p y y). 43 | Proof. 44 | intros [x H]. specialize (H x). tauto. 45 | Qed. 46 | 47 | 48 | (** Lawvere *) 49 | 50 | (* Boolean negation has no fixed pont *) 51 | 52 | Fact negb_no_fp : 53 | ~ exists x, negb x = x. 54 | Proof. 55 | intros [[|] H]. 56 | - cbn in H. congruence. 57 | - cbn in H. congruence. 58 | Qed. 59 | 60 | (* Propositional negation has no fixed pont *) 61 | 62 | Fact not_no_fp : 63 | ~ exists P: Prop, (~P) = P. 64 | Proof. 65 | intros [P H]. 66 | enough (H1: P <-> ~ P). 67 | - tauto. 68 | - rewrite H. tauto. 69 | Qed. 70 | 71 | Definition surjective {X Y} (f: X -> Y) := 72 | forall y, exists x, f x = y. 73 | 74 | Theorem Lawvere X Y (f: X -> X -> Y) (g: Y -> Y) : 75 | surjective f -> exists y, g y = y. 76 | Proof. 77 | intros H. 78 | specialize (H (fun x => g (f x x))) as [x H]. 79 | apply (f_equal (fun f => f x)) in H. 80 | exists (f x x). 81 | easy. 82 | Qed. 83 | 84 | Corollary Lawvere_bool X : 85 | ~ exists f: X -> X -> bool, surjective f. 86 | Proof. 87 | intros [f H]. 88 | apply negb_no_fp. 89 | revert H. apply Lawvere. 90 | Qed. 91 | 92 | Corollary Lawvere_Prop X : 93 | ~ exists f: X -> X -> Prop, surjective f. 94 | Proof. 95 | intros [f H]. 96 | apply not_no_fp. 97 | revert H. apply Lawvere. 98 | Qed. 99 | 100 | Definition void : Type := forall X : Type, X. 101 | 102 | (* Universe inconsistency *) 103 | Fail Check (fun f: void => f void). 104 | -------------------------------------------------------------------------------- /rocq/teaser-gs.v: -------------------------------------------------------------------------------- 1 | (** Step through the teaser file and relate what you see to what is discussed in the first chapter “Getting Started” of MPCTT. You want to understand the inductive proofs first. Next try to understand how the inductive functions “min” and “iter” are defined in Rocq. Note that “Fixpoint” is Rocq’s keyword for defining recursive functions. *) 2 | 3 | (** Addition is commutative *) 4 | 5 | Fact addO x : 6 | x + 0 = x. 7 | Proof. 8 | induction x as [|x IH]. 9 | - reflexivity. 10 | - cbn. rewrite IH. reflexivity. 11 | Qed. 12 | 13 | (* Note that types of variables are derived automatically *) 14 | 15 | Fact addS x y : 16 | x + S y = S (x + y). 17 | Proof. 18 | induction x as [|x IH]. 19 | - reflexivity. 20 | - cbn. rewrite IH. reflexivity. 21 | Qed. 22 | 23 | Fact add_comm x y : 24 | x + y = y + x. 25 | Proof. 26 | induction x as [|x IH]; cbn. 27 | - rewrite addO. reflexivity. 28 | - rewrite IH. rewrite addS. reflexivity. 29 | Qed. 30 | 31 | (** Minimum with quantified IH for commutativity *) 32 | 33 | Fixpoint min x y := 34 | match x, y with 35 | | 0, _ => 0 36 | | S x, 0 => 0 37 | | S x, S y => S (min x y) 38 | end. 39 | 40 | Check min. 41 | 42 | Arguments min : simpl nomatch. 43 | (* Don't expose nested match upon simplification *) 44 | 45 | Fact min_comm x y : 46 | min x y = min y x. 47 | Proof. 48 | revert y. 49 | induction x as [|x IH]; cbn. 50 | - destruct y. 51 | + reflexivity. 52 | + reflexivity. 53 | - destruct y. 54 | + reflexivity. 55 | + cbn. rewrite IH. reflexivity. 56 | Qed. 57 | 58 | (** Iteration *) 59 | 60 | Fixpoint iter (X: Type) (f: X -> X) (n:nat) (x:X) : X := 61 | match n with 62 | | 0 => x 63 | | S n' => f (iter X f n' x) 64 | end. 65 | 66 | Check iter. 67 | 68 | Arguments iter {X}. 69 | (* Make first argument implicit *) 70 | 71 | Fact iter_add n x : 72 | n + x = iter S n x. 73 | Proof. 74 | induction n as [|n IH]; cbn. 75 | - reflexivity. 76 | - rewrite IH. reflexivity. 77 | Qed. 78 | 79 | Fact iter_shift X f n x : 80 | @iter X f (S n) x = iter f n (f x). 81 | (* Help type inference with "@iter X" *) 82 | Proof. 83 | induction n as [|n IH]. 84 | - reflexivity. 85 | - cbn. rewrite <-IH. reflexivity. 86 | Qed. 87 | 88 | (** Fibonacci function with iter and unfolding function *) 89 | 90 | Definition Fib f n := 91 | match n with 92 | | 0 => 0 93 | | 1 => 1 94 | | S(S n) => f n + f (S n) 95 | end. 96 | 97 | Check Fib. 98 | Arguments Fib : simpl nomatch. 99 | 100 | Definition step a := (snd a, fst a + snd a). 101 | 102 | Definition fib n := fst (iter step n (0,1)). 103 | 104 | Compute fib 10. 105 | 106 | Fact fib_correct n : 107 | fib n = Fib fib n. 108 | Proof. 109 | destruct n as [|n]. 110 | - cbn. reflexivity. 111 | - destruct n as [|n]. 112 | + cbn. reflexivity. 113 | + cbn. unfold fib. reflexivity. 114 | Qed. 115 | -------------------------------------------------------------------------------- /rocq/teaser-indind.v: -------------------------------------------------------------------------------- 1 | From Stdlib Require Import Lia. 2 | 3 | (*** Inductive Comparisons *) 4 | 5 | Inductive le (x: nat) : nat -> Prop := 6 | | leR : le x x 7 | | leS y : le x y -> le x (S y). 8 | 9 | Fact le_correct x y : 10 | le x y -> x <= y. 11 | Proof. 12 | induction 1 as [| y H IH]. 13 | - lia. 14 | - lia. 15 | Qed. 16 | 17 | Fact le_complete x y : 18 | x <= y -> le x y. 19 | Proof. 20 | induction y as [|y IH]; intros H. 21 | - assert (x = 0) as -> by lia. 22 | apply leR. 23 | - assert (x = S y \/ x <= y) as [->|H1] by lia. 24 | + apply leR. 25 | + apply leS, IH, H1. 26 | Qed. 27 | 28 | Definition elim (x: nat) (p: nat -> Prop) 29 | : p x -> 30 | (forall y, p y -> p (S y)) -> 31 | forall y, le x y -> p y 32 | := fun e1 e2 => fix f _ a := 33 | match a with 34 | | leR _ => e1 35 | | leS _ y a => e2 y (f y a) 36 | end. 37 | 38 | Fact le_correct' x y : 39 | le x y -> x <= y. 40 | Proof. 41 | revert y. apply elim. 42 | - lia. 43 | - lia. 44 | Qed. 45 | 46 | 47 | (*** Inductive Equality *) 48 | 49 | Module EQ. 50 | 51 | Inductive eq X (x: X) : X -> Prop := 52 | | Q : eq X x x. 53 | 54 | Definition R 55 | : forall X (x y: X) (p: X -> Type), eq X x y -> p x -> p y 56 | := fun X x _ p e => match e with 57 | | Q _ _ => fun a => a 58 | end. 59 | 60 | Goal forall X x y, eq X x y <-> x = y. 61 | Proof. 62 | intros *. split. 63 | - intros H. generalize (eq_refl x). apply R, H. 64 | - intros <-. apply Q. 65 | Qed. 66 | End EQ. 67 | 68 | 69 | (*** Reflexive Transitive Closure *) 70 | 71 | Module Star. 72 | Section Star. 73 | Variable X : Type. 74 | Implicit Type R: X -> X -> Prop. 75 | 76 | Inductive star (R: X -> X -> Prop) (x: X) : X -> Prop := 77 | | Nil : star R x x 78 | | Cons y z : R x y -> star R y z -> star R x z. 79 | 80 | Implicit Type p: X -> X -> Prop. 81 | Definition reflexive p := forall x, p x x. 82 | Definition transitive p := forall x y z, p x y -> p y z -> p x z. 83 | Definition incl p p' := forall x y, p x y -> p' x y. 84 | 85 | Fact star_incl R : 86 | incl R (star R). 87 | Proof. 88 | intros x y r. eapply Cons. exact r. apply Nil. 89 | Qed. 90 | 91 | Definition elim R (p: X -> X -> Prop) 92 | : (forall x, p x x) -> 93 | (forall x y z, R x y -> p y z -> p x z) -> 94 | forall x y, star R x y -> p x y 95 | := fun f1 f2 => fix f x _ a := 96 | match a with 97 | | Nil _ _ => f1 x 98 | | Cons _ _ x' z r a => f2 x x' z r (f x' z a) 99 | end. 100 | 101 | Fact star_trans R : 102 | transitive (star R). 103 | Proof. 104 | intros x y z. 105 | revert x y. 106 | refine (elim _ _ _ _). 107 | - easy. 108 | - intros x x' y r IH H. 109 | eapply Cons. exact r. auto. 110 | Qed. 111 | 112 | (** We may also use the automatically generated eliminator for star. *) 113 | 114 | Check star_ind. 115 | 116 | Goal 117 | forall R, transitive (star R). 118 | Proof. 119 | intros R. 120 | induction 1 as [|x x' y r _ IH]. 121 | - easy. 122 | - intros H. eapply Cons. exact r. auto. 123 | Qed. 124 | 125 | Fact least R p : 126 | reflexive p -> 127 | transitive p -> 128 | incl R p -> 129 | incl (star R) p. 130 | Proof. 131 | intros H1 H2 H3. hnf. 132 | apply elim. 133 | - exact H1. 134 | - intros x y z H%H3. apply H2, H. 135 | Qed. 136 | 137 | Fact idempotent R x y: 138 | star (star R) x y <-> star R x y. 139 | Proof. 140 | split. 141 | - apply least. 142 | + intros z. constructor. 143 | + apply star_trans. 144 | + easy. 145 | - apply star_incl. 146 | Qed. 147 | End Star. 148 | End Star. 149 | 150 | (*** PI -> DPI ***) 151 | 152 | Notation sig := sigT. 153 | Notation Sig := existT. 154 | Notation pi1 := projT1. 155 | Notation pi2 := projT2. 156 | 157 | Definition DPI := forall X (p: X -> Type) x u v, Sig p x u = Sig p x v -> u = v. 158 | 159 | Goal DPI. 160 | Proof. 161 | intros X p x u v e. 162 | change u with (pi2 (Sig p x u)). 163 | Fail rewrite e. 164 | Fail pattern (Sig p x u). 165 | Abort. 166 | 167 | Definition PI := forall (P: Prop) (a b : P), a = b. 168 | 169 | Definition cast {X} [p: X -> Type] {x y: X} 170 | : x = y -> p x -> p y 171 | := fun e a => match e with eq_refl => a end. 172 | 173 | Fact cast_eq X (x y: X) p (a: p x): 174 | cast eq_refl a = a. 175 | Proof. 176 | reflexivity. 177 | Qed. 178 | 179 | Section UIP_DPI. 180 | Variable X: Type. 181 | 182 | Definition UIP := forall (x : X) (e: x = x), e = eq_refl. 183 | Definition K_Streicher := forall (x: X) (p: x = x -> Prop), p (eq_refl x) -> forall e, p e. 184 | Definition CD := forall (p: X -> Type) (x: X) (a: p x) (e: x = x), cast e a = a. 185 | Definition DPI' := forall (p: X -> Type) x u v, Sig p x u = Sig p x v -> u = v. 186 | 187 | Lemma L1 : UIP -> K_Streicher. 188 | Proof. 189 | intros H x p a e. hnf in H. rewrite H. exact a. 190 | Qed. 191 | 192 | Lemma L2 : K_Streicher -> CD. 193 | Proof. 194 | intros H p x a. hnf in H. apply H. reflexivity. 195 | Qed. 196 | 197 | Lemma L3 : CD -> DPI'. 198 | Proof. 199 | intros H p x. 200 | enough (forall a b: sig p, a = b -> forall e: pi1 a = pi1 b, cast e (pi2 a) = pi2 b) as H'. 201 | - intros u v e'. apply (H' _ _ e' (eq_refl x)). 202 | - intros a b <-. hnf in H. apply H. 203 | Qed. 204 | End UIP_DPI. 205 | 206 | Goal PI -> DPI. 207 | Proof. 208 | intros H X. 209 | apply L3, L2, L1. 210 | intros x e. apply H. 211 | Qed. 212 | 213 | 214 | -------------------------------------------------------------------------------- /rocq/teaser-la.v: -------------------------------------------------------------------------------- 1 | (*** MPCTT, Chapter Linear Arithmetic, brute force proofs *) 2 | 3 | Arguments Nat.sub : simpl nomatch. 4 | 5 | Notation "x <= y" := (x - y = 0) : nat_scope. 6 | Notation "x < y" := (S x <= y) : nat_scope. 7 | 8 | Fact lt_le x y : 9 | x < y -> x <= y. 10 | Proof. 11 | revert y. 12 | induction x as [|x IH]; cbn. 13 | - easy. 14 | - destruct y; cbn. easy. 15 | apply IH. 16 | Qed. 17 | 18 | Fact le_refl x : 19 | x <= x. 20 | Proof. 21 | induction x as [|x IH]; cbn; easy. 22 | Qed. 23 | 24 | Fact le_trans x y z: 25 | x <= y -> y <= z -> x <= z. 26 | Proof. 27 | revert y z. 28 | induction x as [|x IH]; cbn. easy. 29 | destruct y. easy. 30 | destruct z. easy. 31 | apply (IH y z). 32 | Qed. 33 | 34 | Fact le_trans' x y z: 35 | x < y -> y <= z -> x < z. 36 | Proof. 37 | apply le_trans. 38 | Qed. 39 | 40 | Fact le_trans'' x y z: 41 | x <= y -> y < z -> x < z. 42 | Proof. 43 | revert z x y. 44 | induction z as [|z IH]; cbn. easy. 45 | destruct x. easy. 46 | destruct y. easy. 47 | apply (IH x y). 48 | Qed. 49 | 50 | Fact le_anti x y : 51 | x <= y -> y <= x -> x = y. 52 | Proof. 53 | revert y. 54 | induction x as [|x IH]; destruct y; cbn. 55 | 1-3:easy. auto. 56 | Qed. 57 | 58 | Fact sub_add_zero x y : 59 | x <= x + y. 60 | Proof. 61 | induction x as [|x IH]; easy. 62 | Qed. 63 | 64 | Fact le_add_char x y : 65 | x <= y -> x + (y - x) = y. 66 | Proof. 67 | revert y. 68 | induction x as [|x IH]; destruct y; cbn. 69 | 1-3:easy. auto. 70 | Qed. 71 | 72 | (** Contra rules *) 73 | 74 | Fact lt_contra_zero x : 75 | x < 0 -> False. 76 | Proof. 77 | easy. 78 | Qed. 79 | 80 | Fact lt_contra_add x y : 81 | x + y < x -> False. 82 | Proof. 83 | induction x as [|x IH]; easy. 84 | Qed. 85 | 86 | Fact lt_contra_self x : 87 | x < x -> False. 88 | Proof. 89 | induction x as [|x IH]; easy. 90 | Qed. 91 | 92 | 93 | (* Decisions *) 94 | 95 | Goal forall x y, (x <= y) + (y < x). 96 | Proof. 97 | induction x; destruct y; cbn; auto. 98 | Qed. 99 | 100 | Goal forall x y, (x <= y) + ~(x <= y). 101 | Proof. 102 | induction x; destruct y; cbn; auto. 103 | Qed. 104 | 105 | Fact lt_contra x y: 106 | (x <= y) <-> ~(y < x). 107 | Proof. 108 | revert y. 109 | induction x as [|x IH]; destruct y; cbn; easy. 110 | Qed. 111 | 112 | Fact le_test x y : 113 | if x - y then x <= y else ~(x <= y). 114 | Proof. 115 | destruct (x - y); easy. 116 | Qed. 117 | 118 | Fact lt_test x y : 119 | if S x - y then x < y else ~(x < y). 120 | Proof. 121 | destruct (S x - y); easy. 122 | Qed. 123 | 124 | Fact eq_test x y : 125 | if (x - y) + (y - x) then x = y else ~(x = y). 126 | Proof. 127 | revert y. 128 | induction x as [|x IH]; destruct y; cbn. 129 | 1-3:easy. 130 | specialize (IH y). destruct (_ + _); congruence. 131 | Qed. 132 | 133 | Fact le_zero x : 134 | x <= 0 -> x = 0. 135 | Proof. 136 | destruct x; easy. 137 | Qed. 138 | 139 | Fact tightness_dec x y : 140 | x <= y -> y <= S x -> (x=y) + (y = S x). 141 | Proof. 142 | revert y. 143 | induction x as [|x IH]; destruct y; cbn. 144 | - auto. 145 | - destruct y; cbn; auto. 146 | - easy. 147 | - specialize (IH y); intuition congruence. 148 | Qed. 149 | 150 | 151 | 152 | 153 | 154 | 155 | 156 | 157 | -------------------------------------------------------------------------------- /rocq/teaser-leibniz.v: -------------------------------------------------------------------------------- 1 | Goal forall X (x y: X), x = y -> y = x. 2 | Proof. 3 | intros * e. 4 | pattern y. (* conversion *) 5 | rewrite <- e. 6 | reflexivity. 7 | Qed. 8 | 9 | Goal True <> False. 10 | intros e. 11 | pattern False. (* conversion *) 12 | rewrite <- e. 13 | exact I. 14 | Qed. 15 | 16 | Goal true <> false. 17 | intros e. 18 | change (if false then True else False). (* conversion *) 19 | pattern false. (* conversion *) 20 | rewrite <- e. 21 | exact I. 22 | Qed. 23 | 24 | Goal forall x y, S x = S y -> x = y. 25 | intros * e. 26 | change (x = match S y with 0 => 0 | S y => y end). 27 | pattern (S y). 28 | rewrite <- e. 29 | reflexivity. 30 | Qed. 31 | 32 | Section AbstractEquality. 33 | Variable eq : forall X, X -> X -> Prop. 34 | Variable Q : forall X x, eq X x x. 35 | Variable R : forall X x y (p: X -> Prop), eq X x y -> p x -> p y. 36 | 37 | Arguments eq {X}. 38 | Arguments Q {X}. 39 | Arguments R {X x y}. 40 | Notation "s = t" := (eq s t) : type_scope. 41 | Notation "s <> t" := (~ s = t) : type_scope. 42 | 43 | Goal forall X (x y: X), x = y -> y = x. 44 | Proof. 45 | intros * e. 46 | pattern y. 47 | apply (R _ e). 48 | apply Q. 49 | Qed. 50 | 51 | Check fun X (x y: X) (e : x = y) => 52 | R (fun y => y = x) e (Q x). 53 | 54 | Goal True <> False. 55 | intros e. 56 | pattern False. 57 | apply (R _ e). 58 | exact I. 59 | Qed. 60 | 61 | Check fun X (x y: X) (e : x = y) => 62 | R (fun y => y = x) e (Q x). 63 | 64 | Goal true <> false. 65 | intros e. 66 | change (if false then True else False). 67 | pattern false. 68 | apply (R _ e). 69 | exact I. 70 | Qed. 71 | 72 | Check fun (e : true = false) => 73 | R (fun b:bool => if b then True else False) e I. 74 | 75 | Goal forall x y, S x = S y -> x = y. 76 | intros * e. 77 | change (x = match S y with 0 => 0 | S y => y end). 78 | pattern (S y). 79 | apply (R _ e). 80 | apply Q. 81 | Qed. 82 | 83 | Check fun x y (e : S x = S y) => 84 | R (fun z => x = match z with 0 => 0 | S y => y end) e (Q x). 85 | -------------------------------------------------------------------------------- /rocq/teaser-list.v: -------------------------------------------------------------------------------- 1 | From Stdlib Require Import Arith Lia List. 2 | Definition dec (X: Type) := sum X (X -> False). 3 | Definition eqdec X := forall x y: X, dec (x = y). 4 | Import ListNotations. 5 | Notation "x 'el' A" := (In x A) (at level 70). 6 | 7 | Section List1. 8 | Variable X : Type. 9 | Implicit Types A B : list X. 10 | 11 | Goal forall x A B, 12 | x el A ++ B <-> x el A \/ x el B. 13 | Proof. 14 | intros x A B. 15 | induction A as [|a A IH]; cbn. 16 | - intuition. 17 | - intuition. 18 | Qed. 19 | 20 | Goal eqdec X -> eqdec (list X). 21 | Proof. 22 | intros d A. 23 | induction A as [|a A IH]; cbn. 24 | - intros [|b B]. 25 | + left. easy. 26 | + right. easy. 27 | - intros [|x B]. 28 | + right. easy. 29 | + destruct (d a x) as [H|H]. 30 | * destruct (IH B) as [H1|H2]. 31 | -- left. congruence. 32 | -- right. congruence. 33 | * right. congruence. 34 | Qed. 35 | -------------------------------------------------------------------------------- /rocq/teaser-pat.v: -------------------------------------------------------------------------------- 1 | Section Demo. 2 | Variables X Y Z: Prop. 3 | 4 | Goal X -> Y -> X. 5 | Proof. 6 | intros x y. exact x. 7 | Show Proof. 8 | Qed. 9 | 10 | Goal X -> Y -> X. 11 | Proof. 12 | refine (fun x y => _). 13 | Show Proof. 14 | exact x. 15 | Qed. 16 | 17 | Goal (X -> Y -> Z) -> X /\ Y -> Z. 18 | Proof. 19 | intros f [x y]. exact (f x y). 20 | Show Proof. 21 | Qed. 22 | 23 | Goal (X -> Y -> Z) <-> (X /\ Y -> Z). 24 | Proof. 25 | split. 26 | Show Proof. 27 | - intros f [x y]. exact (f x y). 28 | - intros f x y. apply f. split. 29 | + exact x. 30 | + exact y. 31 | Show Proof. 32 | Qed. 33 | 34 | Goal X -> ~X -> False. 35 | Proof. 36 | exact (fun x f => f x). 37 | Show Proof. 38 | Qed. 39 | 40 | Goal X -> ~X -> Y. 41 | Proof. 42 | refine (fun x f => _). 43 | refine (False_ind _ _). 44 | exact (f x). 45 | Show Proof. 46 | Qed. 47 | 48 | Goal X -> ~X -> Y. 49 | Proof. 50 | intros x f. 51 | exfalso. 52 | exact (f x). 53 | Show Proof. 54 | Qed. 55 | 56 | Goal (X -> ~X) -> (~X -> X) -> False. 57 | Proof. 58 | refine (fun f g => _). 59 | refine (let x:X := _ in f x x). 60 | refine (g (fun x => _)). 61 | exact (f x x). 62 | Show Proof. 63 | Qed. 64 | 65 | Goal (X -> ~X) -> (~X -> X) -> False. 66 | Proof. 67 | intros f g. 68 | assert (x: X). 69 | - apply g. intros x. exact (f x x). 70 | - exact (f x x). 71 | Show Proof. 72 | Qed. 73 | 74 | Goal (X -> ~X) -> (~X -> X) -> False. 75 | Proof. 76 | refine (fun f g => _). 77 | refine (let x:X := _ in _). 78 | exact (f x x). 79 | Unshelve. 80 | refine (g (fun x => _)). 81 | exact (f x x). 82 | Show Proof. 83 | Qed. 84 | 85 | End Demo. 86 | -------------------------------------------------------------------------------- /rocq/teaser-sizerec.v: -------------------------------------------------------------------------------- 1 | From Stdlib Require Import Arith Lia. 2 | 3 | Definition size_rec {X: Type} (sigma: X -> nat) {p: X -> Type} : 4 | (forall x, (forall y, sigma y < sigma x -> p y) -> p x) -> 5 | (forall x, p x). 6 | Proof. 7 | intros F. 8 | enough (forall n x, sigma x < n -> p x) as H. 9 | { intros x. apply (H (S (sigma x))). lia. } 10 | induction n as [|n IH]; intros x H. 11 | - exfalso. lia. 12 | - apply F. intros y H1. apply IH. lia. 13 | Defined. 14 | 15 | Definition size_rec2 {X Y: Type} (sigma: X -> Y -> nat) {p: X -> Y -> Type} : 16 | (forall x y, (forall x' y', sigma x' y' < sigma x y -> p x' y') -> p x y) -> 17 | (forall x y, p x y). 18 | Proof. 19 | intros F. 20 | pose (p' a := p (fst a) (snd a)). 21 | pose (sigma' a := sigma (fst a) (snd a)). 22 | enough (forall a, p' a) as H. 23 | { intros x y. exact (H (x,y)). } 24 | refine (size_rec sigma' _ ). 25 | intros [x y] IH. 26 | apply F. intros x' y'. cbn. 27 | specialize (IH (x',y')). exact IH. 28 | Defined. 29 | 30 | Definition gcd_rec (p: nat -> nat -> Type) : 31 | (forall y, p 0 y) -> 32 | (forall x, p (S x) 0) -> 33 | (forall x y, 0 < x <= y -> p x (y - x) -> p x y) -> 34 | (forall x y, 0 < y < x -> p (x - y) y -> p x y) -> 35 | forall x y, p x y. 36 | Proof. 37 | intros e1 e2 e3 e4. 38 | refine (size_rec2 (fun x y => x + y) _). 39 | intros [|x] y IH. 40 | - apply e1. 41 | - destruct y as [|y]. 42 | + apply e2. 43 | + destruct (le_lt_dec x y) as [H|H]. 44 | * apply e3. lia. apply IH. lia. 45 | * apply e4. lia. apply IH. lia. 46 | Qed. 47 | 48 | 49 | -------------------------------------------------------------------------------- /rocq/vector.v: -------------------------------------------------------------------------------- 1 | From Stdlib Require Import Arith Lia List. 2 | Import ListNotations. 3 | From Equations Require Import Equations. 4 | Ltac refl := reflexivity. 5 | Definition dec (P: Prop) : Type := P + ~P. 6 | Definition eqdec X : Type := forall x x': X, dec (x = x'). 7 | Notation sig := sigT. 8 | Notation Sig := existT. 9 | Notation pi1 := projT1. 10 | Notation pi2 := projT2. 11 | Notation "'Sigma' x .. y , p" := 12 | (sig (fun x => .. (sig (fun y => p)) ..)) 13 | (at level 200, x binder, right associativity, 14 | format "'[' 'Sigma' '/ ' x .. y , '/ ' p ']'") 15 | : type_scope. 16 | 17 | Inductive num : nat -> Type := 18 | | Zero: forall n, num (S n) 19 | | Next: forall {n}, num n -> num (S n). 20 | 21 | Section Vector. 22 | Variable X : Type. 23 | 24 | Inductive vector : nat -> Type := 25 | | Nil : vector O 26 | | Cons : forall {n}, X -> vector n -> vector (S n). 27 | 28 | Definition vector_elim (p: forall n, vector n -> Type) 29 | : p 0 Nil -> 30 | (forall n x v, p n v -> p (S n) (Cons x v)) -> 31 | forall n v, p n v 32 | := fun e1 e2 => fix f n v := match v with 33 | | Nil => e1 34 | | @Cons n' x v' => e2 n' x v' (f n' v') 35 | end. 36 | 37 | Definition vec_inv : 38 | forall {n} (v: vector n), 39 | match n return vector n -> Type with 40 | | 0 => fun v => v = Nil 41 | | S n' => fun v => Sigma x v', v = Cons x v' 42 | end v. 43 | Proof. 44 | intros n v. 45 | destruct v as [|n x v]. 46 | - refl. 47 | - exists x, v. refl. 48 | Defined. 49 | 50 | Goal forall v: vector 0, 51 | v = Nil. 52 | Proof. 53 | intros v. apply (vec_inv v). 54 | Qed. 55 | 56 | Goal forall n (v: vector (S n)), 57 | Sigma x v', v = Cons x v'. 58 | Proof. 59 | intros n v. apply (vec_inv v). 60 | Qed. 61 | 62 | (** Head, tail, last, sub *) 63 | 64 | Definition hd {n} (v: vector (S n)) : X := 65 | pi1 (vec_inv v). 66 | 67 | Definition tl {n} (v: vector (S n)) : vector n := 68 | pi1 (pi2 (vec_inv v)). 69 | 70 | Fact eta_law n (v:vector (S n)) : 71 | v = Cons (hd v) (tl v). 72 | Proof. 73 | destruct (vec_inv v) as (x&v'&->). 74 | refl. 75 | Qed. 76 | 77 | Fixpoint last {n} : vector (S n) -> X := 78 | match n with 79 | | 0 => hd 80 | | S n' => fun v => @last n' (@tl (S n') v) 81 | end. 82 | 83 | Fixpoint trunc {n} : vector (S n) -> vector n := 84 | match n with 85 | | 0 => fun _ => Nil 86 | | S n' => fun v => Cons (hd v) (trunc (tl v)) 87 | end. 88 | 89 | Fixpoint sub {n} (a: num n) : vector n -> X := 90 | match a with 91 | | Zero n' => @hd n' 92 | | @Next n' a' => fun v => @sub n' a' (@tl n' v) 93 | end. 94 | 95 | (** Constructor injectivity *) 96 | 97 | Fact Cons_inj1 n x x' (v v': vector n) : 98 | Cons x v = Cons x' v' -> x = x'. 99 | Proof. 100 | intros H % (f_equal hd). exact H. 101 | Qed. 102 | 103 | Fact Cons_inj2 n x x' (v v': vector n) : 104 | Cons x v = Cons x' v' -> v = v'. 105 | Proof. 106 | intros H % (f_equal tl). 107 | exact H. 108 | Qed. 109 | 110 | (** Equality Decider *) 111 | 112 | Definition vector_eq n : 113 | eqdec X -> forall v1 v2: vector n, dec (v1 = v2). 114 | Proof. 115 | intros H. 116 | induction v1 as [|n x1 v1 IH]; 117 | intros v2; 118 | generalize (vec_inv v2). 119 | - intros ->. left. refl. 120 | - intros (x2&v2'&->). 121 | specialize (H x1 x2) as [<-|H]. 122 | + destruct (IH v2') as [<-|H1]. 123 | * left. refl. 124 | * right. contradict H1. exact (f_equal tl H1). 125 | + right. contradict H. exact (f_equal hd H). 126 | Defined. 127 | 128 | (** Generalized head and tail *) 129 | 130 | Definition Hd {n} (v: vector n) 131 | : match n return Type with 0 => True | S n' => X end 132 | := match v with Nil => I | Cons x _ => x end. 133 | 134 | Definition Tl {n} (v: vector n) 135 | : match n return Type with 0 => True | S n' => vector n' end 136 | := match v with Nil => I | Cons _ v' => v' end. 137 | 138 | Fact Cons_inj n x x' (v v': vector n) : 139 | Cons x v = Cons x' v' -> x = x' /\ v = v'. 140 | Proof. 141 | intros H. split. 142 | - exact (f_equal Hd H). 143 | - exact (f_equal Tl H). 144 | Qed. 145 | 146 | (** Snoc, reverse *) 147 | 148 | Fixpoint snoc {n} (v: vector n) (x: X) : vector (S n) := 149 | match v with 150 | | Nil => Cons x Nil 151 | | Cons x' v' => Cons x' (snoc v' x) 152 | end. 153 | 154 | Fixpoint rev {n} (v: vector n) : vector n := 155 | match v with 156 | | Nil => Nil 157 | | Cons x v' => snoc (rev v') x 158 | end. 159 | 160 | Goal forall x n (v: vector n), 161 | last (snoc v x) = x. 162 | Proof. 163 | induction v as [|n y v IH]; cbn. 164 | - refl. 165 | - exact IH. 166 | Qed. 167 | 168 | Goal forall n (v: vector (S n)), 169 | v = snoc (trunc v) (last v). 170 | Proof. 171 | induction n as [|n IH]; 172 | intros v; 173 | destruct (vec_inv v) as (x&v'&->); 174 | cbn; f_equal. 175 | - destruct (vec_inv v'); refl. 176 | - apply IH. 177 | Qed. 178 | 179 | Fact rev_snoc x n (v: vector n) : 180 | rev (snoc v x) = Cons x (rev v). 181 | Proof. 182 | induction v as [|n y v IH]; cbn. 183 | - refl. 184 | - rewrite IH. refl. 185 | Qed. 186 | 187 | Goal forall n (v: vector n), 188 | rev (rev v) = v. 189 | Proof. 190 | induction v as [|n x v IH]; cbn. 191 | - refl. 192 | - rewrite rev_snoc, IH. refl. 193 | Qed. 194 | 195 | 196 | (** Converting between vectors and lists *) 197 | 198 | Fixpoint vec_list {n} (v: vector n) : list X := 199 | match v with 200 | | Nil => nil 201 | | Cons x v' => cons x (vec_list v') 202 | end. 203 | 204 | Fact vec_list_length n : 205 | forall v: vector n, length (vec_list v) = n. 206 | Proof. 207 | induction v as [|n x v IH]; cbn. 208 | - refl. 209 | - f_equal. exact IH. 210 | Qed. 211 | 212 | Fact vec_list_injective n (v1 v2: vector n) : 213 | vec_list v1 = vec_list v2 -> v1 = v2. 214 | Proof. 215 | revert n v1 v2. 216 | induction v1 as [|n x1 v1 IH]; 217 | intros v2; 218 | generalize (vec_inv v2). 219 | - intros ->. easy. 220 | - intros (x2&v2'&->); cbn. 221 | intros [= <- H]. 222 | f_equal. apply IH, H. 223 | Qed. 224 | 225 | Definition guarded_list_vec : 226 | forall A, Sigma v: vector (length A), vec_list v = A. 227 | Proof. 228 | induction A as [|x A IH]; cbn. 229 | - exists Nil. refl. 230 | - specialize IH as [v IH]. 231 | exists (Cons x v). cbn. f_equal. exact IH. 232 | Defined. 233 | 234 | Definition list_vec (A: list X) : vector (length A) := 235 | pi1 (guarded_list_vec A). 236 | 237 | Fact vec_list_eq : 238 | forall A, vec_list (list_vec A) = A. 239 | Proof. 240 | intros A. 241 | apply (pi2 (guarded_list_vec A)). 242 | Qed. 243 | 244 | Fail Check forall n (v: vector n), 245 | list_vec (vec_list v) = v. 246 | 247 | Fact vec_cast {n} {v: vector n} : 248 | vector (length (vec_list v)) -> vector n. 249 | Proof. 250 | rewrite vec_list_length. easy. 251 | Qed. 252 | 253 | Check forall n (v: vector n), 254 | vec_cast (list_vec (vec_list v)) = v. 255 | 256 | 257 | (** Smart matches *) 258 | 259 | Definition hd' {n} (v: vector (S n)) : X := 260 | match v with Cons x _ => x end. 261 | 262 | Definition tl' {n} (v: vector (S n)) : vector n := 263 | match v with Cons _ v' => v' end. 264 | 265 | Fact eta_law' n (v: vector (S n)) : 266 | v = Cons (hd' v) (tl' v). 267 | Proof. 268 | refine (match v with Cons x v' => eq_refl end). 269 | Qed. 270 | 271 | Definition hd'' {n} (v: vector (S n)) : X := 272 | match v 273 | in vector n return match n with 0 => True | S _ => X end 274 | with 275 | | Nil => I 276 | | Cons x _ => x 277 | end. 278 | 279 | Goal forall v: vector 0, 280 | v = Nil. 281 | Proof. 282 | refine (fun v => match v with Nil => eq_refl end). 283 | Qed. 284 | 285 | Goal forall n (v: vector (S n)), 286 | Sigma x v', v = Cons x v'. 287 | Proof. 288 | refine (fun n v => match v with Cons x v' => _ end). 289 | exists x, v'. refl. 290 | Qed. 291 | 292 | (** Experiments *) 293 | 294 | Fixpoint revapp {m} (v: vector m) : 295 | forall {n} (w: vector n), vector (tail_plus m n) := 296 | match v with 297 | | Nil => fun n w => w 298 | | Cons x v => fun n w => revapp v (Cons x w) 299 | end. 300 | 301 | Fail Goal forall n (v: vector n), 302 | rev v = revapp v Nil. 303 | 304 | Fixpoint app {m n} (v: vector m) (w: vector n) : vector (m + n) := 305 | match v with 306 | | Nil => w 307 | | Cons x v' => Cons x (app v' w) 308 | end. 309 | 310 | Fixpoint cast {m n} (v: vector (m + S n)) : vector (S (m + n)). 311 | Proof. 312 | rewrite plus_n_Sm. exact v. 313 | Defined. 314 | 315 | Fact snoc_app m n (v: vector m) (w: vector n) x : 316 | snoc (app v w) x = cast (app v (snoc w x)). 317 | Proof. 318 | induction v as [|n' x' v' IH] in n,w |-*. 319 | - cbn. 320 | Abort. 321 | 322 | End Vector. 323 | 324 | Arguments Nil {X}. 325 | Arguments Cons {X n}. 326 | Arguments hd {X n}. 327 | Arguments hd'' {X n}. 328 | Arguments tl {X n}. 329 | Arguments last {X n}. 330 | Arguments rev {X n}. 331 | Arguments sub {X n}. 332 | Arguments vec_list {X n}. 333 | Arguments list_vec {X}. 334 | Arguments vec_list {X n}. 335 | 336 | Compute rev (Cons 5 (Cons 2 Nil)). 337 | Compute hd (tl (Cons 5 (Cons 2 Nil))). 338 | Compute hd'' (tl (Cons 5 (Cons 2 Nil))). 339 | Compute last (Cons 5 (Cons 2 Nil)). 340 | 341 | Compute sub (Zero 1) (Cons 5 (Cons 2 Nil)). 342 | Compute sub (Next (Zero 0)) (Cons 5 (Cons 2 Nil)). 343 | Compute vec_list (Cons 5 (Cons 2 Nil)). 344 | 345 | Compute list_vec [5;3;7]. 346 | Compute vec_list (list_vec [5;3;7]). 347 | 348 | Goal vec_list (list_vec [5;3;7]) = [5;3;7]. 349 | Proof. 350 | refl. 351 | Qed. 352 | 353 | (** Exercises *) 354 | 355 | Fixpoint map {X Y} (f: X -> Y) {n} (v: vector X n) : vector Y n := 356 | match v with 357 | | Nil => Nil 358 | | Cons x v' => Cons (f x) (map f v') 359 | end. 360 | 361 | Fixpoint tab n : vector (num n) n := 362 | match n with 363 | | 0 => Nil 364 | | S n' => Cons (Zero n') (map Next (tab n')) 365 | end. 366 | 367 | Compute tab 7. 368 | Compute rev (tab 7). 369 | Compute sub (Zero 6) (rev (tab 7)). 370 | -------------------------------------------------------------------------------- /rocq/xm.v: -------------------------------------------------------------------------------- 1 | Definition XM: Prop := 2 | forall P, P \/ ~P. 3 | Definition DN: Prop := 4 | forall P, ~~P -> P. 5 | Definition Contra: Prop := 6 | forall P Q, (~Q -> ~P) -> P -> Q. 7 | Definition Peirce: Prop := 8 | forall P Q: Prop, ((P -> Q) -> P) -> P. 9 | Definition XM': Prop := 10 | forall P, (P <-> True) \/ (P <-> False). 11 | 12 | Fact XM_XM' : 13 | XM <-> XM'. 14 | Proof. 15 | split. 16 | - intros H P. specialize (H P). tauto. 17 | - intros H P. specialize (H P). tauto. 18 | Qed. 19 | 20 | Fact XM_DN : 21 | XM -> DN. 22 | Proof. 23 | intros H P H1. specialize (H P) as [H|H]. 24 | exact H. contradict (H1 H). 25 | Qed. 26 | 27 | Fact DN_Contra : 28 | DN -> Contra. 29 | Proof. 30 | intros H P Q H1 x. apply H. clear H. 31 | intros H2. revert x. apply H1, H2. 32 | Qed. 33 | 34 | Fact Contra_Peirce : 35 | Contra -> Peirce. 36 | Proof. 37 | intros H P Q. apply H. clear H. 38 | intros H1 H2. apply H1, H2. intros x. 39 | contradict (H1 x). 40 | Qed. 41 | 42 | Fact Peirce_XM' : 43 | Peirce -> XM. 44 | Proof. 45 | intros H P. apply H with (Q:=False). clear H. 46 | intros H. right. intros x. apply H. left. exact x. 47 | Qed. 48 | 49 | Fact counterexample: 50 | XM <-> forall X (p: X -> Prop), (forall x, p x) \/ (exists x, ~p x). 51 | Proof. 52 | split. 53 | - intros H X p. 54 | destruct (H (exists x, ~p x)) as [H1|H1]. 55 | + right. exact H1. 56 | + left. intros x. 57 | specialize (H (p x)) as [H|H]. exact H. 58 | exfalso. apply H1. exists x. exact H. 59 | - intros H P. 60 | specialize (H True (fun _ => P)) as [H|[_ H]]. 61 | + left. apply H, I. 62 | + right. exact H. 63 | Qed. 64 | 65 | Fact Peirce_DN : 66 | Peirce <-> DN. 67 | Proof. 68 | split; intros H. 69 | - intros P H1. specialize (H P False). (* tauto. *) 70 | apply H. intros H2. exfalso. apply H1. exact H2. 71 | - intros P Q H1. apply H. clear H. (* tauto. *) 72 | intros H2. apply H2. apply H1. intros H3. exfalso. apply H2. exact H3. 73 | Qed. 74 | 75 | (*** Stability *) 76 | 77 | Definition stable (P: Prop) := ~~P -> P. 78 | 79 | Fact stable_not P : 80 | stable (~P). 81 | Proof. 82 | cbv. tauto. 83 | Qed. 84 | 85 | Fact stable_invariant P Q : 86 | (P <-> Q) -> (stable P <-> stable Q). 87 | Proof. 88 | unfold stable. tauto. 89 | Qed. 90 | 91 | Goal forall P Q, stable P -> stable Q -> ~ (P /\ Q) -> ~P \/ ~Q. 92 | Proof. 93 | unfold stable. 94 | Abort. 95 | 96 | Goal forall P Q, stable Q -> stable (P -> Q). 97 | Proof. 98 | unfold stable. tauto. 99 | Qed. 100 | 101 | Goal forall P Q, stable P -> stable Q -> stable (P /\ Q). 102 | Proof. 103 | unfold stable. tauto. 104 | Qed. 105 | 106 | Goal forall A (p: A -> Prop), (forall a, stable (p a)) -> stable (forall a, p a). 107 | Proof. 108 | unfold stable. intros a p H1 H2 b. 109 | apply H1; clear H1. intros H1. 110 | apply H2; clear H2; intros H2. 111 | apply H1, H2. 112 | Qed. 113 | 114 | Goal forall P (p: P -> Prop), 115 | (forall x, stable (p x)) -> ~(forall x, p x) -> ~~(exists x, ~p x). 116 | Proof. 117 | intros P p H H1 H2. apply H1. intros x. 118 | apply H. intros H3. apply H2. exists x. exact H3. 119 | Qed. 120 | 121 | 122 | Goal forall P, stable P <-> exists Q, P <-> ~Q. 123 | Proof. 124 | intros P. 125 | split. 126 | - intros H. exists (~P). split. 2:exact H. tauto. 127 | - intros [Q H]. unfold stable. tauto. 128 | Qed. 129 | 130 | Goal forall P1 P2 Q, stable Q -> ~(P1 /\ P2) -> (~P1 \/ ~P2 -> Q) -> Q. 131 | Proof. 132 | intros *. unfold stable. tauto. 133 | Qed. 134 | 135 | Goal forall P1 P2 Q, stable Q -> (~P1 -> ~P2) -> ((P2 -> P1) -> Q) -> Q. 136 | Proof. 137 | intros *. unfold stable. tauto. 138 | Qed. 139 | 140 | (*** Definiteness *) 141 | 142 | Definition definite (P: Prop) := P \/ ~P. 143 | 144 | Goal forall P, definite P -> stable P. 145 | Proof. 146 | intros P. cbv. tauto. 147 | Qed. 148 | 149 | Goal forall P Q, definite P -> definite Q -> definite (P -> Q). 150 | Proof. 151 | unfold definite. tauto. 152 | Qed. 153 | 154 | Goal forall P Q, definite P -> definite Q -> definite (P /\ Q). 155 | Proof. 156 | unfold definite. tauto. 157 | Qed. 158 | 159 | Goal forall P Q, definite P -> definite Q -> definite (P \/ Q). 160 | Proof. 161 | unfold definite. tauto. 162 | Qed. 163 | 164 | Goal forall P, definite P -> definite (~P). 165 | Proof. 166 | unfold definite. tauto. 167 | Qed. 168 | 169 | Goal forall P Q, definite P \/ definite Q -> ~(P /\ Q) <-> ~P \/ ~Q. 170 | Proof. 171 | unfold definite. tauto. 172 | Qed. 173 | 174 | (*** Truth Value Semantics *) 175 | 176 | Definition TVS : Prop := 177 | forall P: Prop, P = True \/ P = False. 178 | Definition PE : Prop := 179 | forall P Q: Prop, P <-> Q -> P = Q. 180 | 181 | Fact TVS_XM_PE : 182 | TVS <-> XM /\ PE. 183 | Proof. 184 | split. 185 | - intros H. split. 186 | + intros P. specialize (H P) as [-> | ->]; tauto. 187 | + intros P Q. 188 | destruct (H P) as [-> | ->], (H Q) as [-> | ->]; tauto. 189 | - intros [H1 H2] P. 190 | specialize (H1 P) as [H1|H1]. 191 | + left. apply H2. tauto. 192 | + right. apply H2. tauto. 193 | Qed. 194 | --------------------------------------------------------------------------------