├── P1 ├── P1_solution.v └── P1_template.v ├── P2 ├── P2_solution.v └── P2_template.v ├── P3 ├── P3_solution.v └── P3_template.v ├── P4 ├── P4_solution.v └── P4_template.v └── P5 ├── P5_solution.v └── P5_template.v /P1/P1_solution.v: -------------------------------------------------------------------------------- 1 | Set Implicit Arguments. 2 | Require Import Omega. 3 | 4 | Definition decr(f : nat -> nat) := forall n, f (S n) <= f n. 5 | 6 | Definition valley(f : nat -> nat)(n x : nat) := forall y, x <= y -> y <= n+x -> f y = f x. 7 | 8 | Lemma decr_0 : forall f x, decr f -> f x = 0 -> forall y, x <= y -> f y = 0. 9 | Proof. 10 | intros. 11 | induction y. 12 | inversion H1. 13 | rewrite H2 in H0; exact H0. 14 | inversion H1. 15 | rewrite <- H2; exact H0. 16 | pose (H y). 17 | pose (IHy H3). 18 | omega. 19 | Qed. 20 | 21 | Lemma zero_valley : forall f n x, decr f -> f x = 0 -> valley f n x. 22 | Proof. 23 | intros f n x fd fx y Hxy _. 24 | rewrite fx. 25 | exact (decr_0 fd fx Hxy). 26 | Qed. 27 | 28 | Lemma arith_lemma : forall x y, x <= y -> x = y \/ S x <= y. 29 | Proof. 30 | intros. 31 | omega. 32 | Qed. 33 | 34 | Lemma valley_or_drop : forall n f x, decr f -> valley f n x \/ exists y, f y < f x. 35 | Proof. 36 | induction n; intros. 37 | left. 38 | intros y H1 H2. 39 | simpl in H2. 40 | assert (x = y). 41 | omega. 42 | rewrite H0; reflexivity. 43 | destruct (IHn _ (S x) H). 44 | pose (G := H x). 45 | inversion G. 46 | left. 47 | intros y Hy1 Hy2. 48 | destruct (arith_lemma Hy1). 49 | rewrite H1; reflexivity. 50 | rewrite <- H2. 51 | apply H0. 52 | exact H1. 53 | omega. 54 | right. 55 | exists (S x). 56 | omega. 57 | right. 58 | destruct H0 as [y Hy]. 59 | exists y. 60 | pose (G := H x). 61 | omega. 62 | Qed. 63 | 64 | Lemma valley_aux : forall y f x n, f x <= y -> decr f -> exists x', valley f n x'. 65 | Proof. 66 | induction y; intros. 67 | intros. 68 | exists x. 69 | apply zero_valley. 70 | exact H0. 71 | inversion H. 72 | reflexivity. 73 | destruct (valley_or_drop n x H0). 74 | exists x. 75 | exact H1. 76 | destruct H1 as [x' Hx']. 77 | assert (f x' <= y). 78 | omega. 79 | destruct (IHy f x' n H1 H0). 80 | exists x0; exact H2. 81 | Qed. 82 | 83 | Theorem decr_valleys : forall n f, decr f -> exists x, valley f n x. 84 | Proof. 85 | intros. 86 | exact (valley_aux 0 n (le_refl (f 0)) H). 87 | Qed. 88 | -------------------------------------------------------------------------------- /P1/P1_template.v: -------------------------------------------------------------------------------- 1 | Definition decr(f : nat -> nat) := forall n, f (S n) <= f n. 2 | 3 | Definition valley(f : nat -> nat)(n x : nat) := forall y, x <= y -> y <= n+x -> f y = f x. 4 | 5 | Theorem decr_valleys : forall n f, decr f -> exists x, valley f n x. 6 | Admitted. 7 | -------------------------------------------------------------------------------- /P2/P2_solution.v: -------------------------------------------------------------------------------- 1 | Set Implicit Arguments. 2 | Require Import Omega. 3 | Require Import Coq.Arith.Max. 4 | 5 | Definition LPO := forall f : nat -> bool, (exists x, f x = true) \/ (forall x, f x = false). 6 | 7 | Fixpoint true_below(f : nat -> bool)(x : nat) : bool := 8 | match x with 9 | |0 => f 0 10 | |S y => f x || true_below f y 11 | end. 12 | 13 | Lemma true_below_correct1 : forall f x, true_below f x = true -> exists y, f y = true. 14 | Proof. 15 | intros. 16 | induction x. 17 | exists 0. 18 | exact H. 19 | simpl in H. 20 | destruct (f (S x)) eqn:G. 21 | exists (S x); exact G. 22 | apply IHx. 23 | exact H. 24 | Qed. 25 | 26 | Lemma f_true_below : forall f x, f x = true -> true_below f x = true. 27 | Proof. 28 | intros. 29 | destruct x. 30 | exact H. 31 | simpl. 32 | rewrite H; reflexivity. 33 | Qed. 34 | 35 | Lemma true_below_correct2 : forall f x y, true_below f x = true -> x <= y -> true_below f y = true. 36 | Proof. 37 | intros f x. 38 | induction x. 39 | intros. 40 | induction y. 41 | exact H. 42 | simpl. 43 | rewrite IHy. 44 | destruct (f (S y)); reflexivity. 45 | omega. 46 | induction y. 47 | intros. 48 | inversion H0. 49 | intros. 50 | inversion H0. 51 | rewrite <- H2; exact H. 52 | simpl. 53 | rewrite IHy. 54 | destruct (f (S y)); reflexivity. 55 | exact H. 56 | exact H2. 57 | Qed. 58 | 59 | Definition to_nat(f : nat -> bool) : nat -> nat := 60 | fun x => if f x then 0 else 1. 61 | 62 | Definition decr(f : nat -> nat) := forall n, f (S n) <= f n. 63 | 64 | Lemma decr_shift : forall y f, decr f -> decr (fun x => f (x + y)). 65 | Proof. 66 | induction y; simpl; intros; intro x. 67 | apply H. 68 | simpl. 69 | apply H. 70 | Qed. 71 | 72 | Lemma decr_antitone : forall f, decr f -> forall x y, x <= y -> f y <= f x. 73 | Proof. 74 | intros. 75 | induction y. 76 | inversion H0. 77 | omega. 78 | inversion H0. 79 | omega. 80 | pose (IHy H2). 81 | pose (H y). 82 | omega. 83 | Qed. 84 | 85 | Lemma decr_0 : forall f, decr f -> f 0 = 0 -> forall x, f x = 0. 86 | Proof. 87 | intros. 88 | induction x. 89 | exact H0. 90 | pose (G := H x). 91 | omega. 92 | Qed. 93 | 94 | Lemma to_nat_decr : forall f, decr (to_nat (true_below f)). 95 | Proof. 96 | intros f x. 97 | unfold to_nat. 98 | simpl. 99 | destruct (f (S x)); simpl; omega. 100 | Qed. 101 | 102 | Definition infvalley(f : nat -> nat)(x : nat) := forall y, x <= y -> f y = f x. 103 | 104 | Theorem infvalley_LPO : (forall f, decr f -> exists x, infvalley f x) -> LPO. 105 | Proof. 106 | intros H f. 107 | destruct (H _ (to_nat_decr f)). 108 | unfold to_nat,infvalley in H0. 109 | destruct (true_below f x) eqn:G. 110 | left. 111 | exact (true_below_correct1 _ _ G). 112 | right. 113 | intro y. 114 | destruct (f y) eqn:fy. 115 | destruct (true_below f (max x y)) eqn:G1. 116 | pose (G2 := H0 _(le_max_l x y)). 117 | rewrite G1 in G2. 118 | discriminate G2. 119 | pose (G2 := f_true_below _ _ fy). 120 | rewrite (true_below_correct2 _ G2 (le_max_r x y)) in G1. 121 | exact G1. 122 | reflexivity. 123 | Qed. 124 | 125 | Definition Slt(f : nat -> nat) : nat -> bool := 126 | fun x => match lt_dec (f (S x)) (f x) with 127 | |left _ => true 128 | |right _ => false 129 | end. 130 | 131 | (*to avoid creating ill-typed terms*) 132 | Lemma obvious : forall f x, f (S x) < f x -> Slt f x = true. 133 | Proof. 134 | intros. 135 | unfold Slt. 136 | destruct (lt_dec (f (S x)) (f x)). 137 | reflexivity. 138 | contradiction. 139 | Qed. 140 | 141 | Lemma Slt_correct1 : forall f, decr f -> (forall x, Slt f x = false) -> forall x, f x = f 0. 142 | Proof. 143 | intros. 144 | induction x. 145 | reflexivity. 146 | destruct (lt_dec (f (S x)) (f x)). 147 | pose (G := H0 x). 148 | rewrite (obvious _ _ l) in G. 149 | discriminate G. 150 | pose (G := H x). 151 | omega. 152 | Qed. 153 | 154 | Lemma Slt_correct2 : forall f, decr f -> forall y, Slt f y = true -> f (S y) < f 0. 155 | Proof. 156 | intros. 157 | unfold Slt in H0. 158 | destruct (lt_dec (f (S y)) (f y)). 159 | assert (0 <= y). 160 | omega. 161 | pose (decr_antitone H H1). 162 | omega. 163 | discriminate H0. 164 | Qed. 165 | 166 | Lemma arith_lemma : forall x y z, x + y <= z -> exists u, z = u + y. 167 | Proof. 168 | intros. 169 | exists (z - y). 170 | omega. 171 | Qed. 172 | 173 | Lemma infvalley_shift : forall s f x, infvalley (fun x => f (x + s)) x -> infvalley f (x + s). 174 | Proof. 175 | intros s f x v y Hy. 176 | destruct (arith_lemma _ _ Hy) as [y' Hy']. 177 | rewrite Hy'. 178 | apply v. 179 | omega. 180 | Qed. 181 | 182 | Lemma infvalley_aux : LPO -> forall n f, f 0 <= n -> decr f -> exists x, infvalley f x. 183 | Proof. 184 | intros H n; induction n. 185 | intros. 186 | assert (f 0 = 0). 187 | omega. 188 | exists 0. 189 | intros x _. 190 | rewrite H2. 191 | rewrite (decr_0 H1 H2). 192 | reflexivity. 193 | intros. 194 | destruct (H (Slt f)). 195 | destruct H2. 196 | pose (Slt_correct2 H1 _ H2). 197 | assert (f (S x) <= n). 198 | omega. 199 | destruct (IHn _ H3 (decr_shift _ H1)). 200 | exists (x0 + S x). 201 | apply infvalley_shift. 202 | exact H4. 203 | exists 0. 204 | pose (Slt_correct1 H1 H2). 205 | intros y _. 206 | apply e. 207 | Qed. 208 | 209 | Theorem LPO_infvalley : LPO -> forall f, decr f -> exists x, infvalley f x. 210 | Proof. 211 | intros H f fd. 212 | exact (infvalley_aux H (le_refl (f 0)) fd). 213 | Qed. 214 | -------------------------------------------------------------------------------- /P2/P2_template.v: -------------------------------------------------------------------------------- 1 | Definition LPO := forall f : nat -> bool, (exists x, f x = true) \/ (forall x, f x = false). 2 | 3 | Definition decr(f : nat -> nat) := forall n, f (S n) <= f n. 4 | 5 | Definition infvalley(f : nat -> nat)(x : nat) := forall y, x <= y -> f y = f x. 6 | 7 | Theorem infvalley_LPO : (forall f, decr f -> exists x, infvalley f x) -> LPO. 8 | Admitted. 9 | 10 | Theorem LPO_infvalley : LPO -> forall f, decr f -> exists x, infvalley f x. 11 | Admitted. 12 | -------------------------------------------------------------------------------- /P3/P3_solution.v: -------------------------------------------------------------------------------- 1 | Definition inj{X Y}(f : X -> Y) := forall x x', f x = f x' -> x = x'. 2 | 3 | Definition surj{X Y}(f : X -> Y) := forall y, {x : X & f x = y}. 4 | 5 | Definition ded_fin(X : Set) := forall f : X -> X, inj f -> surj f. 6 | 7 | Section df_inh_cancel_sgroups. 8 | 9 | Variable X : Set. 10 | Variable x0 : X. 11 | Variable m : X -> X -> X. 12 | 13 | Infix "*" := m. 14 | 15 | Hypothesis X_df : ded_fin X. 16 | Hypothesis assoc : forall x y z, x * (y * z) = (x * y) * z. 17 | Hypothesis l_cancel : forall x y z, x * y = x * z -> y = z. 18 | Hypothesis r_cancel : forall x y z, y * x = z * x -> y = z. 19 | 20 | Lemma l_mult_inj : forall x, inj (m x). 21 | Proof. 22 | intros x y z H. 23 | exact (l_cancel _ _ _ H). 24 | Qed. 25 | 26 | Lemma r_mult_inj : forall x, inj (fun y => m y x). 27 | Proof. 28 | intros x y z H. 29 | exact (r_cancel _ _ _ H). 30 | Qed. 31 | 32 | Lemma r_eq_solve : forall a b, {x : X & a * x = b}. 33 | Proof. 34 | intros. 35 | destruct (X_df (m a) (l_mult_inj a) b). 36 | exists x; exact e. 37 | Qed. 38 | 39 | Lemma l_eq_solve : forall a b, {x : X & x * a = b}. 40 | Proof. 41 | intros. 42 | destruct (X_df (fun y => m y a) (r_mult_inj a) b). 43 | exists x; exact e. 44 | Qed. 45 | 46 | (* the identity *) 47 | Definition e : X := projT1 (l_eq_solve x0 x0). 48 | 49 | Lemma l_x0_id : e * x0 = x0. 50 | Proof. 51 | unfold e. 52 | destruct (l_eq_solve x0 x0). 53 | simpl. 54 | exact e0. 55 | Qed. 56 | 57 | Theorem l_id : forall x, e * x = x. 58 | Proof. 59 | intro. 60 | destruct (r_eq_solve x0 x) as [y Hy]. 61 | rewrite <- Hy. 62 | rewrite assoc. 63 | rewrite l_x0_id. 64 | reflexivity. 65 | Qed. 66 | 67 | Theorem r_id : forall x, x * e = x. 68 | Proof. 69 | intro. 70 | apply (r_cancel e). 71 | rewrite <- assoc. 72 | rewrite l_id. 73 | reflexivity. 74 | Qed. 75 | 76 | (* the inverse operations *) 77 | Definition inv : X -> X := fun x => projT1 (l_eq_solve x e). 78 | 79 | Theorem l_inv : forall x, (inv x) * x = e. 80 | Proof. 81 | intro. 82 | unfold inv. 83 | destruct (l_eq_solve x e). 84 | simpl. 85 | exact e0. 86 | Qed. 87 | 88 | Theorem r_inv : forall x, x * (inv x) = e. 89 | Proof. 90 | intro. 91 | apply (r_cancel x). 92 | rewrite <- assoc. 93 | rewrite l_inv. 94 | rewrite l_id, r_id. 95 | reflexivity. 96 | Qed. 97 | 98 | End df_inh_cancel_sgroups. -------------------------------------------------------------------------------- /P3/P3_template.v: -------------------------------------------------------------------------------- 1 | Definition inj{X Y}(f : X -> Y) := forall x x', f x = f x' -> x = x'. 2 | 3 | Definition surj{X Y}(f : X -> Y) := forall y, {x : X & f x = y}. 4 | 5 | Definition ded_fin(X : Set) := forall f : X -> X, inj f -> surj f. 6 | 7 | Section df_inh_cancel_sgroups. 8 | 9 | Variable X : Set. 10 | Variable x0 : X. 11 | Variable m : X -> X -> X. 12 | 13 | Infix "*" := m. 14 | 15 | Hypothesis X_df : ded_fin X. 16 | Hypothesis assoc : forall x y z, x * (y * z) = (x * y) * z. 17 | Hypothesis l_cancel : forall x y z, x * y = x * z -> y = z. 18 | Hypothesis r_cancel : forall x y z, y * x = z * x -> y = z. 19 | 20 | (* the identity *) 21 | Definition e : X. 22 | Admitted. 23 | 24 | Theorem l_id : forall x, e * x = x. 25 | Admitted. 26 | 27 | Theorem r_id : forall x, x * e = x. 28 | Admitted. 29 | 30 | (* the inverse operation *) 31 | Definition inv : X -> X. 32 | Admitted. 33 | 34 | Theorem l_inv : forall x, (inv x) * x = e. 35 | Admitted. 36 | 37 | Theorem r_inv : forall x, x * (inv x) = e. 38 | Admitted. 39 | 40 | End df_inh_cancel_sgroups. 41 | -------------------------------------------------------------------------------- /P4/P4_solution.v: -------------------------------------------------------------------------------- 1 | Set Implicit Arguments. 2 | 3 | Definition AC{X Y}(P : X -> Y -> Prop) : 4 | (forall x, {y : Y & P x y}) -> {f : X -> Y & forall x, P x (f x)}. 5 | Proof. 6 | intro. 7 | exists (fun x => projT1 (X0 x)). 8 | intro. 9 | destruct (X0 x); simpl. 10 | exact p. 11 | Qed. 12 | 13 | Definition inj{X Y}(f : X -> Y) := forall x x', f x = f x' -> x = x'. 14 | 15 | Definition surj{X Y}(f : X -> Y) := forall y, {x : X & f x = y}. 16 | 17 | Lemma inr_inj{X Y} : inj (@inr X Y). 18 | Proof. 19 | intros x x' H. 20 | inversion H. 21 | reflexivity. 22 | Qed. 23 | 24 | Fixpoint Fin(n : nat) : Set := 25 | match n with 26 | |0 => Empty_set 27 | |S m => unit + Fin m 28 | end. 29 | 30 | Lemma Fin_dec_eq : forall (n : nat)(i j : Fin n), {i = j} + {i <> j}. 31 | Proof. 32 | induction n; intros. 33 | destruct i. 34 | destruct i as [[]|i], j as [[]|j]. 35 | left; reflexivity. 36 | right; discriminate. 37 | right; discriminate. 38 | destruct (IHn i j). 39 | left; rewrite e; reflexivity. 40 | right; intro; apply n0. 41 | inversion H; reflexivity. 42 | Qed. 43 | 44 | Definition transpose(n : nat)(i j : Fin n) : Fin n -> Fin n := 45 | fun x => match Fin_dec_eq n i x with 46 | |left _ => j 47 | |right _ => match Fin_dec_eq n j x with 48 | |left _ => i 49 | |right _ => x 50 | end 51 | end. 52 | 53 | Lemma transpose_inj(n : nat) : forall i j, inj (transpose n i j). 54 | Proof. 55 | unfold transpose. 56 | intros i j x x' H. 57 | destruct (Fin_dec_eq n i x). 58 | destruct (Fin_dec_eq n i x'). 59 | transitivity i; auto. 60 | destruct (Fin_dec_eq n j x'). 61 | rewrite <- H in n0; contradiction. 62 | contradiction. 63 | destruct (Fin_dec_eq n j x). 64 | destruct (Fin_dec_eq n i x'). 65 | rewrite H in n0; contradiction. 66 | destruct (Fin_dec_eq n j x'). 67 | rewrite <- e; exact e0. 68 | contradiction. 69 | destruct (Fin_dec_eq n i x'). 70 | symmetry in H; contradiction. 71 | destruct (Fin_dec_eq n j x'). 72 | symmetry in H; contradiction. 73 | exact H. 74 | Qed. 75 | 76 | Lemma transpose_correct(n : nat) : forall i j x : Fin n, transpose n i j x = j -> i = x. 77 | Proof. 78 | unfold transpose; intros. 79 | destruct (Fin_dec_eq n i x). 80 | exact e. 81 | destruct (Fin_dec_eq n j x). 82 | transitivity j; auto. 83 | symmetry in H; contradiction. 84 | Qed. 85 | 86 | Theorem inj_to_surj(n : nat) : forall f : Fin n -> Fin n, inj f -> surj f. 87 | Proof. 88 | induction n; intros f finj. 89 | intro y; destruct y. 90 | assert (forall x, {y : Fin n & transpose (S n) (f (inl tt)) (inl tt) (f (inr x)) = inr y}). 91 | intro. 92 | destruct (transpose (S n) (f (inl tt)) (inl tt) (f (inr x))) as [[]|y] eqn:G. 93 | pose (transpose_correct _ _ _ G). 94 | absurd (inl tt = inr x). 95 | discriminate. 96 | exact (finj _ _ e). 97 | exists y; reflexivity. 98 | destruct (AC _ H) as [fhat Hf]. 99 | assert (inj fhat). 100 | intros x x' Hxx'. 101 | apply (@inr_inj unit). 102 | apply finj. 103 | apply (transpose_inj _ (f (inl tt)) (inl tt)). 104 | rewrite Hf; rewrite Hf. 105 | rewrite Hxx'; reflexivity. 106 | intro y. 107 | destruct (Fin_dec_eq (S n) (f (inl tt)) y). 108 | exists (inl tt); exact e. 109 | destruct (transpose _ (f (inl tt)) (inl tt) y) as [[]|y0] eqn:G. 110 | pose (transpose_correct _ _ _ G). 111 | contradiction. 112 | destruct (IHn _ H0 y0). 113 | exists (inr x). 114 | unfold transpose in G. 115 | rewrite <- e in G. 116 | rewrite <- Hf in G. 117 | destruct (Fin_dec_eq _ (f (inl tt)) y). 118 | contradiction. 119 | destruct (Fin_dec_eq (S n) (inl tt) y). 120 | unfold transpose in G. 121 | destruct (Fin_dec_eq (S n) (f (inl tt)) (f (inr x))). 122 | absurd (inl tt = inr x). 123 | discriminate. 124 | exact (finj _ _ e1). 125 | destruct (Fin_dec_eq (S n) (inl tt) (f (inr x))). 126 | rewrite <- e1; exact e0. 127 | contradiction. 128 | unfold transpose in G. 129 | destruct (Fin_dec_eq (S n) (f (inl tt)) (f (inr x))). 130 | symmetry in G; contradiction. 131 | destruct (Fin_dec_eq (S n) (inl tt) (f (inr x))). 132 | symmetry in G; contradiction. 133 | symmetry in G; auto. 134 | Qed. 135 | 136 | Lemma no_surj_n_Sn : forall (n : nat)(f : Fin n -> Fin (S n)), surj f -> False. 137 | Proof. 138 | induction n; intros. 139 | destruct (X (inl tt)). 140 | destruct x. 141 | destruct (f (inl tt)) as [[]|x] eqn:G. 142 | pose (g := fun x => match f (inr x) with 143 | |inl _ => inl tt 144 | |inr y => y 145 | end). 146 | apply (IHn g). 147 | intro y. 148 | destruct (X (inr y)). 149 | destruct x as [[]|x]. 150 | rewrite e in G; discriminate G. 151 | exists x. 152 | unfold g. 153 | rewrite e. 154 | reflexivity. 155 | pose (g := fun z => match f (inr z) with 156 | |inl _ => x 157 | |inr y => y 158 | end). 159 | apply (IHn g). 160 | intro y. 161 | destruct (X (inr y)). 162 | destruct x0 as [[]|x0]. 163 | destruct (X (inl tt)). 164 | destruct x0 as [[]|x0]. 165 | rewrite e0 in e; discriminate e. 166 | exists x0. 167 | unfold g. 168 | rewrite e0. 169 | rewrite e in G; inversion G; reflexivity. 170 | exists x0. 171 | unfold g. 172 | rewrite e. 173 | reflexivity. 174 | Qed. 175 | 176 | Lemma surj_to_inj(n : nat) : forall f : Fin n -> Fin n, surj f -> inj f. 177 | Proof. 178 | intros f fsurj x x' H. 179 | destruct n. 180 | destruct x. 181 | destruct (Fin_dec_eq _ x x'). 182 | exact e. 183 | destruct x as [[]|x]. 184 | destruct x' as [[]|x']. 185 | elim n0; reflexivity. 186 | pose (g := fun x => f (inr x)). 187 | assert (surj g). 188 | unfold g. 189 | intro y. 190 | destruct (fsurj y) as [[[]|x] fx]. 191 | exists x'. 192 | rewrite <- H; exact fx. 193 | exists x; exact fx. 194 | destruct (no_surj_n_Sn X). 195 | destruct x' as [[]|x']. 196 | pose (g := fun x => f (inr x)). 197 | assert (surj g). 198 | unfold g. 199 | intro y. 200 | destruct (fsurj y) as [[[]|x'] fx']. 201 | exists x; rewrite <- fx'; exact H. 202 | exists x'; exact fx'. 203 | destruct (no_surj_n_Sn X). 204 | pose (g := fun z => match Fin_dec_eq _ z x' with 205 | |left _ => f (inl tt) 206 | |right _ => f (inr z) 207 | end). 208 | assert (surj g). 209 | intro y. 210 | unfold g. 211 | destruct (Fin_dec_eq _ (f (inl tt)) y). 212 | exists x'. 213 | destruct (Fin_dec_eq n x' x'). 214 | exact e. 215 | elim n1; reflexivity. 216 | destruct (fsurj y) as [[[]|x''] fx'']. 217 | contradiction. 218 | destruct (Fin_dec_eq _ x'' x'). 219 | exists x. 220 | destruct (Fin_dec_eq n x x'). 221 | elim n0; rewrite e0; reflexivity. 222 | rewrite H; rewrite <- e; exact fx''. 223 | exists x''. 224 | destruct (Fin_dec_eq n x'' x'). 225 | contradiction. 226 | exact fx''. 227 | destruct (no_surj_n_Sn X). 228 | Qed. 229 | -------------------------------------------------------------------------------- /P4/P4_template.v: -------------------------------------------------------------------------------- 1 | Definition inj{X Y}(f : X -> Y) := forall x x', f x = f x' -> x = x'. 2 | 3 | Definition surj{X Y}(f : X -> Y) := forall y, {x : X & f x = y}. 4 | 5 | (* feel free to change this if you prefer other formalizations of Fin *) 6 | Fixpoint Fin(n : nat) : Set := 7 | match n with 8 | |0 => Empty_set 9 | |S m => unit + Fin m 10 | end. 11 | 12 | Theorem inj_to_surj(n : nat) : forall f : Fin n -> Fin n, inj f -> surj f. 13 | Admitted. 14 | 15 | Theorem surj_to_inj(n : nat) : forall f : Fin n -> Fin n, surj f -> inj f. 16 | Admitted. 17 | -------------------------------------------------------------------------------- /P5/P5_solution.v: -------------------------------------------------------------------------------- 1 | Set Implicit Arguments. 2 | Require Import Omega. 3 | Require Import Coq.Arith.Max. 4 | 5 | (* some arithmetical lemmas *) 6 | 7 | Lemma max_lemma_l : forall x y z, x <> S (max x y + z). 8 | Proof. 9 | induction x; intros. 10 | omega. 11 | destruct y. 12 | simpl; omega. 13 | simpl. 14 | pose (IHx y z). 15 | omega. 16 | Qed. 17 | 18 | Lemma max_lemma_r : forall x y z, y <> S (max x y + z). 19 | Proof. 20 | intros; rewrite max_comm. 21 | apply max_lemma_l. 22 | Qed. 23 | 24 | Lemma incr_global : forall f, (forall x, f x <= f (S x)) -> forall x y, x <= y -> f x <= f y. 25 | Proof. 26 | intros f H x. 27 | destruct x. 28 | induction y; intros. 29 | omega. 30 | assert (0 <= y). 31 | omega. 32 | pose (IHy H1). 33 | pose (H y). 34 | omega. 35 | intros. 36 | induction y. 37 | omega. 38 | inversion H0. 39 | omega. 40 | pose (IHy H2). 41 | pose (H y). 42 | omega. 43 | Qed. 44 | 45 | Lemma incr_global_strict : forall f, (forall x, f x < f (S x)) -> forall x y, x < y -> f x < f y. 46 | Proof. 47 | intros. 48 | assert (forall x, f x <= f (S x)). 49 | intro z; pose (H z). omega. 50 | assert (x <= y). omega. 51 | pose (incr_global _ H1). 52 | pose (l _ _ H2). 53 | inversion l0. 54 | inversion H0. 55 | rewrite <- H3 in H4. 56 | pose (H x). 57 | omega. 58 | rewrite <- H5 in H4. 59 | pose (H x). 60 | pose (l _ _ H3). 61 | pose (H m). 62 | omega. 63 | omega. 64 | Qed. 65 | 66 | Lemma incr_inj : forall f, (forall x, f x < f (S x)) -> forall x y, x <> y -> f x <> f y. 67 | Proof. 68 | intros. 69 | destruct (lt_dec x y). 70 | pose (incr_global_strict _ H l). 71 | omega. 72 | assert (y < x). 73 | omega. 74 | pose (incr_global_strict _ H H1). 75 | omega. 76 | Qed. 77 | 78 | (* streamless stuff *) 79 | 80 | Definition streamless(X : Set) := forall f : nat -> X, 81 | {i : nat & {j : nat & i <> j /\ f i = f j}}. 82 | 83 | Definition hat(X : Set)(x0 : X)(f : nat -> unit + X) : nat -> X := 84 | fun n => match f n with 85 | |inl _ => x0 86 | |inr x => x 87 | end. 88 | 89 | Lemma streamless_lemma : forall X, streamless X -> forall (x0 : X)(i j : nat)(f : nat -> unit + X), 90 | f i = inl tt -> f j = inr x0 -> {i0 : nat & {j0 : nat & i0 <> j0 /\ f i0 = f j0}}. 91 | Proof. 92 | intros X Xstr x0 i j f fi fj. 93 | destruct (Xstr (fun n => hat x0 f (S (max i j + n)))) as [k [l [klneq Hkl]]]. 94 | unfold hat in Hkl. 95 | destruct (f (S (max i j + k))) as [[]|xk] eqn:fk. 96 | exists i,(S (max i j + k)). 97 | split. 98 | apply max_lemma_l. 99 | rewrite fk; exact fi. 100 | destruct (f (S (max i j + l))) as [[]|xl] eqn:fl. 101 | exists i,(S (max i j + l)). 102 | split. 103 | apply max_lemma_l. 104 | rewrite fl; exact fi. 105 | exists (S (max i j + k)),(S (max i j + l)). 106 | split. 107 | omega. 108 | rewrite fl; rewrite <- Hkl; exact fk. 109 | Qed. 110 | 111 | Lemma streamless_plus_one_inh : forall X, streamless X -> X -> streamless (unit + X). 112 | Proof. 113 | intros X Xstr x0 f. 114 | destruct (Xstr (hat x0 f)) as [i [j [ijneq gijeq]]]. 115 | unfold hat in gijeq. 116 | destruct (f i) as [[]|xi] eqn:fi. 117 | destruct (f j) as [[]|xj] eqn:fj. 118 | exists i,j. 119 | split. 120 | exact ijneq. 121 | rewrite fj; exact fi. 122 | exact (streamless_lemma Xstr _ _ _ fi fj). 123 | destruct (f j) as [[]|xj] eqn:fj. 124 | exact (streamless_lemma Xstr _ _ _ fj fi). 125 | exists i,j. 126 | split. 127 | exact ijneq. 128 | rewrite fj; rewrite <- gijeq; exact fi. 129 | Qed. 130 | 131 | Lemma streamless_plus_one : forall X, streamless X -> streamless (unit + X). 132 | Proof. 133 | intros X Xstr f. 134 | destruct (f 0) as [[]|x0] eqn:f0. 135 | destruct (f 1) as [[]|x1] eqn:f1. 136 | exists 0,1; split. 137 | discriminate. 138 | rewrite f1; exact f0. 139 | exact (streamless_plus_one_inh Xstr x1 _). 140 | exact (streamless_plus_one_inh Xstr x0 _). 141 | Qed. 142 | 143 | Definition hat_l(X Y : Set)(f : nat -> X + Y) : nat -> unit + X := 144 | fun n => match f n with 145 | |inl x => inr x 146 | |inr _ => inl tt 147 | end. 148 | 149 | Definition hat_r(X Y : Set)(f : nat -> X + Y) : nat -> unit + Y := 150 | fun n => match f n with 151 | |inl _ => inl tt 152 | |inr y => inr y 153 | end. 154 | 155 | Lemma str_lt_wlog : forall X, streamless X -> forall f : nat -> X, 156 | {i : nat & {j : nat & i < j /\ f i = f j}}. 157 | Proof. 158 | intros X Xstr f. 159 | destruct (Xstr f) as [i [j [ijneq Hij]]]. 160 | destruct (lt_dec i j). 161 | exists i,j. 162 | split. 163 | exact l. 164 | exact Hij. 165 | exists j,i. 166 | split. 167 | omega. 168 | symmetry; exact Hij. 169 | Qed. 170 | 171 | Lemma str_lt : forall (X : Set)(n : nat)(f : nat -> X), streamless X -> 172 | {i : nat & {j : nat & n < i /\ i < j /\ f i = f j}}. 173 | Proof. 174 | intros X n f Xstr. 175 | destruct (str_lt_wlog Xstr (fun m => f (S (m + n)))) as [i [j [ijleq Hij]]]. 176 | exists (S (i + n)),(S (j + n)). 177 | split. 178 | omega. 179 | split. 180 | omega. 181 | exact Hij. 182 | Qed. 183 | 184 | Fixpoint subseq(X : Set)(Xstr : streamless X)(f : nat -> X)(n : nat) : nat := 185 | match n with 186 | |0 => projT1 (str_lt 0 f Xstr) 187 | |S m => projT1 (str_lt (subseq Xstr f m) f Xstr) 188 | end. 189 | 190 | Lemma subseq_incr : forall (X : Set)(Xstr : streamless X)(f : nat -> X)(n : nat), 191 | subseq Xstr f n < subseq Xstr f (S n). 192 | Proof. 193 | intros. 194 | simpl. 195 | destruct (str_lt (subseq Xstr f n) f Xstr) as [i [j [H1 [H2 H3]]]]. 196 | simpl. 197 | omega. 198 | Qed. 199 | 200 | Theorem streamless_sum : forall X Y, streamless X -> streamless Y -> streamless (X + Y). 201 | Proof. 202 | intros X Y Xstr Ystr f. 203 | pose (is := subseq (streamless_plus_one Xstr) (hat_l f)). 204 | destruct (streamless_plus_one Ystr (fun x => hat_r f (is x))) as [k [l [klneq fkl]]]. 205 | unfold hat_r in fkl. 206 | destruct (f (is k)) as [xk|yk] eqn:fi. 207 | unfold is in fi. 208 | destruct k. 209 | simpl in fi. 210 | destruct (str_lt 0 (hat_l f) (streamless_plus_one Xstr)) as [i [j [ijneq [H1 H2]]]]. 211 | simpl in fi. 212 | unfold hat_l in H2. 213 | rewrite fi in H2. 214 | destruct (f j) as [xj|yj] eqn:fj. 215 | exists i,j. 216 | split. 217 | omega. 218 | rewrite fi,fj. 219 | inversion H2; reflexivity. 220 | discriminate H2. 221 | simpl in fi. 222 | destruct (str_lt (subseq (streamless_plus_one Xstr) (hat_l f) k) (hat_l f) (streamless_plus_one Xstr)) 223 | as [i [j [ijneq [H1 H2]]]]. 224 | simpl in fi. 225 | unfold hat_l in H2. 226 | rewrite fi in H2. 227 | destruct (f j) as [xj|yj] eqn:fj. 228 | exists i,j. 229 | split. 230 | omega. 231 | rewrite fi,fj. 232 | inversion H2; reflexivity. 233 | discriminate H2. 234 | destruct (f (is l)) as [xl|yl] eqn:fj. 235 | discriminate fkl. 236 | exists (is k),(is l). 237 | split. 238 | apply incr_inj. 239 | apply subseq_incr. 240 | exact klneq. 241 | rewrite fi,fj. 242 | inversion fkl; reflexivity. 243 | Qed. 244 | -------------------------------------------------------------------------------- /P5/P5_template.v: -------------------------------------------------------------------------------- 1 | Definition streamless(X : Set) := forall f : nat -> X, 2 | {i : nat & {j : nat & i <> j /\ f i = f j}}. 3 | 4 | Theorem streamless_sum : forall X Y, streamless X -> streamless Y -> streamless (X + Y). 5 | Admitted. 6 | --------------------------------------------------------------------------------