diff --git a/coq/src/putnam_1994_b3.v b/coq/src/putnam_1994_b3.v index a5b53c1..cd8ea93 100644 --- a/coq/src/putnam_1994_b3.v +++ b/coq/src/putnam_1994_b3.v @@ -1,6 +1,18 @@ -Require Import Reals Coquelicot.Coquelicot. -Open Scope R. -Definition putnam_1994_b3_solution (k: R) := k < 1. -Theorem putnam_1994_b3 - : forall k : R, (forall f : R -> R, (forall x : R, f x > 0 /\ ex_derive f x /\ (Derive f) x > f x) -> exists N : R, forall x : R, x > N -> f x > exp (k * x)) <-> putnam_1994_b3_solution k. -Proof. Admitted. +From mathcomp Require Import all_ssreflect ssralg ssrnum. +From mathcomp Require Import reals normedtype derive topology sequences. +From mathcomp Require Import classical_sets. +Import numFieldNormedType.Exports. + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Local Open Scope ring_scope. +Local Open Scope classical_set_scope. + +Variable R : realType. +Definition putnam_1993_b3_solution : set R := [set k | k < 1]. +Theorem putnam_1993_b3 + : [set k | forall f (hf : forall x, differentiable f x /\ 0 < f x < f^`() x), + exists N : R, forall x, N < x -> expR (k * x) < f x] = putnam_1993_b3_solution. +Proof. Admitted. \ No newline at end of file diff --git a/lean4/src/putnam_1963_b3.lean b/lean4/src/putnam_1963_b3.lean index 69511d4..d69c049 100644 --- a/lean4/src/putnam_1963_b3.lean +++ b/lean4/src/putnam_1963_b3.lean @@ -8,8 +8,8 @@ abbrev putnam_1963_b3_solution : Set (ℝ → ℝ) := sorry Find every twice-differentiable real-valued function $f$ with domain the set of all real numbers and satisfying the functional equation $(f(x))^2-(f(y))^2=f(x+y)f(x-y)$ for all real numbers $x$ and $y$. -/ theorem putnam_1963_b3 -(f : ℝ → ℝ) -(fdiff : Prop) -(hfdiff : fdiff ↔ ContDiff ℝ 1 f ∧ Differentiable ℝ (deriv f)) -: (fdiff ∧ ∀ x y : ℝ, (f x) ^ 2 - (f y) ^ 2 = f (x + y) * f (x - y)) ↔ f ∈ putnam_1963_b3_solution := -sorry + (f : ℝ → ℝ) : + f ∈ putnam_1963_b3_solution ↔ + (ContDiff ℝ 1 f ∧ Differentiable ℝ (deriv f) ∧ + ∀ x y : ℝ, (f x) ^ 2 - (f y) ^ 2 = f (x + y) * f (x - y)) := + sorry diff --git a/lean4/src/putnam_1964_a5.lean b/lean4/src/putnam_1964_a5.lean index 53b2861..455684f 100644 --- a/lean4/src/putnam_1964_a5.lean +++ b/lean4/src/putnam_1964_a5.lean @@ -9,7 +9,8 @@ Prove that there exists a constant $k$ such that for any sequence $a_i$ of posit \] -/ theorem putnam_1964_a5 -(pa : (ℕ → ℝ) → Prop) -(hpa : pa = fun a ↦ (∀ n : ℕ, a n > 0) ∧ ∃ L : ℝ, Tendsto (fun N ↦ ∑ n in Finset.range N, 1 / a n) atTop (𝓝 L)) -: (∃ k : ℝ, ∀ a : ℕ → ℝ, pa a → ∑' n : ℕ, (n + 1) / (∑ i in Finset.range (n + 1), a i) ≤ k * ∑' n : ℕ, 1 / a n) := -sorry + (pa : (ℕ → ℝ) → Prop) + (hpa : ∀ a, pa a ↔ (∀ n : ℕ, a n > 0) ∧ ∃ L : ℝ, Tendsto (fun N ↦ ∑ n in Finset.range N, 1 / a n) atTop (𝓝 L)) : + ∃ k : ℝ, ∀ a : ℕ → ℝ, pa a → + ∑' n : ℕ, (n + 1) / (∑ i in Finset.range (n + 1), a i) ≤ k * ∑' n : ℕ, 1 / a n := + sorry diff --git a/lean4/src/putnam_1964_b6.lean b/lean4/src/putnam_1964_b6.lean index 32f70a5..e919f1d 100644 --- a/lean4/src/putnam_1964_b6.lean +++ b/lean4/src/putnam_1964_b6.lean @@ -6,9 +6,9 @@ open Set Function Filter Topology Let $D$ be the unit disk in the plane. Show that we cannot find congruent sets $A, B$ with $A \cap B = \emptyset$ and $A \cup B = D$. -/ theorem putnam_1964_b6 -(D : Set (EuclideanSpace ℝ (Fin 2))) -(hD : D = {v : EuclideanSpace ℝ (Fin 2) | dist 0 v ≤ 1}) -(cong : Set (EuclideanSpace ℝ (Fin 2)) → Set (EuclideanSpace ℝ (Fin 2)) → Prop) -(hcong : cong = fun A B ↦ ∃ f : (EuclideanSpace ℝ (Fin 2)) → (EuclideanSpace ℝ (Fin 2)), B = f '' A ∧ ∀ v w : EuclideanSpace ℝ (Fin 2), dist v w = dist (f v) (f w)) -: (¬∃ A B : Set (Fin 2 → ℝ), cong A B ∧ A ∩ B = ∅ ∧ A ∪ B = D) := -sorry + (D : Set (EuclideanSpace ℝ (Fin 2))) + (hD : D = {v : EuclideanSpace ℝ (Fin 2) | dist 0 v ≤ 1}) + (cong : Set (EuclideanSpace ℝ (Fin 2)) → Set (EuclideanSpace ℝ (Fin 2)) → Prop) + (hcong : ∀ A B, cong A B ↔ ∃ f : (EuclideanSpace ℝ (Fin 2)) → (EuclideanSpace ℝ (Fin 2)), B = f '' A ∧ ∀ v w : EuclideanSpace ℝ (Fin 2), dist v w = dist (f v) (f w)) + : (¬∃ A B : Set (Fin 2 → ℝ), cong A B ∧ A ∩ B = ∅ ∧ A ∪ B = D) := + sorry diff --git a/lean4/src/putnam_1965_b6.lean b/lean4/src/putnam_1965_b6.lean index 4a9abab..ce77717 100644 --- a/lean4/src/putnam_1965_b6.lean +++ b/lean4/src/putnam_1965_b6.lean @@ -6,13 +6,14 @@ open EuclideanGeometry Topology Filter Complex SimpleGraph.Walk Let $A$, $B$, $C$, and $D$ be four distinct points for which every circle through $A$ and $B$ intersects every circle through $C$ and $D$. Prove that $A$, $B$, $C$ and $D$ are either collinear (all lying on the same line) or cocyclic (all lying on the same circle). -/ theorem putnam_1965_b6 -(A B C D : EuclideanSpace ℝ (Fin 2)) -(S : Set (EuclideanSpace ℝ (Fin 2))) -(hS : S = {A, B, C, D}) -(hdistinct : S.ncard = 4) -(through : (ℝ × (EuclideanSpace ℝ (Fin 2))) → (EuclideanSpace ℝ (Fin 2)) → Prop) -(through_def : through = fun (r, P) => fun Q => dist P Q = r) -(hABCD : ∀ r s : ℝ, ∀ P Q : EuclideanSpace ℝ (Fin 2), through (r, P) A ∧ through (r, P) B ∧ through (s, Q) C ∧ through (s, Q) D → -∃ I : EuclideanSpace ℝ (Fin 2), through (r, P) I ∧ through (s, Q) I) -: Collinear ℝ S ∨ ∃ r : ℝ, ∃ P : EuclideanSpace ℝ (Fin 2), ∀ Q ∈ S, through (r, P) Q := -sorry + (A B C D : EuclideanSpace ℝ (Fin 2)) + (S : Set (EuclideanSpace ℝ (Fin 2))) + (hS : S = {A, B, C, D}) + (hdistinct : S.ncard = 4) + (through : (ℝ × (EuclideanSpace ℝ (Fin 2))) → (EuclideanSpace ℝ (Fin 2)) → Prop) + (through_def : through = fun (r, P) => fun Q => dist P Q = r) + (hABCD : ∀ r s : ℝ, ∀ P Q, + through (r, P) A ∧ through (r, P) B ∧ through (s, Q) C ∧ through (s, Q) D → + ∃ I, through (r, P) I ∧ through (s, Q) I) : + Collinear ℝ S ∨ ∃ r : ℝ, ∃ P, ∀ Q ∈ S, through (r, P) Q := + sorry diff --git a/lean4/src/putnam_1967_a1.lean b/lean4/src/putnam_1967_a1.lean index 66e062e..246d1cb 100644 --- a/lean4/src/putnam_1967_a1.lean +++ b/lean4/src/putnam_1967_a1.lean @@ -6,11 +6,10 @@ open Nat Topology Filter Let $f(x)=a_1\sin x+a_2\sin 2x+\dots+a_n\sin nx$, where $a_1,a_2,\dots,a_n$ are real numbers and where $n$ is a positive integer. Given that $|f(x)| \leq |\sin x|$ for all real $x$, prove that $|a_1|+|2a_2|+\dots+|na_n| \leq 1$. -/ theorem putnam_1967_a1 -(n : ℕ) +(n : ℕ) (hn : n > 0) (a : Set.Icc 1 n → ℝ) (f : ℝ → ℝ) (hf : f = (fun x : ℝ => ∑ i : Set.Icc 1 n, a i * Real.sin (i * x))) -(npos : n > 0) (flesin : ∀ x : ℝ, abs (f x) ≤ abs (Real.sin x)) : abs (∑ i : Set.Icc 1 n, i * a i) ≤ 1 := sorry diff --git a/lean4/src/putnam_1967_a2.lean b/lean4/src/putnam_1967_a2.lean index f62f4ef..a36b685 100644 --- a/lean4/src/putnam_1967_a2.lean +++ b/lean4/src/putnam_1967_a2.lean @@ -10,8 +10,9 @@ Define $S_0$ to be $1$. For $n \geq 1$, let $S_n$ be the number of $n \times n$ \end{enumerate} -/ theorem putnam_1967_a2 -(S : ℕ → ℤ) -(hS0 : S 0 = 1) -(hSn : ∀ n ≥ 1, S n = {A : Matrix (Fin n) (Fin n) ℕ | (∀ i j : Fin n, A i j = A j i) ∧ (∀ j : Fin n, (∑ i : Fin n, A i j) = 1)}.ncard) -: (∀ n ≥ 1, S (n + 1) = S n + n * S (n - 1)) ∧ (∀ x : ℝ, (∑' n : ℕ, S n * (x ^ n / (n)!)) = Real.exp (x + x ^ 2 / 2)) := -sorry + (S : ℕ → ℤ) + (hS0 : S 0 = 1) + (hSn : ∀ n ≥ 1, S n = {A : Matrix (Fin n) (Fin n) ℕ | (∀ i j, A i j = A j i) ∧ (∀ j, (∑ i, A i j) = 1)}.ncard) : + (∀ n ≥ 1, S (n + 1) = S n + n * S (n - 1)) ∧ + (∀ x : ℝ, (∑' n : ℕ, S n * (x ^ n / (n)!)) = Real.exp (x + x ^ 2 / 2)) := + sorry diff --git a/lean4/src/putnam_1967_a3.lean b/lean4/src/putnam_1967_a3.lean index cc30318..21f987f 100644 --- a/lean4/src/putnam_1967_a3.lean +++ b/lean4/src/putnam_1967_a3.lean @@ -1,16 +1,17 @@ import Mathlib -open Nat Topology Filter +open Polynomial abbrev putnam_1967_a3_solution : ℕ := sorry -- 5 /-- Consider polynomial forms $ax^2-bx+c$ with integer coefficients which have two distinct zeros in the open interval $0 0) -: (∃ p : Polynomial ℝ, pall p ∧ p.coeff 2 = putnam_1967_a3_solution) ∧ (∀ p : Polynomial ℝ, pall p → p.coeff 2 ≥ putnam_1967_a3_solution) := -sorry +theorem putnam_1967_a3 : + IsLeast + {a | ∃ P : Polynomial ℤ, + P.degree = 2 ∧ + (∃ z1 z2 : Set.Ioo (0 : ℝ) 1, z1 ≠ z2 ∧ aeval (z1 : ℝ) P = 0 ∧ aeval (z2 : ℝ) P = 0) ∧ + P.coeff 2 = a ∧ a > 0} + putnam_1967_a3_solution := + sorry diff --git a/lean4/src/putnam_1970_a4.lean b/lean4/src/putnam_1970_a4.lean index 4a68a3c..87aac2f 100644 --- a/lean4/src/putnam_1970_a4.lean +++ b/lean4/src/putnam_1970_a4.lean @@ -7,6 +7,6 @@ Suppose $(x_n)$ is a sequence such that $\lim_{n \to \infty} (x_n - x_{n-2} = 0$ -/ theorem putnam_1970_a4 (x : ℕ → ℝ) -(hxlim : Tendsto (fun n => x n - x (n-2)) atTop (𝓝 0)) -: Tendsto (fun n => (x n - x (n-1))/n) atTop (𝓝 0) := +(hxlim : Tendsto (fun n => x (n+2) - x n) atTop (𝓝 0)) +: Tendsto (fun n => (x (n+1) - x (n))/(n+1)) atTop (𝓝 0) := sorry diff --git a/lean4/src/putnam_1971_a2.lean b/lean4/src/putnam_1971_a2.lean index 6aa00af..5e57db7 100644 --- a/lean4/src/putnam_1971_a2.lean +++ b/lean4/src/putnam_1971_a2.lean @@ -8,5 +8,6 @@ abbrev putnam_1971_a2_solution : Set (Polynomial ℝ) := sorry Determine all polynomials $P(x)$ such that $P(x^2 + 1) = (P(x))^2 + 1$ and $P(0) = 0$. -/ theorem putnam_1971_a2 -: ∀ P : Polynomial ℝ, (P.eval 0 = 0 ∧ (∀ x : ℝ, P.eval (x^2 + 1) = (P.eval x)^2 + 1)) ↔ P ∈ putnam_1971_a2_solution := -sorry + (P : Polynomial ℝ) : + (P.eval 0 = 0 ∧ (∀ x : ℝ, P.eval (x^2 + 1) = (P.eval x)^2 + 1)) ↔ P ∈ putnam_1971_a2_solution := + sorry diff --git a/lean4/src/putnam_1971_a3.lean b/lean4/src/putnam_1971_a3.lean index 0b0f0a8..f11d9ae 100644 --- a/lean4/src/putnam_1971_a3.lean +++ b/lean4/src/putnam_1971_a3.lean @@ -6,13 +6,14 @@ open Set The three vertices of a triangle of sides $a,b,c$ are lattice points and lie on a circle of radius $R$. Show that $abc \geq 2R$. -/ theorem putnam_1971_a3 -(a b c : ℝ × ℝ) -(R : ℝ) -(habclattice : a.1 = round a.1 ∧ a.2 = round a.2 ∧ b.1 = round b.1 ∧ b.2 = round b.2 ∧ c.1 = round c.1 ∧ c.2 = round c.2) -(habcneq : a ≠ b ∧ a ≠ c ∧ b ≠ c) -(hR : R > 0) -(oncircle : (ℝ × ℝ) → ℝ → (ℝ × ℝ) → Prop) -(oncircle_def : oncircle = fun (C : ℝ × ℝ) (r : ℝ) (p : ℝ × ℝ) => Real.sqrt ((p.1 - C.1)^2 + (p.2 - C.2)^2) = r) -(hcircle : ∃ C : ℝ × ℝ, oncircle C R a ∧ oncircle C R b ∧ oncircle C R c) -: (Real.sqrt ((a.1 - b.1)^2 + (a.2 - b.2)^2)) * (Real.sqrt ((a.1 - c.1)^2 + (a.2 - c.2)^2)) * (Real.sqrt ((b.1 - c.1)^2 + (b.2 - c.2)^2)) ≥ 2 * R := -sorry + (a b c : ℝ × ℝ) + (R : ℝ) + (habclattice : a.1 = round a.1 ∧ a.2 = round a.2 ∧ b.1 = round b.1 ∧ b.2 = round b.2 ∧ c.1 = round c.1 ∧ c.2 = round c.2) + (habcneq : a ≠ b ∧ a ≠ c ∧ b ≠ c) + (hR : R > 0) + (hcircle : ∃ C : ℝ × ℝ, + √((a.1 - C.1)^2 + (a.2 - C.2)^2) = R ∧ + √((b.1 - C.1)^2 + (b.2 - C.2)^2) = R ∧ + √((c.1 - C.1)^2 + (c.2 - C.2)^2) = R) : + (√((a.1 - b.1)^2 + (a.2 - b.2)^2)) * (√((a.1 - c.1)^2 + (a.2 - c.2)^2)) * (√((b.1 - c.1)^2 + (b.2 - c.2)^2)) ≥ 2 * R := + sorry diff --git a/lean4/src/putnam_1971_a5.lean b/lean4/src/putnam_1971_a5.lean index a134900..8765e15 100644 --- a/lean4/src/putnam_1971_a5.lean +++ b/lean4/src/putnam_1971_a5.lean @@ -8,9 +8,11 @@ abbrev putnam_1971_a5_solution : ℤ × ℤ := sorry After each play of a certain game of solitaire, the player receives either $a$ or $b$ points, where $a$ and $b$ are positive integers with $a > b$; scores accumulate from play to play. If there are $35$ unattainable scores, one of which is $58$, find $a$ and $b$. -/ theorem putnam_1971_a5 -(a b : ℤ) -(hab : a > 0 ∧ b > 0 ∧ a > b) -(pab : ℤ → ℤ → Prop) -(hpab : pab = fun x y : ℤ ↦ {s : ℕ | ¬∃ m n : ℕ, m*x + n*y = s}.ncard = 35 ∧ ¬∃ m n : ℕ, m*x + n*y = 58) -: pab a b ↔ a = putnam_1971_a5_solution.1 ∧ b = putnam_1971_a5_solution.2 := -sorry + (a b : ℤ) + (hab : a > 0 ∧ b > 0 ∧ a > b) + (pab : ℤ → ℤ → Prop) + (hpab : ∀ x y, pab x y ↔ + {s : ℕ | ¬∃ m n : ℕ, m*x + n*y = s}.ncard = 35 ∧ + ¬∃ m n : ℕ, m*x + n*y = 58) : + pab a b ↔ a = putnam_1971_a5_solution.1 ∧ b = putnam_1971_a5_solution.2 := + sorry diff --git a/lean4/src/putnam_1972_a1.lean b/lean4/src/putnam_1972_a1.lean index 7d4ffdf..ff30bea 100644 --- a/lean4/src/putnam_1972_a1.lean +++ b/lean4/src/putnam_1972_a1.lean @@ -6,9 +6,8 @@ open EuclideanGeometry Filter Topology Set Show that there are no four consecutive binomial coefficients ${n \choose r}, {n \choose (r+1)}, {n \choose (r+2)}, {n \choose (r+3)}$ where $n,r$ are positive integers and $r+3 \leq n$, which are in arithmetic progression. -/ theorem putnam_1972_a1 -(n : ℕ) -(hn : n > 0) -(fourAP : ℤ → ℤ → ℤ → ℤ → Prop) -(hfourAP : fourAP = fun n1 n2 n3 n4 => n4-n3 = n3-n2 ∧ n3-n2 = n2-n1) -: ¬ ∃ r : ℕ, r > 0 ∧ r + 3 ≤ n ∧ fourAP (n.choose r) (n.choose (r+1)) (n.choose (r+2)) (n.choose (r+3)) := -sorry + (n : ℕ) (hn : n > 0) + (fourAP : ℤ → ℤ → ℤ → ℤ → Prop) + (hfourAP : ∀ n1 n2 n3 n4, fourAP n1 n2 n3 n4 ↔ n4-n3 = n3-n2 ∧ n3-n2 = n2-n1) : + ¬ ∃ r : ℕ, r > 0 ∧ r + 3 ≤ n ∧ fourAP (n.choose r) (n.choose (r+1)) (n.choose (r+2)) (n.choose (r+3)) := + sorry diff --git a/lean4/src/putnam_1972_a3.lean b/lean4/src/putnam_1972_a3.lean index cac3468..3243ff2 100644 --- a/lean4/src/putnam_1972_a3.lean +++ b/lean4/src/putnam_1972_a3.lean @@ -9,9 +9,9 @@ abbrev putnam_1972_a3_solution : Set (ℝ → ℝ) := sorry We call a function $f$ from $[0,1]$ to the reals to be supercontinuous on $[0,1]$ if the Cesaro-limit exists for the sequence $f(x_1), f(x_2), f(x_3), \dots$ whenever it does for the sequence $x_1, x_2, x_3 \dots$. Find all supercontinuous functions on $[0,1]$. -/ theorem putnam_1972_a3 -(climit_exists : (ℕ → ℝ) → Prop) -(hclimit_exists : climit_exists = fun x => ∃ C : ℝ, Tendsto (fun n => (∑ i in Finset.range n, (x i))/(n : ℝ)) atTop (𝓝 C)) -(supercontinuous : (ℝ → ℝ) → Prop) -(hsupercontinuous : supercontinuous = fun f => ∀ (x : ℕ → ℝ), (∀ i : ℕ, x i ∈ Icc 0 1) → climit_exists x → climit_exists (fun i => f (x i))) -: {f | supercontinuous f} = putnam_1972_a3_solution := -sorry + (climit_exists : (ℕ → ℝ) → Prop) + (supercontinuous : (ℝ → ℝ) → Prop) + (hclimit_exists : ∀ x, climit_exists x ↔ ∃ C : ℝ, Tendsto (fun n => (∑ i in Finset.range n, (x i))/(n : ℝ)) atTop (𝓝 C)) + (hsupercontinuous : ∀ f, supercontinuous f ↔ ∀ (x : ℕ → ℝ), (∀ i : ℕ, x i ∈ Icc 0 1) → climit_exists x → climit_exists (fun i => f (x i))) : + {f | supercontinuous f} = putnam_1972_a3_solution := + sorry diff --git a/lean4/src/putnam_1972_b2.lean b/lean4/src/putnam_1972_b2.lean index f5cbab9..7c13b1e 100644 --- a/lean4/src/putnam_1972_b2.lean +++ b/lean4/src/putnam_1972_b2.lean @@ -8,13 +8,13 @@ noncomputable abbrev putnam_1972_b2_solution : ℝ → ℝ → ℝ := sorry Let $x : \mathbb{R} \to \mathbb{R}$ be a twice differentiable function whose second derivative is nonstrictly decreasing. If $x(t) - x(0) = s$, $x'(0) = 0$, and $x'(t) = v$ for some $t > 0$, find the maximum possible value of $t$ in terms of $s$ and $v$. -/ theorem putnam_1972_b2 -(s v : ℝ) -(hs : s > 0) -(hv : v > 0) -(valid : ℝ → (ℝ → ℝ) → Prop) -(hvalid : valid = fun (t : ℝ) (x : ℝ → ℝ) ↦ -DifferentiableOn ℝ x (Set.Icc 0 t) ∧ DifferentiableOn ℝ (deriv x) (Set.Icc 0 t) - ∧ AntitoneOn (deriv (deriv x)) (Set.Icc 0 t) ∧ deriv x 0 = 0 ∧ deriv x t = v ∧ x t - x 0 = s) -: (∃ x : ℝ → ℝ, valid (putnam_1972_b2_solution s v) x) ∧ -∀ t > putnam_1972_b2_solution s v, ¬(∃ x : ℝ → ℝ, valid t x) := -sorry + (s v : ℝ) + (hs : s > 0) + (hv : v > 0) + (valid : ℝ → (ℝ → ℝ) → Prop) + (hvalid : ∀ t x, valid t x ↔ + DifferentiableOn ℝ x (Set.Icc 0 t) ∧ DifferentiableOn ℝ (deriv x) (Set.Icc 0 t) ∧ + AntitoneOn (deriv (deriv x)) (Set.Icc 0 t) ∧ + deriv x 0 = 0 ∧ deriv x t = v ∧ x t - x 0 = s) + : IsGreatest {t | ∃ x : ℝ → ℝ, valid t x} (putnam_1972_b2_solution s v) := + sorry diff --git a/lean4/src/putnam_1974_a1.lean b/lean4/src/putnam_1974_a1.lean index e685ac6..f00c909 100644 --- a/lean4/src/putnam_1974_a1.lean +++ b/lean4/src/putnam_1974_a1.lean @@ -8,7 +8,7 @@ abbrev putnam_1974_a1_solution : ℕ := sorry Call a set of positive integers 'conspiratorial' if no three of them are pairwise relatively prime. What is the largest number of elements in any conspiratorial subset of the integers 1 through 16? -/ theorem putnam_1974_a1 -(conspiratorial : Set ℤ → Prop) -(hconspiratorial : conspiratorial = fun S => ∀ a ∈ S, ∀ b ∈ S, ∀ c ∈ S, (a > 0 ∧ b > 0 ∧ c > 0) ∧ ((a ≠ b ∧ b ≠ c ∧ a ≠ c) → (Int.gcd a b > 1 ∨ Int.gcd b c > 1 ∨ Int.gcd a c > 1))) -: (∀ S : Set ℤ, S ⊆ Icc 1 16 → conspiratorial S → S.encard ≤ putnam_1974_a1_solution) ∧ (∃ S : Set ℤ, S ⊆ Icc 1 16 ∧ conspiratorial S ∧ S.encard = putnam_1974_a1_solution) := -sorry + (conspiratorial : Set ℤ → Prop) + (hconspiratorial : ∀ S, conspiratorial S ↔ ∀ a ∈ S, ∀ b ∈ S, ∀ c ∈ S, (a > 0 ∧ b > 0 ∧ c > 0) ∧ ((a ≠ b ∧ b ≠ c ∧ a ≠ c) → (Int.gcd a b > 1 ∨ Int.gcd b c > 1 ∨ Int.gcd a c > 1))) : + IsGreatest {k | ∃ S, S ⊆ Icc 1 16 ∧ conspiratorial S ∧ S.encard = k} putnam_1974_a1_solution := + sorry diff --git a/lean4/src/putnam_1976_b6.lean b/lean4/src/putnam_1976_b6.lean index 7564703..398d7bf 100644 --- a/lean4/src/putnam_1976_b6.lean +++ b/lean4/src/putnam_1976_b6.lean @@ -9,6 +9,6 @@ theorem putnam_1976_b6 (σ : ℕ → ℤ) (hσ : σ = fun N : ℕ => ∑ d in Nat.divisors N, (d : ℤ)) (quasiperfect : ℕ → Prop) -(quasiperfect_def : quasiperfect = fun N : ℕ => σ N = 2*N + 1) +(quasiperfect_def : ∀ N, quasiperfect N ↔ σ N = 2*N + 1) : ∀ N : ℕ, quasiperfect N → ∃ m : ℤ, Odd m ∧ m^2 = N := sorry diff --git a/lean4/src/putnam_1977_a2.lean b/lean4/src/putnam_1977_a2.lean index 8c01350..c5ea460 100644 --- a/lean4/src/putnam_1977_a2.lean +++ b/lean4/src/putnam_1977_a2.lean @@ -5,6 +5,7 @@ abbrev putnam_1977_a2_solution : ℝ → ℝ → ℝ → ℝ → Prop := sorry /-- Find all real solutions $(a, b, c, d)$ to the equations $a + b + c = d$, $\frac{1}{a} + \frac{1}{b} + \frac{1}{c} = \frac{1}{d}$. -/ -theorem putnam_1977_a2 -: (∀ a b c d : ℝ, a ≠ 0 → b ≠ 0 → c ≠ 0 → d ≠ 0 → ((a + b + c = d ∧ 1 / a + 1 / b + 1 / c = 1 / d) ↔ putnam_1977_a2_solution a b c d)) := -sorry +theorem putnam_1977_a2 : + ∀ a b c d : ℝ, putnam_1977_a2_solution a b c d ↔ + a ≠ 0 → b ≠ 0 → c ≠ 0 → d ≠ 0 → (a + b + c = d ∧ 1 / a + 1 / b + 1 / c = 1 / d) := + sorry diff --git a/lean4/src/putnam_1977_a4.lean b/lean4/src/putnam_1977_a4.lean index 349be51..3a7ad07 100644 --- a/lean4/src/putnam_1977_a4.lean +++ b/lean4/src/putnam_1977_a4.lean @@ -7,6 +7,7 @@ noncomputable abbrev putnam_1977_a4_solution : RatFunc ℝ := sorry /-- Find $\sum_{n=0}^{\infty} \frac{x^{2^n}}{1 - x^{2^{n+1}}}$ as a rational function of $x$ for $x \in (0, 1)$. -/ -theorem putnam_1977_a4 -: (∀ x ∈ Ioo 0 1, putnam_1977_a4_solution.eval (id ℝ) x = ∑' n : ℕ, x ^ 2 ^ n / (1 - x ^ 2 ^ (n + 1))) := -sorry +theorem putnam_1977_a4 : + ∀ x ∈ Ioo 0 1, + putnam_1977_a4_solution.eval (id ℝ) x = ∑' n : ℕ, x ^ 2 ^ n / (1 - x ^ 2 ^ (n + 1)) := + sorry diff --git a/lean4/src/putnam_1978_a3.lean b/lean4/src/putnam_1978_a3.lean index ec08e9f..bc68443 100644 --- a/lean4/src/putnam_1978_a3.lean +++ b/lean4/src/putnam_1978_a3.lean @@ -12,9 +12,9 @@ I_k = \int_0^{\infty} \frac{x^k}{p(x)} \, dx. For which $k$ is $I_k$ smallest? -/ theorem putnam_1978_a3 -(p : Polynomial ℝ) -(hp : p = 2 * (X ^ 6 + 1) + 4 * (X ^ 5 + X) + 3 * (X ^ 4 + X ^ 2) + 5 * X ^ 3) -(I : ℕ → ℝ) -(hI : I = fun k ↦ ∫ x in Ioi 0, x ^ k / p.eval x) -: (putnam_1978_a3_solution ∈ Ioo 0 5 ∧ ∀ k ∈ Ioo 0 5, I putnam_1978_a3_solution ≤ I k) := -sorry + (p : Polynomial ℝ) + (hp : p = 2 * (X ^ 6 + 1) + 4 * (X ^ 5 + X) + 3 * (X ^ 4 + X ^ 2) + 5 * X ^ 3) + (I : ℕ → ℝ) + (hI : I = fun k ↦ ∫ x in Ioi 0, x ^ k / p.eval x) : + IsLeast {y | ∃ k ∈ Ioo 0 5, I k = y} putnam_1978_a3_solution := + sorry diff --git a/lean4/src/putnam_1978_b4.lean b/lean4/src/putnam_1978_b4.lean index 08856fe..3bd0441 100644 --- a/lean4/src/putnam_1978_b4.lean +++ b/lean4/src/putnam_1978_b4.lean @@ -5,6 +5,9 @@ open Set Real Filter Topology Polynomial /-- Show that we can find integers $a, b, c, d$ such that $a^2 + b^2 + c^2 + d^2 = abc + abd + acd + bcd$, and the smallest of $a, b, c, d$ is arbitrarily large. -/ -theorem putnam_1978_b4 -: (∀ N : ℝ, ∃ a b c d : ℤ, a > N ∧ b > N ∧ c > N ∧ d > N ∧ a ^ 2 + b ^ 2 + c ^ 2 + d ^ 2 = a * b * c + a * b * d + a * c * d + b * c * d) := +theorem putnam_1978_b4 : + ∀ N : ℝ, + ∃ a b c d : ℤ, + a > N ∧ b > N ∧ c > N ∧ d > N ∧ + a ^ 2 + b ^ 2 + c ^ 2 + d ^ 2 = a * b * c + a * b * d + a * c * d + b * c * d := sorry diff --git a/lean4/src/putnam_1979_a1.lean b/lean4/src/putnam_1979_a1.lean index 81905d0..1a42b72 100644 --- a/lean4/src/putnam_1979_a1.lean +++ b/lean4/src/putnam_1979_a1.lean @@ -6,7 +6,7 @@ abbrev putnam_1979_a1_solution : Multiset ℕ := sorry For which positive integers $n$ and $a_1, a_2, \dots, a_n$ with $\sum_{i = 1}^{n} a_i = 1979$ does $\prod_{i = 1}^{n} a_i$ attain the greatest value? -/ theorem putnam_1979_a1 -(P : Multiset ℕ → Prop) -(hP : P = fun a => Multiset.card a > 0 ∧ (∀ i ∈ a, i > 0) ∧ a.sum = 1979) -: P putnam_1979_a1_solution ∧ ∀ a : Multiset ℕ, P a → putnam_1979_a1_solution.prod ≥ a.prod := -sorry + (P : Multiset ℕ → Prop) + (hP : ∀ a, P a ↔ Multiset.card a > 0 ∧ (∀ i ∈ a, i > 0) ∧ a.sum = 1979) : + P putnam_1979_a1_solution ∧ ∀ a : Multiset ℕ, P a → putnam_1979_a1_solution.prod ≥ a.prod := + sorry diff --git a/lean4/src/putnam_1979_a4.lean b/lean4/src/putnam_1979_a4.lean index f4d706d..81b9cb1 100644 --- a/lean4/src/putnam_1979_a4.lean +++ b/lean4/src/putnam_1979_a4.lean @@ -8,13 +8,15 @@ abbrev putnam_1979_a4_solution : Prop := sorry Let $A$ be a set of $2n$ points in the plane, $n$ colored red and $n$ colored blue, such that no three points in $A$ are collinear. Must there exist $n$ closed straight line segments, each connecting one red and one blue point in $A$, such that no two of the $n$ line segments intersect? -/ theorem putnam_1979_a4 -(A : Finset (Fin 2 → ℝ) × Finset (Fin 2 → ℝ) → Prop) -(hA : A = fun (R, B) => R.card = B.card ∧ R ∩ B = ∅ ∧ -∀ u : Finset (Fin 2 → ℝ), u ⊆ R ∪ B ∧ u.card = 3 → ¬Collinear ℝ (u : Set (Fin 2 → ℝ))) -(w : (Fin 2 → ℝ) × (Fin 2 → ℝ) → ℝ → (Fin 2 → ℝ)) -(hw : w = fun (P, Q) => fun x : ℝ => fun i : Fin 2 => x * P i + (1 - x) * Q i) -: (∀ R : Finset (Fin 2 → ℝ), ∀ B : Finset (Fin 2 → ℝ), A (R, B) → ∃ v : Finset ((Fin 2 → ℝ) × (Fin 2 → ℝ)), -(∀ L ∈ v, ∀ M ∈ v, L ≠ M → ∀ x ∈ Icc 0 1, ∀ y ∈ Icc 0 1, -Real.sqrt ((w (L.1, L.2) x 0 - w (M.1, M.2) y 0)^2 + (w (L.1, L.2) x 1 - w (M.1, M.2) y 1)^2) ≠ 0) ∧ -v.card = R.card ∧ ∀ L ∈ v, L.1 ∈ R ∧ L.2 ∈ B) ↔ putnam_1979_a4_solution := -sorry + (A : Finset (Fin 2 → ℝ) × Finset (Fin 2 → ℝ) → Prop) + (hA : A = fun (R, B) => R.card = B.card ∧ R ∩ B = ∅ ∧ + ∀ u : Finset (Fin 2 → ℝ), u ⊆ R ∪ B → u.card = 3 → ¬Collinear ℝ (u : Set (Fin 2 → ℝ))) + (w : (Fin 2 → ℝ) × (Fin 2 → ℝ) → ℝ → (Fin 2 → ℝ)) + (hw : w = fun (P, Q) => fun x : ℝ => fun i : Fin 2 => x * P i + (1 - x) * Q i) : + putnam_1979_a4_solution ↔ + (∀ R B, A (R, B) → + ∃ v : Finset ((Fin 2 → ℝ) × (Fin 2 → ℝ)), + (∀ L ∈ v, ∀ M ∈ v, L ≠ M → ∀ x ∈ Icc 0 1, ∀ y ∈ Icc 0 1, + Real.sqrt ((w (L.1, L.2) x 0 - w (M.1, M.2) y 0)^2 + (w (L.1, L.2) x 1 - w (M.1, M.2) y 1)^2) ≠ 0) ∧ + v.card = R.card ∧ ∀ L ∈ v, L.1 ∈ R ∧ L.2 ∈ B) := + sorry diff --git a/lean4/src/putnam_1979_a5.lean b/lean4/src/putnam_1979_a5.lean index 68b3ae6..c008dbf 100644 --- a/lean4/src/putnam_1979_a5.lean +++ b/lean4/src/putnam_1979_a5.lean @@ -9,6 +9,6 @@ theorem putnam_1979_a5 (S : ℝ → ℕ → ℤ) (hS : S = fun x : ℝ => fun n : ℕ => Int.floor (n*x)) (P : ℝ → Prop) -(hP : P = fun x : ℝ => x^3 - 10*x^2 + 29*x - 25 = 0) +(hP : ∀ x, P x ↔ x^3 - 10*x^2 + 29*x - 25 = 0) : ∃ α β : ℝ, α ≠ β ∧ P α ∧ P β ∧ ∀ n : ℕ, ∃ m : ℤ, m > n ∧ ∃ c d : ℕ, S α c = m ∧ S β d = m := sorry diff --git a/lean4/src/putnam_1980_a2.lean b/lean4/src/putnam_1980_a2.lean index db78578..6132bf5 100644 --- a/lean4/src/putnam_1980_a2.lean +++ b/lean4/src/putnam_1980_a2.lean @@ -6,9 +6,14 @@ abbrev putnam_1980_a2_solution : ℕ → ℕ → ℕ := sorry Let $r$ and $s$ be positive integers. Derive a formula for the number of ordered quadruples $(a,b,c,d)$ of positive integers such that $3^r \cdot 7^s=\text{lcm}[a,b,c]=\text{lcm}[a,b,d]=\text{lcm}[a,c,d]=\text{lcm}[b,c,d]$. The answer should be a function of $r$ and $s$. (Note that $\text{lcm}[x,y,z]$ denotes the least common multiple of $x,y,z$.) -/ theorem putnam_1980_a2 -(r s : ℕ) -(abcdlcm : ℕ → ℕ → ℕ → ℕ → Prop) -(rspos : r > 0 ∧ s > 0) -(habcdlcm : ∀ a b c d : ℕ, abcdlcm a b c d = (a > 0 ∧ b > 0 ∧ c > 0 ∧ d > 0 ∧ (3 ^ r * 7 ^ s = Nat.lcm (Nat.lcm a b) c) ∧ (3 ^ r * 7 ^ s = Nat.lcm (Nat.lcm a b) d) ∧ (3 ^ r * 7 ^ s = Nat.lcm (Nat.lcm a c) d) ∧ (3 ^ r * 7 ^ s = Nat.lcm (Nat.lcm b c) d))) -: {(a, b, c, d) : ℕ × ℕ × ℕ × ℕ | abcdlcm a b c d}.encard = putnam_1980_a2_solution r s := -sorry + (r s : ℕ) + (abcdlcm : ℕ → ℕ → ℕ → ℕ → Prop) + (rspos : r > 0 ∧ s > 0) + (habcdlcm : ∀ a b c d : ℕ, abcdlcm a b c d ↔ + (a > 0 ∧ b > 0 ∧ c > 0 ∧ d > 0 ∧ + (3 ^ r * 7 ^ s = Nat.lcm (Nat.lcm a b) c) ∧ + (3 ^ r * 7 ^ s = Nat.lcm (Nat.lcm a b) d) ∧ + (3 ^ r * 7 ^ s = Nat.lcm (Nat.lcm a c) d) ∧ + (3 ^ r * 7 ^ s = Nat.lcm (Nat.lcm b c) d))) : + {(a, b, c, d) : ℕ × ℕ × ℕ × ℕ | abcdlcm a b c d}.encard = putnam_1980_a2_solution r s := + sorry diff --git a/lean4/src/putnam_1980_a4.lean b/lean4/src/putnam_1980_a4.lean index 4bbb535..07cd883 100644 --- a/lean4/src/putnam_1980_a4.lean +++ b/lean4/src/putnam_1980_a4.lean @@ -7,7 +7,10 @@ import Mathlib \end{enumerate} -/ theorem putnam_1980_a4 -(abcvals : ℤ → ℤ → ℤ → Prop) -(habcvals : ∀ a b c : ℤ, abcvals a b c = (¬(a = 0 ∧ b = 0 ∧ c = 0) ∧ |a| < 1000000 ∧ |b| < 1000000 ∧ |c| < 1000000)) -: (∃ a b c : ℤ, abcvals a b c ∧ |a + b * Real.sqrt 2 + c * Real.sqrt 3| < 10 ^ (-(11 : ℝ))) ∧ (∀ a b c : ℤ, abcvals a b c → |a + b * Real.sqrt 2 + c * Real.sqrt 3| > 10 ^ (-(21 : ℝ))) := -sorry + (abcvals : ℤ → ℤ → ℤ → Prop) + (habcvals : ∀ a b c : ℤ, abcvals a b c ↔ (a = 0 ∧ b = 0 ∧ c = 0) ∧ |a| < 1000000 ∧ |b| < 1000000 ∧ |c| < 1000000) : + (∃ a b c : ℤ, + abcvals a b c ∧ + |a + b * Real.sqrt 2 + c * Real.sqrt 3| < 10 ^ (-(11 : ℝ))) ∧ + (∀ a b c : ℤ, abcvals a b c → |a + b * Real.sqrt 2 + c * Real.sqrt 3| > 10 ^ (-(21 : ℝ))) := + sorry diff --git a/lean4/src/putnam_1980_a5.lean b/lean4/src/putnam_1980_a5.lean index 1faf07c..041638e 100644 --- a/lean4/src/putnam_1980_a5.lean +++ b/lean4/src/putnam_1980_a5.lean @@ -4,9 +4,9 @@ import Mathlib Let $P(t)$ be a nonconstant polynomial with real coefficients. Prove that the system of simultaneous equations $0=\int_0^xP(t)\sin t\,dt=\int_0^xP(t)\cos t\,dt$ has only finitely many real solutions $x$. -/ theorem putnam_1980_a5 -(P : Polynomial ℝ) -(xeqs : ℝ → Prop) -(Pnonconst : P.degree > 0) -(hxeqs : ∀ x : ℝ, xeqs x = (0 = (∫ t in (0)..x, P.eval t * Real.sin t) ∧ 0 = (∫ t in (0)..x, P.eval t * Real.cos t))) -: {x : ℝ | xeqs x}.Finite := -sorry + (P : Polynomial ℝ) + (Pnonconst : P.degree > 0) : + Set.Finite {x : ℝ | + 0 = (∫ t in (0)..x, P.eval t * Real.sin t) ∧ + 0 = (∫ t in (0)..x, P.eval t * Real.cos t)} := + sorry diff --git a/lean4/src/putnam_1980_a6.lean b/lean4/src/putnam_1980_a6.lean index 65f2ac4..4fab548 100644 --- a/lean4/src/putnam_1980_a6.lean +++ b/lean4/src/putnam_1980_a6.lean @@ -7,9 +7,9 @@ noncomputable abbrev putnam_1980_a6_solution : ℝ := sorry Let $C$ be the class of all real valued continuously differentiable functions $f$ on the interval $0 \leq x \leq 1$ with $f(0)=0$ and $f(1)=1$. Determine the largest real number $u$ such that $u \leq \int_0^1|f'(x)-f(x)|\,dx$ for all $f$ in $C$. -/ theorem putnam_1980_a6 -(C : Set (ℝ → ℝ)) -(uleint : ℝ → Prop) -(hC : C = {f : ℝ → ℝ | ContDiffOn ℝ 1 f (Set.Icc 0 1) ∧ f 0 = 0 ∧ f 1 = 1}) -(huleint : ∀ u : ℝ, uleint u = ∀ f ∈ C, u ≤ (∫ x in Set.Ioo 0 1, |deriv f x - f x|)) -: uleint putnam_1980_a6_solution ∧ (∀ u : ℝ, uleint u → u ≤ putnam_1980_a6_solution) := -sorry + (C : Set (ℝ → ℝ)) + (hC : C = {f : ℝ → ℝ | ContDiffOn ℝ 1 f (Set.Icc 0 1) ∧ f 0 = 0 ∧ f 1 = 1}) : + IsGreatest + {(u : ℝ) | ∀ f ∈ C, u ≤ (∫ x in Set.Ioo 0 1, |deriv f x - f x|)} + putnam_1980_a6_solution := + sorry diff --git a/lean4/src/putnam_1980_b5.lean b/lean4/src/putnam_1980_b5.lean index c9d7ba0..8085dee 100644 --- a/lean4/src/putnam_1980_b5.lean +++ b/lean4/src/putnam_1980_b5.lean @@ -10,13 +10,14 @@ Let $S_t$ denote the set of all nonnegative increasing convex continuous functio For which real numbers $t \ge 0$ is $S_t$ closed under multiplication? -/ theorem putnam_1980_b5 -(T : Set ℝ) -(hT : T = Icc 0 1) -(P : ℝ → (ℝ → ℝ) → Prop) -(Convex : (ℝ → ℝ) → Prop) -(S : ℝ → Set (ℝ → ℝ)) -(hP : P = fun t : ℝ => fun f : ℝ → ℝ => f 1 - 2*f (2/3) + f (1/3) ≥ t*(f (2/3) - 2*f (1/3) + f 0)) -(hConvex : Convex = fun f : ℝ → ℝ => ∀ u ∈ T, ∀ v ∈ T, ∀ s ∈ T, f (s*u + (1 - s)*v) ≤ s*(f u) + (1 - s)*(f v)) -(hS : S = fun t : ℝ => {f : ℝ → ℝ | (∀ x ∈ T, f x ≥ 0) ∧ StrictMonoOn f T ∧ Convex f ∧ ContinuousOn f T ∧ P t f}) -: ∀ t : ℝ, t ≥ 0 → ((∀ f ∈ S t, ∀ g ∈ S t, f * g ∈ S t) ↔ putnam_1980_b5_solution t) := -sorry + (T : Set ℝ) + (hT : T = Icc 0 1) + (P : ℝ → (ℝ → ℝ) → Prop) + (IsConvex : (ℝ → ℝ) → Prop) + (S : ℝ → Set (ℝ → ℝ)) + (P_def : ∀ t f, P t f ↔ f 1 - 2*f (2/3) + f (1/3) ≥ t*(f (2/3) - 2*f (1/3) + f 0)) + (IsConvex_def : ∀ f, IsConvex f ↔ ∀ u ∈ T, ∀ v ∈ T, ∀ s ∈ T, f (s*u + (1 - s)*v) ≤ s*(f u) + (1 - s)*(f v)) + (hS : S = fun t : ℝ => {f : ℝ → ℝ | (∀ x ∈ T, f x ≥ 0) ∧ StrictMonoOn f T ∧ IsConvex f ∧ ContinuousOn f T ∧ P t f}) + (t : ℝ) (ht : t ≥ 0) : + putnam_1980_b5_solution t ↔ (∀ f ∈ S t, ∀ g ∈ S t, f * g ∈ S t) := + sorry diff --git a/lean4/src/putnam_1981_a1.lean b/lean4/src/putnam_1981_a1.lean index 911f927..f8cdc35 100644 --- a/lean4/src/putnam_1981_a1.lean +++ b/lean4/src/putnam_1981_a1.lean @@ -8,9 +8,9 @@ noncomputable abbrev putnam_1981_a1_solution : ℝ := sorry Let $E(n)$ be the greatest integer $k$ such that $5^k$ divides $1^1 2^2 3^3 \cdots n^n$. Find $\lim_{n \rightarrow \infty} \frac{E(n)}{n^2}$. -/ theorem putnam_1981_a1 -(P : ℕ → ℕ → Prop) -(hP : P = fun n k : ℕ => 5^k ∣ ∏ m in Finset.Icc 1 n, (m^m : ℤ)) -(E : ℕ → ℕ) -(hE : ∀ n ∈ Ici 1, P n (E n) ∧ ∀ k : ℕ, P n k → k ≤ E n) -: Tendsto (fun n : ℕ => ((E n) : ℝ)/n^2) atTop (𝓝 putnam_1981_a1_solution) := -sorry + (P : ℕ → ℕ → Prop) + (hP : ∀ n k, P n k ↔ 5^k ∣ ∏ m in Finset.Icc 1 n, (m^m : ℤ)) + (E : ℕ → ℕ) + (hE : ∀ n ∈ Ici 1, P n (E n) ∧ ∀ k : ℕ, P n k → k ≤ E n) : + Tendsto (fun n : ℕ => ((E n) : ℝ)/n^2) atTop (𝓝 putnam_1981_a1_solution) := + sorry diff --git a/lean4/src/putnam_1981_b2.lean b/lean4/src/putnam_1981_b2.lean index c74c123..9b55355 100644 --- a/lean4/src/putnam_1981_b2.lean +++ b/lean4/src/putnam_1981_b2.lean @@ -8,10 +8,9 @@ noncomputable abbrev putnam_1981_b2_solution : ℝ := sorry Determine the minimum value attained by $$(r - 1)^2 + (\frac{s}{r} - 1)^2 + (\frac{t}{s} - 1)^2 + (\frac{4}{t} - 1)^2$$ across all choices of real $r$, $s$, and $t$ that satisfy $1 \le r \le s \le t \le 4$. -/ theorem putnam_1981_b2 -(P : ℝ × ℝ × ℝ → Prop) -(hP : P = fun (r, s, t) => 1 ≤ r ∧ r ≤ s ∧ s ≤ t ∧ t ≤ 4) -(f : ℝ × ℝ × ℝ → ℝ) -(hf : f = fun (r, s, t) => (r - 1)^2 + (s/r - 1)^2 + (t/s - 1)^2 + (4/t - 1)^2) -: (∃ r : ℝ, ∃ s : ℝ, ∃ t : ℝ, P (r, s, t) ∧ f (r, s, t) = putnam_1981_b2_solution) ∧ -∀ r : ℝ, ∀ s : ℝ, ∀ t : ℝ, P (r, s, t) → f (r, s, t) ≥ putnam_1981_b2_solution := -sorry + (P : ℝ × ℝ × ℝ → Prop) + (hP : P = fun (r, s, t) => 1 ≤ r ∧ r ≤ s ∧ s ≤ t ∧ t ≤ 4) + (f : ℝ × ℝ × ℝ → ℝ) + (hf : f = fun (r, s, t) => (r - 1)^2 + (s/r - 1)^2 + (t/s - 1)^2 + (4/t - 1)^2) : + IsLeast {y | ∃ r s t, P (r, s, t) ∧ f (r, s, t) = y} putnam_1981_b2_solution := + sorry diff --git a/lean4/src/putnam_1981_b3.lean b/lean4/src/putnam_1981_b3.lean index e6aedd7..85bfc0f 100644 --- a/lean4/src/putnam_1981_b3.lean +++ b/lean4/src/putnam_1981_b3.lean @@ -6,7 +6,9 @@ open Topology Filter Set Polynomial Function Prove that, for infinitely many positive integers $n$, all primes $p$ that divide $n^2 + 3$ also divide $k^2 + 3$ for some integer $k$ such that $k^2 < n$. -/ theorem putnam_1981_b3 -(P : ℕ → Prop) -(hP : P = fun n : ℕ => ∀ p : ℕ, (Nat.Prime p ∧ p ∣ n^2 + 3) → ∃ k : ℕ, (p : ℤ) ∣ (k : ℤ)^2 + 3 ∧ k^2 < n) -: ∀ n : ℕ, ∃ m : ℕ, (m : ℤ) > n ∧ P m := -sorry + (P : ℕ → Prop) + (hP : ∀ n, P n ↔ + ∀ p : ℕ, (Nat.Prime p ∧ p ∣ n^2 + 3) → + ∃ k : ℕ, (p : ℤ) ∣ (k : ℤ)^2 + 3 ∧ k^2 < n) : + ∀ n : ℕ, ∃ m : ℕ, (m : ℤ) > n ∧ P m := + sorry diff --git a/lean4/src/putnam_1981_b4.lean b/lean4/src/putnam_1981_b4.lean index 5cb5f78..021b37e 100644 --- a/lean4/src/putnam_1981_b4.lean +++ b/lean4/src/putnam_1981_b4.lean @@ -8,9 +8,13 @@ abbrev putnam_1981_b4_solution : Prop := sorry Let $V$ be a set of $5$ by $7$ matrices, with real entries and with the property that $rA+sB \in V$ whenever $A,B \in V$ and $r$ and $s$ are scalars (i.e., real numbers). \emph{Prove or disprove} the following assertion: If $V$ contains matrices of ranks $0$, $1$, $2$, $4$, and $5$, then it also contains a matrix of rank $3$. [The rank of a nonzero matrix $M$ is the largest $k$ such that the entries of some $k$ rows and some $k$ columns form a $k$ by $k$ matrix with a nonzero determinant.] -/ theorem putnam_1981_b4 -(VAB : Set (Matrix (Fin 5) (Fin 7) ℝ) → Prop) -(Vrank : Set (Matrix (Fin 5) (Fin 7) ℝ) → ℕ → Prop) -(hVAB : ∀ V : Set (Matrix (Fin 5) (Fin 7) ℝ), VAB V = (∀ A ∈ V, ∀ B ∈ V, ∀ r s : ℝ, r • A + s • B ∈ V)) -(hVrank : ∀ (V : Set (Matrix (Fin 5) (Fin 7) ℝ)) (k : ℕ), Vrank V k = ∃ A ∈ V, A.rank = k) -: (∀ V : Set (Matrix (Fin 5) (Fin 7) ℝ), (VAB V ∧ Vrank V 0 ∧ Vrank V 1 ∧ Vrank V 2 ∧ Vrank V 4 ∧ Vrank V 5) → Vrank V 3) ↔ putnam_1981_b4_solution := -sorry + (VAB : Set (Matrix (Fin 5) (Fin 7) ℝ) → Prop) + (Vrank : Set (Matrix (Fin 5) (Fin 7) ℝ) → ℕ → Prop) + (hVAB : ∀ V, VAB V = (∀ A ∈ V, ∀ B ∈ V, ∀ r s : ℝ, r • A + s • B ∈ V)) + (hVrank : ∀ V k, Vrank V k = ∃ A ∈ V, A.rank = k) : + putnam_1981_b4_solution ↔ + (∀ V, + VAB V → Vrank V 0 → Vrank V 1 → + Vrank V 2 → Vrank V 4 → Vrank V 5 → + Vrank V 3) := + sorry diff --git a/lean4/src/putnam_1982_a3.lean b/lean4/src/putnam_1982_a3.lean index c0627f9..bab8494 100644 --- a/lean4/src/putnam_1982_a3.lean +++ b/lean4/src/putnam_1982_a3.lean @@ -7,6 +7,6 @@ noncomputable abbrev putnam_1982_a3_solution : ℝ := sorry /-- Evaluate $\int_0^{\infty} \frac{\tan^{-1}(\pi x) - \tan^{-1} x}{x} \, dx$. -/ -theorem putnam_1982_a3 -: (Tendsto (fun t ↦ ∫ x in (0)..t, (arctan (Real.pi * x) - arctan x) / x) atTop (𝓝 putnam_1982_a3_solution)) := +theorem putnam_1982_a3 : + Tendsto (fun t ↦ ∫ x in (0)..t, (arctan (Real.pi * x) - arctan x) / x) atTop (𝓝 putnam_1982_a3_solution) := sorry diff --git a/lean4/src/putnam_1982_a4.lean b/lean4/src/putnam_1982_a4.lean index 613fef7..3da8463 100644 --- a/lean4/src/putnam_1982_a4.lean +++ b/lean4/src/putnam_1982_a4.lean @@ -6,10 +6,14 @@ open Function Filter Topology Assume that the system of simultaneous differentiable equations \[y' = -z^3, z' = y^3\] with the initial conditions $y(0) = 1, z(0) = 0$ has a unique solution $y = f(x), z = g(x)$ defined for all real $x$. Prove that there exists a positive constant $L$ such that for all real $x$, \[f(x) + L = f(x), g(x + L) = g(x).\] -/ theorem putnam_1982_a4 -(hdiffeq : (ℝ → ℝ) → (ℝ → ℝ) → Prop) -(hdiffeq_def : hdiffeq = fun y z => y 0 = 1 ∧ z 0 = 0 ∧ ContDiff ℝ 1 y ∧ ContDiff ℝ 1 z ∧ (∀ x : ℝ, deriv y x = -1 * (z x)^3 ∧ deriv z x = (y x)^3)) -(f g : ℝ → ℝ) -(hfgsat : hdiffeq f g) -(hfgonly : ¬(∃ f' g' : ℝ → ℝ, (f ≠ f' ∨ g ≠ g') ∧ hdiffeq f' g')) -: ∃ L : ℝ, L > 0 ∧ Function.Periodic f L ∧ Function.Periodic g L := -sorry + (hdiffeq : (ℝ → ℝ) → (ℝ → ℝ) → Prop) + (hdiffeq_def : ∀ y z, + hdiffeq y z ↔ + y 0 = 1 ∧ z 0 = 0 ∧ + ContDiff ℝ 1 y ∧ ContDiff ℝ 1 z ∧ + (∀ x : ℝ, deriv y x = -1 * (z x)^3 ∧ deriv z x = (y x)^3)) + (f g : ℝ → ℝ) + (hfgsat : hdiffeq f g) + (hfgonly : ¬(∃ f' g' : ℝ → ℝ, (f ≠ f' ∨ g ≠ g') ∧ hdiffeq f' g')) : + ∃ L : ℝ, L > 0 ∧ Function.Periodic f L ∧ Function.Periodic g L := + sorry diff --git a/lean4/src/putnam_1982_a6.lean b/lean4/src/putnam_1982_a6.lean index d76da02..812a11b 100644 --- a/lean4/src/putnam_1982_a6.lean +++ b/lean4/src/putnam_1982_a6.lean @@ -16,17 +16,13 @@ $\lim_{n \rightarrow \infty}\sum_{k = 1}^{n} x_k = 1$. \end{enumerate} Prove or disprove: these conditions imply that $$\lim_{n \rightarrow \infty} \sum_{k = 1}^{n} x_{b(k)} = 1.$$ -/ -theorem putnam_1982_a6 -(S : Set ℕ) -(S_def : S = Ici 1) -(hb : (ℕ → ℕ) → Prop) -(hb_def : hb = fun b : ℕ → ℕ => BijOn b S S) -(hx : (ℕ → ℝ) → Prop) -(hx_def : hx = fun x : ℕ → ℝ => StrictAntiOn (fun n : ℕ => |x n|) S) -(limb : (ℕ → ℕ) × (ℕ → ℝ) → Prop) -(hlimb : limb = fun (b, x) => Tendsto (fun n : ℕ => |b n - (n : ℤ)| * |x n|) atTop (𝓝 0)) -(limx : (ℕ → ℝ) → Prop) -(hlimx : limx = fun x : ℕ → ℝ => Tendsto (fun n : ℕ => ∑ k in Finset.Icc 1 n, x k) atTop (𝓝 1)) -: (∀ b : ℕ → ℕ, ∀ x : ℕ → ℝ, hb b ∧ hx x ∧ limb (b, x) ∧ limx x → -Tendsto (fun n : ℕ => ∑ k in Finset.Icc 1 n, x (b k)) atTop (𝓝 1)) ↔ putnam_1982_a6_solution := +theorem putnam_1982_a6 : + (∀ b : ℕ → ℕ, + ∀ x : ℕ → ℝ, + BijOn b (Ici 1) (Ici 1) → + StrictAntiOn (fun n : ℕ => |x n|) (Ici 1) → + Tendsto (fun n : ℕ => |b n - (n : ℤ)| * |x n|) atTop (𝓝 0) → + Tendsto (fun n : ℕ => ∑ k in Finset.Icc 1 n, x k) atTop (𝓝 1) → + Tendsto (fun n : ℕ => ∑ k in Finset.Icc 1 n, x (b k)) atTop (𝓝 1)) + ↔ putnam_1982_a6_solution := sorry diff --git a/lean4/src/putnam_1983_a5.lean b/lean4/src/putnam_1983_a5.lean index ab17938..6d6090e 100644 --- a/lean4/src/putnam_1983_a5.lean +++ b/lean4/src/putnam_1983_a5.lean @@ -7,6 +7,6 @@ abbrev putnam_1983_a5_solution : Prop := sorry /-- Prove or disprove that there exists a positive real number $\alpha$ such that $[\alpha_n] - n$ is even for all integers $n > 0$. (Here $[x]$ denotes the greatest integer less than or equal to $x$.) -/ -theorem putnam_1983_a5 -: ((∃ α : ℝ, α > 0 ∧ ∀ n : ℕ, n > 0 → Even (⌊α ^ n⌋ - n)) ↔ putnam_1983_a5_solution) := -sorry +theorem putnam_1983_a5 : + (∃ α : ℝ, α > 0 ∧ ∀ n : ℕ, n > 0 → Even (⌊α ^ n⌋ - n)) ↔ putnam_1983_a5_solution := + sorry diff --git a/lean4/src/putnam_1983_b2.lean b/lean4/src/putnam_1983_b2.lean index e018606..05148a2 100644 --- a/lean4/src/putnam_1983_b2.lean +++ b/lean4/src/putnam_1983_b2.lean @@ -8,7 +8,12 @@ abbrev putnam_1983_b2_solution : Prop := sorry Let $f(n)$ be the number of ways of representing $n$ as a sum of powers of $2$ with no power being used more than $3$ times. For example, $f(7) = 4$ (the representations are $4 + 2 + 1$, $4 + 1 + 1 + 1$, $2 + 2 + 2 + 1$, $2 + 2 + 1 + 1 + 1$). Can we find a real polynomial $p(x)$ such that $f(n) = [p(n)]$, where $[u]$ denotes the greatest integer less than or equal to $u$? -/ theorem putnam_1983_b2 -(f : ℕ+ → ℕ) -(hf : f = fun (n : ℕ+) ↦ {M : Multiset ℕ | (∀ m ∈ M, ∃ k : ℕ, m = (2 ^ k : ℤ)) ∧ (∀ m ∈ M, M.count m ≤ 3) ∧ (M.sum : ℤ) = n}.ncard) -: ((∃ p : Polynomial ℝ, ∀ n : ℕ+, ⌊p.eval (n : ℝ)⌋ = f n) ↔ putnam_1983_b2_solution) := -sorry + (f : ℕ+ → ℕ) + (hf : f = fun (n : ℕ+) ↦ + Set.ncard {M : Multiset ℕ | + (∀ m ∈ M, ∃ k : ℕ, m = (2 ^ k : ℤ)) ∧ + (∀ m ∈ M, M.count m ≤ 3) ∧ + (M.sum : ℤ) = n}) : + putnam_1983_b2_solution ↔ + (∃ p : Polynomial ℝ, ∀ n : ℕ+, ⌊p.eval (n : ℝ)⌋ = f n) := + sorry diff --git a/lean4/src/putnam_1983_b4.lean b/lean4/src/putnam_1983_b4.lean index f05e9ad..34569ef 100644 --- a/lean4/src/putnam_1983_b4.lean +++ b/lean4/src/putnam_1983_b4.lean @@ -6,10 +6,10 @@ open Nat Real Let $f(n) = n + [\sqrt n]$, where $[x]$ denotes the greatest integer less than or equal to $x$. Define the sequence $a_i$ by $a_0 = m$, $a_{n+1} = f(a_n)$. Prove that it contains at least one square. -/ theorem putnam_1983_b4 -(f : ℕ → ℤ) -(a : ℕ → ℕ) -(hf : f = fun (n : ℕ) ↦ n + Int.floor (Real.sqrt n)) -(ha0 : a 0 > 0) -(han : ∀ n : ℕ, a (n + 1) = f (a n)) -: (∃ i : ℕ, ∃ s : ℤ, a i = s ^ 2) := -sorry + (f : ℕ → ℤ) + (a : ℕ → ℕ) + (hf : f = fun (n : ℕ) ↦ n + Int.floor (√n)) + (ha0 : a 0 > 0) + (han : ∀ n : ℕ, a (n + 1) = f (a n)) : + (∃ i : ℕ, ∃ s : ℤ, a i = s ^ 2) := + sorry diff --git a/lean4/src/putnam_1984_a5.lean b/lean4/src/putnam_1984_a5.lean index 3d35816..adf953e 100644 --- a/lean4/src/putnam_1984_a5.lean +++ b/lean4/src/putnam_1984_a5.lean @@ -8,9 +8,11 @@ abbrev putnam_1984_a5_solution : ℕ × ℕ × ℕ × ℕ × ℕ := sorry Let $R$ be the region consisting of all triples $(x,y,z)$ of nonnegative real numbers satisfying $x+y+z \leq 1$. Let $w=1-x-y-z$. Express the value of the triple integral $\iiint_R x^1y^9z^8w^4\,dx\,dy\,dz$ in the form $a!b!c!d!/n!$, where $a$, $b$, $c$, $d$, and $n$ are positive integers. -/ theorem putnam_1984_a5 -(R : Set (Fin 3 → ℝ)) -(w : (Fin 3 → ℝ) → ℝ) -(hR : R = {p : Fin 3 → ℝ | (∀ i : Fin 3, p i ≥ 0) ∧ p 0 + p 1 + p 2 ≤ 1}) -(hw : ∀ p : Fin 3 → ℝ, w p = 1 - p 0 - p 1 - p 2) -: let (a, b, c, d, n) := putnam_1984_a5_solution; a > 0 ∧ b > 0 ∧ c > 0 ∧ d > 0 ∧ n > 0 ∧ (∫ p in R, (p 0) ^ 1 * (p 1) ^ 9 * (p 2) ^ 8 * (w p) ^ 4 = ((a)! * (b)! * (c)! * (d)! : ℝ) / (n)!) := -sorry + (R : Set (Fin 3 → ℝ)) + (w : (Fin 3 → ℝ) → ℝ) + (hR : R = {p | (∀ i : Fin 3, p i ≥ 0) ∧ p 0 + p 1 + p 2 ≤ 1}) + (hw : ∀ p, w p = 1 - p 0 - p 1 - p 2) : + let (a, b, c, d, n) := putnam_1984_a5_solution; + a > 0 ∧ b > 0 ∧ c > 0 ∧ d > 0 ∧ n > 0 ∧ + (∫ p in R, (p 0) ^ 1 * (p 1) ^ 9 * (p 2) ^ 8 * (w p) ^ 4 = ((a)! * (b)! * (c)! * (d)! : ℝ) / (n)!) := + sorry diff --git a/lean4/src/putnam_1984_b2.lean b/lean4/src/putnam_1984_b2.lean index 2652970..1ceac1e 100644 --- a/lean4/src/putnam_1984_b2.lean +++ b/lean4/src/putnam_1984_b2.lean @@ -8,7 +8,7 @@ abbrev putnam_1984_b2_solution : ℝ := sorry Find the minimum value of $(u-v)^2+(\sqrt{2-u^2}-\frac{9}{v})^2$ for $00$. -/ theorem putnam_1984_b2 -(f : ℝ → ℝ → ℝ) -(hf : ∀ u v : ℝ, f u v = (u - v) ^ 2 + (Real.sqrt (2 - u ^ 2) - 9 / v) ^ 2) -: (∃ u : Set.Ioo 0 (Real.sqrt 2), ∃ v > 0, f u v = putnam_1984_b2_solution) ∧ (∀ u : Set.Ioo 0 (Real.sqrt 2), ∀ v > 0, f u v ≥ putnam_1984_b2_solution) := -sorry + (f : ℝ → ℝ → ℝ) + (hf : ∀ u v : ℝ, f u v = (u - v) ^ 2 + (Real.sqrt (2 - u ^ 2) - 9 / v) ^ 2) : + IsLeast {y | ∃ᵉ (u : Set.Ioo 0 √2) (v > 0), f u v = y} putnam_1984_b2_solution := + sorry diff --git a/lean4/src/putnam_1984_b5.lean b/lean4/src/putnam_1984_b5.lean index 244c189..0eed150 100644 --- a/lean4/src/putnam_1984_b5.lean +++ b/lean4/src/putnam_1984_b5.lean @@ -8,11 +8,11 @@ noncomputable abbrev putnam_1984_b5_solution : ℤ × Polynomial ℝ × Polynomi For each nonnegative integer $k$, let $d(k)$ denote the number of $1$'s in the binary expansion of $k$ (for example, $d(0)=0$ and $d(5)=2$). Let $m$ be a positive integer. Express $\sum_{k=0}^{2^m-1} (-1)^{d(k)}k^m$ in the form $(-1)^ma^{f(m)}(g(m))!$, where $a$ is an integer and $f$ and $g$ are polynomials. -/ theorem putnam_1984_b5 -(sumbits : List ℕ → ℕ) -(d : ℕ → ℕ) -(m : ℕ) -(hsumbits : ∀ bits : List ℕ, sumbits bits = ∑ i : Fin bits.length, bits[i]) -(hd : ∀ k : ℕ, d k = sumbits (Nat.digits 2 k)) -(mpos : m > 0) -: let (a, f, g) := putnam_1984_b5_solution; ∑ k : Set.Icc 0 (2 ^ m - 1), (-(1 : ℤ)) ^ (d k) * (k : ℕ) ^ m = (-1) ^ m * (a : ℝ) ^ (f.eval (m : ℝ)) * (g.eval m)! := -sorry + (m : ℕ) (mpos : m > 0) + (d : ℕ → ℕ) + (sumbits : List ℕ → ℕ) + (hsumbits : ∀ bits : List ℕ, sumbits bits = ∑ i : Fin bits.length, bits[i]) + (hd : ∀ k : ℕ, d k = sumbits (Nat.digits 2 k)) : + let (a, f, g) := putnam_1984_b5_solution; + ∑ k : Set.Icc 0 (2 ^ m - 1), (-(1 : ℤ)) ^ (d k) * (k : ℕ) ^ m = (-1) ^ m * (a : ℝ) ^ (f.eval (m : ℝ)) * (g.eval m)! := + sorry diff --git a/lean4/src/putnam_1985_a1.lean b/lean4/src/putnam_1985_a1.lean index 5bb6342..9210567 100644 --- a/lean4/src/putnam_1985_a1.lean +++ b/lean4/src/putnam_1985_a1.lean @@ -12,6 +12,7 @@ Determine, with proof, the number of ordered triples $(A_1, A_2, A_3)$ of sets w \end{enumerate} Express your answer in the form $2^a 3^b 5^c 7^d$, where $a,b,c,d$ are nonnegative integers. -/ -theorem putnam_1985_a1 -: (let (a, b, c, d) := putnam_1985_a1_solution; {(A1, A2, A3) : Set ℤ × Set ℤ × Set ℤ | A1 ∪ A2 ∪ A3 = Icc 1 10 ∧ A1 ∩ A2 ∩ A3 = ∅}.ncard = 2 ^ a * 3 ^ b * 5 ^ c * 7 ^ d) := +theorem putnam_1985_a1 : + let (a, b, c, d) := putnam_1985_a1_solution; + {(A1, A2, A3) : Set ℤ × Set ℤ × Set ℤ | A1 ∪ A2 ∪ A3 = Icc 1 10 ∧ A1 ∩ A2 ∩ A3 = ∅}.ncard = 2 ^ a * 3 ^ b * 5 ^ c * 7 ^ d := sorry diff --git a/lean4/src/putnam_1985_a4.lean b/lean4/src/putnam_1985_a4.lean index 3aeccd4..2e2cef0 100644 --- a/lean4/src/putnam_1985_a4.lean +++ b/lean4/src/putnam_1985_a4.lean @@ -8,8 +8,8 @@ abbrev putnam_1985_a4_solution : Set (Fin 100) := sorry Define a sequence $\{a_i\}$ by $a_1=3$ and $a_{i+1}=3^{a_i}$ for $i \geq 1$. Which integers between $00$ and $99$ inclusive occur as the last two digits in the decimal expansion of infinitely many $a_i$? -/ theorem putnam_1985_a4 -(a : ℕ → ℕ) -(ha1 : a 1 = 3) -(ha : ∀ i ≥ 1, a (i + 1) = 3 ^ a i) -: ({k : Fin 100 | ∀ N : ℕ, ∃ i ≥ N, a i % 100 = k} = putnam_1985_a4_solution) := -sorry + (a : ℕ → ℕ) + (ha1 : a 1 = 3) + (ha : ∀ i ≥ 1, a (i + 1) = 3 ^ a i) : + {k : Fin 100 | ∀ N : ℕ, ∃ i ≥ N, a i % 100 = k} = putnam_1985_a4_solution := + sorry diff --git a/lean4/src/putnam_1985_a5.lean b/lean4/src/putnam_1985_a5.lean index 0375f1e..cd67264 100644 --- a/lean4/src/putnam_1985_a5.lean +++ b/lean4/src/putnam_1985_a5.lean @@ -8,7 +8,7 @@ abbrev putnam_1985_a5_solution : Set ℕ := sorry Let $I_m = \int_0^{2\pi} \cos(x)\cos(2x)\cdots \cos(mx)\,dx$. For which integers $m$, $1 \leq m \leq 10$ is $I_m \neq 0$? -/ theorem putnam_1985_a5 -(I : ℕ → ℝ) -(hI : I = fun (m : ℕ) ↦ ∫ x in (0)..(2 * Real.pi), ∏ k in Finset.Icc 1 m, cos (k * x)) -: ({m ∈ Finset.Icc 1 10 | I m ≠ 0} = putnam_1985_a5_solution) := -sorry + (I : ℕ → ℝ) + (hI : I = fun (m : ℕ) ↦ ∫ x in (0)..(2 * Real.pi), ∏ k in Finset.Icc 1 m, cos (k * x)) : + {m ∈ Finset.Icc 1 10 | I m ≠ 0} = putnam_1985_a5_solution := + sorry diff --git a/lean4/src/putnam_1985_a6.lean b/lean4/src/putnam_1985_a6.lean index c543fb2..2a1326e 100644 --- a/lean4/src/putnam_1985_a6.lean +++ b/lean4/src/putnam_1985_a6.lean @@ -17,9 +17,10 @@ Let $F(x) = 3x^2+7x+2$. Find, with proof, a polynomial $g(x)$ with real coeffici for every integer $n \geq 1$. -/ theorem putnam_1985_a6 -(Γ : Polynomial ℝ → ℝ) -(hΓ : Γ = fun p ↦ ∑ k in Finset.range (p.natDegree + 1), coeff p k ^ 2) -(f : Polynomial ℝ) -(hf : f = 3 * X ^ 2 + 7 * X + 2) -: (let g := putnam_1985_a6_solution; g.eval 0 = 1 ∧ ∀ n : ℕ, n ≥ 1 → Γ (f ^ n) = Γ (g ^ n)) := + (Γ : Polynomial ℝ → ℝ) + (f : Polynomial ℝ) + (hΓ : Γ = fun p ↦ ∑ k in Finset.range (p.natDegree + 1), coeff p k ^ 2) + (hf : f = 3 * X ^ 2 + 7 * X + 2) : + let g := putnam_1985_a6_solution; + g.eval 0 = 1 ∧ ∀ n : ℕ, n ≥ 1 → Γ (f ^ n) = Γ (g ^ n) := sorry diff --git a/lean4/src/putnam_1986_a1.lean b/lean4/src/putnam_1986_a1.lean index 30fd679..10a536f 100644 --- a/lean4/src/putnam_1986_a1.lean +++ b/lean4/src/putnam_1986_a1.lean @@ -6,9 +6,10 @@ abbrev putnam_1986_a1_solution : ℝ := sorry Find, with explanation, the maximum value of $f(x)=x^3-3x$ on the set of all real numbers $x$ satisfying $x^4+36 \leq 13x^2$. -/ theorem putnam_1986_a1 -(S : Set ℝ) -(hS : S = {x : ℝ | x ^ 4 + 36 ≤ 13 * x ^ 2}) -(f : ℝ → ℝ) -(hf : f = fun x ↦ x ^ 3 - 3 * x) -: (∀ x ∈ S, f x ≤ putnam_1986_a1_solution) ∧ (∃ x ∈ S, f x = putnam_1986_a1_solution) := -sorry + (S : Set ℝ) (f : ℝ → ℝ) + (hS : S = {x : ℝ | x ^ 4 + 36 ≤ 13 * x ^ 2}) + (hf : f = fun x ↦ x ^ 3 - 3 * x) : + IsGreatest + {f x | x ∈ S} + putnam_1986_a1_solution := + sorry diff --git a/lean4/src/putnam_1986_a4.lean b/lean4/src/putnam_1986_a4.lean index d686959..35433aa 100644 --- a/lean4/src/putnam_1986_a4.lean +++ b/lean4/src/putnam_1986_a4.lean @@ -24,8 +24,11 @@ f(n) = a_1 b_1^n + a_2 b_2^n + a_3 b_3^n + a_4, where the $a_i$'s and $b_i$'s are rational numbers. -/ theorem putnam_1986_a4 -(f : ℕ → ℕ) -(hf : f = fun n ↦ {A : Matrix (Fin n) (Fin n) ℤ | (∀ i j : Fin n, A i j ∈ ({-1, 0, 1} : Set ℤ)) ∧ ∃ S : ℤ, ∀ ϕ : Perm (Fin n), ∑ i : Fin n, A i (ϕ i) = S}.ncard) -: let (a1, b1, a2, b2, a3, b3, a4) := putnam_1986_a4_solution; -(∀ n > 0, f n = a1 * b1 ^ n + a2 * b2 ^ n + a3 * b3 ^ n + a4) := + (f : ℕ → ℕ) + (hf : f = fun n ↦ + Set.ncard {A : Matrix (Fin n) (Fin n) ℤ | + (∀ i j : Fin n, A i j ∈ ({-1, 0, 1} : Set ℤ)) ∧ + ∃ S : ℤ, ∀ ϕ : Perm (Fin n), ∑ i : Fin n, A i (ϕ i) = S}) : + let (a1, b1, a2, b2, a3, b3, a4) := putnam_1986_a4_solution; + (∀ n > 0, f n = a1 * b1 ^ n + a2 * b2 ^ n + a3 * b3 ^ n + a4) := sorry diff --git a/lean4/src/putnam_1986_b3.lean b/lean4/src/putnam_1986_b3.lean index d7cc676..7b6b019 100644 --- a/lean4/src/putnam_1986_b3.lean +++ b/lean4/src/putnam_1986_b3.lean @@ -6,13 +6,13 @@ open Real Equiv Polynomial Let $\Gamma$ consist of all polynomials in $x$ with integer coefficients. For $f$ and $g$ in $\Gamma$ and $m$ a positive integer, let $f \equiv g \pmod{m}$ mean that every coefficient of $f-g$ is an integral multiple of $m$. Let $n$ and $p$ be positive integers with $p$ prime. Given that $f,g,h,r$ and $s$ are in $\Gamma$ with $rf+sg\equiv 1 \pmod{p}$ and $fg \equiv h \pmod{p}$, prove that there exist $F$ and $G$ in $\Gamma$ with $F \equiv f \pmod{p}$, $G \equiv g \pmod{p}$, and $FG \equiv h \pmod{p^n}$. -/ theorem putnam_1986_b3 -(cong : Polynomial ℤ → Polynomial ℤ → ℤ → Prop) -(hcong : cong = fun f g m ↦ ∀ i : ℕ, m ∣ (f - g).coeff i) -(n p : ℕ) -(nppos : n > 0 ∧ p > 0) -(pprime : Nat.Prime p) -(f g h r s : Polynomial ℤ) -(hcoprime : cong (r * f + s * g) 1 p) -(hprod : cong (f * g) h p) -: (∃ F G : Polynomial ℤ, cong F f p ∧ cong G g p ∧ cong (F * G) h (p ^ n)) := -sorry + (n p : ℕ) + (nppos : n > 0 ∧ p > 0) + (pprime : Nat.Prime p) + (cong : Polynomial ℤ → Polynomial ℤ → ℤ → Prop) + (hcong : ∀ f g m, cong f g m ↔ ∀ i : ℕ, m ∣ (f - g).coeff i) + (f g h r s : Polynomial ℤ) + (hcoprime : cong (r * f + s * g) 1 p) + (hprod : cong (f * g) h p) + : (∃ F G : Polynomial ℤ, cong F f p ∧ cong G g p ∧ cong (F * G) h (p ^ n)) := + sorry diff --git a/lean4/src/putnam_1986_b5.lean b/lean4/src/putnam_1986_b5.lean index 9c392bf..252eda2 100644 --- a/lean4/src/putnam_1986_b5.lean +++ b/lean4/src/putnam_1986_b5.lean @@ -12,9 +12,11 @@ f(p(x,y,z), q(x,y,z), r(x,y,z)) = f(x,y,z). Prove or disprove the assertion that the sequence $p,q,r$ consists of some permutation of $\pm x, \pm y, \pm z$, where the number of minus signs is $0$ or $2$. -/ theorem putnam_1986_b5 -(f : MvPolynomial (Fin 3) ℝ) -(hf : f = (X 0) ^ 2 + (X 1) ^ 2 + (X 2) ^ 2 + (X 0) * (X 1) * (X 2)) -(perms : Set (Set (MvPolynomial (Fin 3) ℝ))) -(hperms : perms = {{X 0, X 1, X 2}, {X 0, -X 1, -X 2}, {-X 0, X 1, -X 2}, {-X 0, -X 1, X 2}}) -: ((∀ pqr : Fin 3 → MvPolynomial (Fin 3) ℝ, (∀ xyz : Fin 3 → ℝ, MvPolynomial.eval (fun i ↦ MvPolynomial.eval xyz (pqr i)) f = MvPolynomial.eval xyz f) → ({pqr 0, pqr 1, pqr 2} ∈ perms)) ↔ putnam_1986_b5_solution) := -sorry + (f : MvPolynomial (Fin 3) ℝ) + (perms : Set (Set (MvPolynomial (Fin 3) ℝ))) + (hf : f = (X 0) ^ 2 + (X 1) ^ 2 + (X 2) ^ 2 + (X 0) * (X 1) * (X 2)) + (hperms : perms = {{X 0, X 1, X 2}, {X 0, -X 1, -X 2}, {-X 0, X 1, -X 2}, {-X 0, -X 1, X 2}}) : + putnam_1986_b5_solution ↔ (∀ pqr : Fin 3 → MvPolynomial (Fin 3) ℝ, + (∀ xyz : Fin 3 → ℝ, MvPolynomial.eval (fun i ↦ MvPolynomial.eval xyz (pqr i)) f = MvPolynomial.eval xyz f) → + ({pqr 0, pqr 1, pqr 2} ∈ perms)) := + sorry diff --git a/lean4/src/putnam_1987_a2.lean b/lean4/src/putnam_1987_a2.lean index 6ced9af..ea4c6bd 100644 --- a/lean4/src/putnam_1987_a2.lean +++ b/lean4/src/putnam_1987_a2.lean @@ -6,11 +6,9 @@ abbrev putnam_1987_a2_solution : ℕ := sorry The sequence of digits $123456789101112131415161718192021 \dots$ is obtained by writing the positive integers in order. If the $10^n$-th digit in this sequence occurs in the part of the sequence in which the $m$-digit numbers are placed, define $f(n)$ to be $m$. For example, $f(2)=2$ because the $100$th digit enters the sequence in the placement of the two-digit integer $55$. Find, with proof, $f(1987)$. -/ theorem putnam_1987_a2 -(seqind : ℕ → ℕ) -(seqsize : ℕ → ℕ) -(f : ℕ → ℕ) -(hseqind : seqind 1 = 1 ∧ ∀ i ≥ 2, seqind i = seqind (i - 1) + (Nat.digits 10 (i - 1)).length) -(hseqsize : ∀ i ≥ 1, ∀ j : Fin ((Nat.digits 10 i).length), seqsize (seqind i + j) = (Nat.digits 10 i).length) -(hf : ∀ n : ℕ, f n = seqsize (10 ^ n)) -: f 1987 = putnam_1987_a2_solution := -sorry + (seqind seqsize f : ℕ → ℕ) + (hseqind : seqind 1 = 1 ∧ ∀ i ≥ 2, seqind i = seqind (i - 1) + (Nat.digits 10 (i - 1)).length) + (hseqsize : ∀ i ≥ 1, ∀ j : Fin ((Nat.digits 10 i).length), seqsize (seqind i + j) = (Nat.digits 10 i).length) + (hf : ∀ n : ℕ, f n = seqsize (10 ^ n)) : + f 1987 = putnam_1987_a2_solution := + sorry diff --git a/lean4/src/putnam_1987_b4.lean b/lean4/src/putnam_1987_b4.lean index e04325c..2feca16 100644 --- a/lean4/src/putnam_1987_b4.lean +++ b/lean4/src/putnam_1987_b4.lean @@ -8,10 +8,13 @@ abbrev putnam_1987_b4_solution : Prop × ℝ × Prop × ℝ := sorry Let $(x_1,y_1) = (0.8, 0.6)$ and let $x_{n+1} = x_n \cos y_n - y_n \sin y_n$ and $y_{n+1}= x_n \sin y_n + y_n \cos y_n$ for $n=1,2,3,\dots$. For each of $\lim_{n\to \infty} x_n$ and $\lim_{n \to \infty} y_n$, prove that the limit exists and find it or prove that the limit does not exist. -/ theorem putnam_1987_b4 -(x y : ℕ → ℝ) -(hxy1 : (x 1, y 1) = (0.8, 0.6)) -(hx : ∀ n ≥ 1, x (n + 1) = (x n) * cos (y n) - (y n) * sin (y n)) -(hy : ∀ n ≥ 1, y (n + 1) = (x n) * sin (y n) + (y n) * cos (y n)) -: let (existsx, limx, existsy, limy) := putnam_1987_b4_solution -(((∃ c : ℝ, Tendsto x atTop (𝓝 c)) → existsx) ∧ (existsx → Tendsto x atTop (𝓝 limx)) ∧ ((∃ c : ℝ, Tendsto y atTop (𝓝 c)) → existsy) ∧ (existsy → Tendsto y atTop (𝓝 limy))) := -sorry + (x y : ℕ → ℝ) + (hxy1 : (x 1, y 1) = (0.8, 0.6)) + (hx : ∀ n ≥ 1, x (n + 1) = (x n) * cos (y n) - (y n) * sin (y n)) + (hy : ∀ n ≥ 1, y (n + 1) = (x n) * sin (y n) + (y n) * cos (y n)) : + let (existsx, limx, existsy, limy) := putnam_1987_b4_solution + ((∃ c : ℝ, Tendsto x atTop (𝓝 c)) → existsx) ∧ + (existsx → Tendsto x atTop (𝓝 limx)) ∧ + ((∃ c : ℝ, Tendsto y atTop (𝓝 c)) → existsy) ∧ + (existsy → Tendsto y atTop (𝓝 limy)) := + sorry diff --git a/lean4/src/putnam_1987_b6.lean b/lean4/src/putnam_1987_b6.lean index 188bf38..0ad80df 100644 --- a/lean4/src/putnam_1987_b6.lean +++ b/lean4/src/putnam_1987_b6.lean @@ -6,15 +6,13 @@ open MvPolynomial Real Nat Filter Topology Let $F$ be the field of $p^2$ elements, where $p$ is an odd prime. Suppose $S$ is a set of $(p^2-1)/2$ distinct nonzero elements of $F$ with the property that for each $a\neq 0$ in $F$, exactly one of $a$ and $-a$ is in $S$. Let $N$ be the number of elements in the intersection $S \cap \{2a: a \in S\}$. Prove that $N$ is even. -/ theorem putnam_1987_b6 -(p : ℕ) -(podd : Odd p ∧ Nat.Prime p) -(F : Type*) [Field F] [Fintype F] -(Fcard : Fintype.card F = p ^ 2) -(S : Set F) -(Snz : ∀ x ∈ S, x ≠ 0) -(Scard : S.ncard = ((p : ℤ) ^ 2 - 1) / 2) -(hS : ∀ a : F, a ≠ 0 → Xor' (a ∈ S) (-a ∈ S)) -(N : ℕ) -(hN : N = (S ∩ {x | ∃ a ∈ S, x = 2 * a}).ncard) -: (Even N) := -sorry + (p : ℕ) + (F : Type*) [Field F] [Fintype F] + (S : Set F) + (hp : Odd p ∧ Nat.Prime p) + (Fcard : Fintype.card F = p ^ 2) + (Snz : ∀ x ∈ S, x ≠ 0) + (Scard : S.ncard = ((p : ℤ) ^ 2 - 1) / 2) + (hS : ∀ a : F, a ≠ 0 → Xor' (a ∈ S) (-a ∈ S)) : + (Even ((S ∩ {x | ∃ a ∈ S, x = 2 * a}).ncard)) := + sorry diff --git a/lean4/src/putnam_1988_a1.lean b/lean4/src/putnam_1988_a1.lean index e2c14f1..140da84 100644 --- a/lean4/src/putnam_1988_a1.lean +++ b/lean4/src/putnam_1988_a1.lean @@ -1,12 +1,14 @@ import Mathlib +open MeasureTheory + abbrev putnam_1988_a1_solution : ℝ := sorry -- 6 /-- Let $R$ be the region consisting of the points $(x,y)$ of the cartesian plane satisfying both $|x|-|y| \leq 1$ and $|y| \leq 1$. Find the area of $R$. -/ theorem putnam_1988_a1 -(R : Set (Fin 2 → ℝ)) -(hR : R = {p : Fin 2 → ℝ | |p 0| - |p 1| ≤ 1 ∧ |p 1| ≤ 1}) -: (MeasureTheory.volume R).toReal = putnam_1988_a1_solution := -sorry + (R : Set (Fin 2 → ℝ)) + (hR : R = {p | |p 0| - |p 1| ≤ 1 ∧ |p 1| ≤ 1}) : + (volume R).toReal = putnam_1988_a1_solution := + sorry diff --git a/lean4/src/putnam_1988_a2.lean b/lean4/src/putnam_1988_a2.lean index 84e97a9..4d356d1 100644 --- a/lean4/src/putnam_1988_a2.lean +++ b/lean4/src/putnam_1988_a2.lean @@ -8,7 +8,13 @@ abbrev putnam_1988_a2_solution : Prop := sorry A not uncommon calculus mistake is to believe that the product rule for derivatives says that $(fg)' = f'g'$. If $f(x)=e^{x^2}$, determine, with proof, whether there exists an open interval $(a,b)$ and a nonzero function $g$ defined on $(a,b)$ such that this wrong product rule is true for $x$ in $(a,b)$. -/ theorem putnam_1988_a2 -(f : ℝ → ℝ) -(hf : f = fun x ↦ Real.exp (x ^ 2)) -: ((∃ a b : ℝ, a < b ∧ ∃ g : ℝ → ℝ, (∃ x ∈ Ioo a b, g x ≠ 0) ∧ DifferentiableOn ℝ g (Ioo a b) ∧ ∀ x ∈ Ioo a b, deriv (fun y ↦ f y * g y) x = (deriv f x) * (deriv g x)) ↔ putnam_1988_a2_solution) := -sorry + (f : ℝ → ℝ) + (hf : f = fun x ↦ Real.exp (x ^ 2)) : + putnam_1988_a2_solution ↔ + (∃ a b : ℝ, + a < b ∧ + ∃ g : ℝ → ℝ, + (∃ x ∈ Ioo a b, g x ≠ 0) ∧ + DifferentiableOn ℝ g (Ioo a b) ∧ + ∀ x ∈ Ioo a b, deriv (fun y ↦ f y * g y) x = (deriv f x) * (deriv g x)) := + sorry diff --git a/lean4/src/putnam_1988_a3.lean b/lean4/src/putnam_1988_a3.lean index 2cb1b4e..5cfc796 100644 --- a/lean4/src/putnam_1988_a3.lean +++ b/lean4/src/putnam_1988_a3.lean @@ -12,5 +12,5 @@ Determine, with proof, the set of real numbers $x$ for which converges. -/ theorem putnam_1988_a3 -: ({x : ℝ | ∃ L : ℝ, Tendsto (fun t ↦ ∑ n in Finset.Icc (1 : ℕ) t, (((1 / n) / Real.sin (1 / n) - 1) ^ x)) atTop (𝓝 L)} = putnam_1988_a3_solution) := +: {x : ℝ | ∃ L : ℝ, Tendsto (fun t ↦ ∑ n in Finset.Icc (1 : ℕ) t, (((1 / n) / Real.sin (1 / n) - 1) ^ x)) atTop (𝓝 L)} = putnam_1988_a3_solution := sorry diff --git a/lean4/src/putnam_1988_a4.lean b/lean4/src/putnam_1988_a4.lean index f9284d1..c69b313 100644 --- a/lean4/src/putnam_1988_a4.lean +++ b/lean4/src/putnam_1988_a4.lean @@ -11,7 +11,10 @@ abbrev putnam_1988_a4_solution : Prop × Prop := sorry \end{enumerate} -/ theorem putnam_1988_a4 -(p : ℕ → Prop) -(hp : p = fun n ↦ ∀ color : (EuclideanSpace ℝ (Fin 2)) → Fin n, ∃ p q : EuclideanSpace ℝ (Fin 2), color p = color q ∧ dist p q = 1) -: (let (a, b) := putnam_1988_a4_solution; (p 3 ↔ a) ∧ (p 9 ↔ b)) := -sorry + (p : ℕ → Prop) + (hp : ∀ n, p n ↔ + ∀ color : (EuclideanSpace ℝ (Fin 2)) → Fin n, + ∃ p q : EuclideanSpace ℝ (Fin 2), + color p = color q ∧ dist p q = 1) : + (let (a, b) := putnam_1988_a4_solution; (p 3 ↔ a) ∧ (p 9 ↔ b)) := + sorry diff --git a/lean4/src/putnam_1988_b3.lean b/lean4/src/putnam_1988_b3.lean index 06ff361..876d537 100644 --- a/lean4/src/putnam_1988_b3.lean +++ b/lean4/src/putnam_1988_b3.lean @@ -8,7 +8,11 @@ noncomputable abbrev putnam_1988_b3_solution : ℝ := sorry For every $n$ in the set $N=\{1,2,\dots\}$ of positive integers, let $r_n$ be the minimum value of $|c-d \sqrt{3}|$ for all nonnegative integers $c$ and $d$ with $c+d=n$. Find, with proof, the smallest positive real number $g$ with $r_n \leq g$ for all $n \in N$. -/ theorem putnam_1988_b3 -(r : ℤ → ℝ) -(hr : ∀ n ≥ 1, (∃ c d : ℤ, (c ≥ 0 ∧ d ≥ 0) ∧ c + d = n ∧ r n = |c - d * Real.sqrt 3|) ∧ (∀ c d : ℤ, (c ≥ 0 ∧ d ≥ 0 ∧ c + d = n) → |c - d * Real.sqrt 3| ≥ r n)) -: putnam_1988_b3_solution > 0 ∧ (∀ n : ℤ, n ≥ 1 → r n ≤ putnam_1988_b3_solution) ∧ (∀ g > 0, (∀ n ≥ 1, r n ≤ g) → g ≥ putnam_1988_b3_solution) := -sorry + (r : ℤ → ℝ) + (hr : ∀ n ≥ 1, + (∃ c d : ℤ, + (c ≥ 0 ∧ d ≥ 0) ∧ + c + d = n ∧ r n = |c - d * Real.sqrt 3|) ∧ + (∀ c d : ℤ, (c ≥ 0 ∧ d ≥ 0 ∧ c + d = n) → |c - d * Real.sqrt 3| ≥ r n)) + : IsLeast {g : ℝ | g > 0 ∧ (∀ n : ℤ, n ≥ 1 → r n ≤ g)} putnam_1988_b3_solution := + sorry diff --git a/lean4/src/putnam_1988_b4.lean b/lean4/src/putnam_1988_b4.lean index bf27f2a..97bf15e 100644 --- a/lean4/src/putnam_1988_b4.lean +++ b/lean4/src/putnam_1988_b4.lean @@ -6,12 +6,10 @@ open Set Filter Topology Prove that if $\sum_{n=1}^\infty a_n$ is a convergent series of positive real numbers, then so is $\sum_{n=1}^\infty (a_n)^{n/(n+1)}$. -/ theorem putnam_1988_b4 -(a : ℕ → ℝ) -(appos : (ℕ → ℝ) → Prop) -(apconv : (ℕ → ℝ) → Prop) -(apposconv : (ℕ → ℝ) → Prop) -(happos : ∀ a' : ℕ → ℝ, appos a' = ∀ n ≥ 1, a' n > 0) -(hapconv : ∀ a' : ℕ → ℝ, apconv a' = ∃ s : ℝ, Tendsto (fun N : ℕ => ∑ n : Set.Icc 1 N, a' n) atTop (𝓝 s)) -(happosconv : ∀ a' : ℕ → ℝ, apposconv a' = (appos a' ∧ apconv a')) -: apposconv a → apposconv (fun n : ℕ => (a n) ^ ((n : ℝ) / (n + 1))) := -sorry + (a : ℕ → ℝ) + (IsPosConv : (ℕ → ℝ) → Prop) + (IsPosConv_def : ∀ a' : ℕ → ℝ, IsPosConv a' ↔ + (∀ n ≥ 1, a' n > 0) ∧ + (∃ s : ℝ, Tendsto (fun N : ℕ => ∑ n : Set.Icc 1 N, a' n) atTop (𝓝 s))) : + (IsPosConv a) → IsPosConv (fun n : ℕ => (a n) ^ ((n : ℝ) / (n + 1))) := + sorry diff --git a/lean4/src/putnam_1988_b5.lean b/lean4/src/putnam_1988_b5.lean index 5ac0360..b0ccde7 100644 --- a/lean4/src/putnam_1988_b5.lean +++ b/lean4/src/putnam_1988_b5.lean @@ -12,11 +12,10 @@ M_2&=\begin{pmatrix} 0 & -1 & -1 & 1 & 1 \\ 1 & 0 & -1 & -1 & 1 \\ 1 & 1 & 0 & - \end{align*} -/ theorem putnam_1988_b5 -(n : ℕ) -(Mn : Matrix (Fin (2 * n + 1)) (Fin (2 * n + 1)) ℝ) -(npos : n > 0) -(Mnskewsymm : ∀ i j : Fin (2 * n + 1), Mn i j = -(Mn j i)) -(hMn1 : ∀ i j : Fin (2 * n + 1), (1 ≤ (i.1 : ℤ) - j.1 ∧ (i.1 : ℤ) - j.1 ≤ n) → Mn i j = 1) -(hMnn1 : ∀ i j : Fin (2 * n + 1), (i.1 : ℤ) - j.1 > n → Mn i j = -1) -: Mn.rank = putnam_1988_b5_solution n := -sorry + (n : ℕ) (hn : n > 0) + (Mn : Matrix (Fin (2 * n + 1)) (Fin (2 * n + 1)) ℝ) + (Mnskewsymm : ∀ i j, Mn i j = -(Mn j i)) + (hMn1 : ∀ i j, (1 ≤ (i.1 : ℤ) - j.1 ∧ (i.1 : ℤ) - j.1 ≤ n) → Mn i j = 1) + (hMnn1 : ∀ i j, (i.1 : ℤ) - j.1 > n → Mn i j = -1) : + Mn.rank = putnam_1988_b5_solution n := + sorry diff --git a/lean4/src/putnam_1989_b1.lean b/lean4/src/putnam_1989_b1.lean index 3fc5bc3..06ed6ba 100644 --- a/lean4/src/putnam_1989_b1.lean +++ b/lean4/src/putnam_1989_b1.lean @@ -1,6 +1,6 @@ import Mathlib -open Nat +open Nat MeasureTheory abbrev putnam_1989_b1_solution : ℤ × ℤ × ℤ × ℤ := sorry -- (4, 2, -5, 3) @@ -8,14 +8,13 @@ abbrev putnam_1989_b1_solution : ℤ × ℤ × ℤ × ℤ := sorry A dart, thrown at random, hits a square target. Assuming that any two parts of the target of equal area are equally likely to be hit, find the probability that the point hit is nearer to the center than to any edge. Express your answer in the form $(a\sqrt{b}+c)/d$, where $a$, $b$, $c$, $d$ are integers and $b$, $d$ are positive. -/ theorem putnam_1989_b1 -(square : Set (EuclideanSpace ℝ (Fin 2))) -(square_def : square = {p : EuclideanSpace ℝ (Fin 2) | ∀ i : Fin 2, p i ∈ Set.Icc 0 1}) -(perimeter : Set (EuclideanSpace ℝ (Fin 2))) -(perimeter_def : perimeter = {p ∈ square | p 0 = 0 ∨ p 0 = 1 ∨ p 1 = 0 ∨ p 1 = 1}) -(center : EuclideanSpace ℝ (Fin 2)) -(center_def : center = (fun i : Fin 2 => 1 / 2)) -(Scloser : Set (EuclideanSpace ℝ (Fin 2))) -(hScloser : Scloser = {p ∈ square | ∀ q ∈ perimeter, dist p center < dist p q}) -: let (a, b, c, d) := putnam_1989_b1_solution; b > 0 ∧ d > 0 ∧ (¬∃ n : ℤ, n^2 = b) ∧ -(MeasureTheory.volume Scloser).toReal / (MeasureTheory.volume square).toReal = (a * Real.sqrt b + c) / d := -sorry + (square Scloser perimeter: Set (EuclideanSpace ℝ (Fin 2))) + (center : EuclideanSpace ℝ (Fin 2)) + (square_def : square = {p | ∀ i : Fin 2, p i ∈ Set.Icc 0 1}) + (perimeter_def : perimeter = {p ∈ square | p 0 = 0 ∨ p 0 = 1 ∨ p 1 = 0 ∨ p 1 = 1}) + (center_def : center = (fun i : Fin 2 => 1 / 2)) + (hScloser : Scloser = {p ∈ square | ∀ q ∈ perimeter, dist p center < dist p q}) : + let (a, b, c, d) := putnam_1989_b1_solution; + b > 0 ∧ d > 0 ∧ (¬∃ n : ℤ, n^2 = b) ∧ + (volume Scloser).toReal / (volume square).toReal = (a * Real.sqrt b + c) / d := + sorry diff --git a/lean4/src/putnam_1989_b4.lean b/lean4/src/putnam_1989_b4.lean index fc3673e..488c8b6 100644 --- a/lean4/src/putnam_1989_b4.lean +++ b/lean4/src/putnam_1989_b4.lean @@ -7,6 +7,12 @@ abbrev putnam_1989_b4_solution : Prop := sorry /-- Can a countably infinite set have an uncountable collection of non-empty subsets such that the intersection of any two of them is finite? -/ -theorem putnam_1989_b4 -: ((∃ S : Type, Countable S ∧ Infinite S ∧ ∃ C : Set (Set S), ¬Countable C ∧ (∀ R ∈ C, R ≠ ∅) ∧ (∀ A ∈ C, ∀ B ∈ C, A ≠ B → (A ∩ B).Finite)) ↔ putnam_1989_b4_solution) := +theorem putnam_1989_b4 : + (∃ S : Type, + Countable S ∧ Infinite S ∧ + ∃ C : Set (Set S), + ¬Countable C ∧ + (∀ R ∈ C, R ≠ ∅) ∧ + (∀ A ∈ C, ∀ B ∈ C, A ≠ B → (A ∩ B).Finite) + ) ↔ putnam_1989_b4_solution := sorry diff --git a/lean4/src/putnam_1990_a1.lean b/lean4/src/putnam_1990_a1.lean index 336fc49..059e179 100644 --- a/lean4/src/putnam_1990_a1.lean +++ b/lean4/src/putnam_1990_a1.lean @@ -8,8 +8,8 @@ abbrev putnam_1990_a1_solution : (ℕ → ℤ) × (ℕ → ℤ) := sorry Let $T_0=2,T_1=3,T_2=6$, and for $n \geq 3$, $T_n=(n+4)T_{n-1}-4nT_{n-2}+(4n-8)T_{n-3}$. The first few terms are $2,3,6,14,40,152,784,5168,40576$. Find, with proof, a formula for $T_n$ of the form $T_n=A_n+B_n$, where $\{A_n\}$ and $\{B_n\}$ are well-known sequences. -/ theorem putnam_1990_a1 -(T : ℕ → ℤ) -(hT012 : T 0 = 2 ∧ T 1 = 3 ∧ T 2 = 6) -(hTn : ∀ n ≥ 3, T n = (n + 4) * T (n - 1) - 4 * n * T (n - 2) + (4 * n - 8) * T (n - 3)) -: T = putnam_1990_a1_solution.1 + putnam_1990_a1_solution.2 := -sorry + (T : ℕ → ℤ) + (hT012 : T 0 = 2 ∧ T 1 = 3 ∧ T 2 = 6) + (hTn : ∀ n, T (n + 3) = (n + 7) * T (n + 2) - 4 * (n + 3) * T (n + 1) + (4 * n + 4) * T n) : + T = putnam_1990_a1_solution.1 + putnam_1990_a1_solution.2 := + sorry diff --git a/lean4/src/putnam_1990_a2.lean b/lean4/src/putnam_1990_a2.lean index 7a79cc1..77ad74f 100644 --- a/lean4/src/putnam_1990_a2.lean +++ b/lean4/src/putnam_1990_a2.lean @@ -8,7 +8,10 @@ abbrev putnam_1990_a2_solution : Prop := sorry Is $\sqrt{2}$ the limit of a sequence of numbers of the form $\sqrt[3]{n}-\sqrt[3]{m}$ ($n,m=0,1,2,\dots$)? -/ theorem putnam_1990_a2 -(numform : ℝ → Prop) -(hnumform : ∀ x : ℝ, numform x = ∃ n m : ℕ, x = n ^ ((1 : ℝ) / 3) - m ^ ((1 : ℝ) / 3)) -: (∃ s : ℕ → ℝ, (∀ i : ℕ, numform (s i)) ∧ Tendsto s atTop (𝓝 (Real.sqrt 2))) ↔ putnam_1990_a2_solution := + (numform : ℝ → Prop) + (hnumform : ∀ x : ℝ, numform x ↔ ∃ n m : ℕ, x = n ^ ((1 : ℝ) / 3) - m ^ ((1 : ℝ) / 3)) : + putnam_1990_a2_solution ↔ + (∃ s : ℕ → ℝ, + (∀ i : ℕ, numform (s i)) ∧ + Tendsto s atTop (𝓝 (Real.sqrt 2))) := sorry diff --git a/lean4/src/putnam_1990_a5.lean b/lean4/src/putnam_1990_a5.lean index 55713e9..dc1e2ec 100644 --- a/lean4/src/putnam_1990_a5.lean +++ b/lean4/src/putnam_1990_a5.lean @@ -7,6 +7,8 @@ abbrev putnam_1990_a5_solution : Prop := sorry /-- If $\mathbf{A}$ and $\mathbf{B}$ are square matrices of the same size such that $\mathbf{ABAB}=\mathbf{0}$, does it follow that $\mathbf{BABA}=\mathbf{0}$? -/ -theorem putnam_1990_a5 -: (∀ n ≥ 1, ∀ A B : Matrix (Fin n) (Fin n) ℝ, A * B * A * B = 0 → B * A * B * A = 0) ↔ putnam_1990_a5_solution := +theorem putnam_1990_a5 : + putnam_1990_a5_solution ↔ + (∀ n ≥ 1, ∀ A B : Matrix (Fin n) (Fin n) ℝ, + A * B * A * B = 0 → B * A * B * A = 0) := sorry diff --git a/lean4/src/putnam_1991_a2.lean b/lean4/src/putnam_1991_a2.lean index 6b941f0..8bcdb40 100644 --- a/lean4/src/putnam_1991_a2.lean +++ b/lean4/src/putnam_1991_a2.lean @@ -8,7 +8,9 @@ abbrev putnam_1991_a2_solution : Prop := sorry Let $\mathbf{A}$ and $\mathbf{B}$ be different $n \times n$ matrices with real entries. If $\mathbf{A}^3=\mathbf{B}^3$ and $\mathbf{A}^2\mathbf{B}=\mathbf{B}^2\mathbf{A}$, can $\mathbf{A}^2+\mathbf{B}^2$ be invertible? -/ theorem putnam_1991_a2 -(n : ℕ) -(npos : n ≥ 1) -: (∃ A B : Matrix (Fin n) (Fin n) ℝ, A ≠ B ∧ A ^ 3 = B ^ 3 ∧ A ^ 2 * B = B ^ 2 * A ∧ Nonempty (Invertible (A ^ 2 + B ^ 2))) ↔ putnam_1991_a2_solution := + (n : ℕ) (hn : 1 ≤ n) : + putnam_1991_a2_solution ↔ (∃ A B : Matrix (Fin n) (Fin n) ℝ, + A ≠ B ∧ A ^ 3 = B ^ 3 ∧ + A ^ 2 * B = B ^ 2 * A ∧ + Nonempty (Invertible (A ^ 2 + B ^ 2))) := sorry diff --git a/lean4/src/putnam_1991_a3.lean b/lean4/src/putnam_1991_a3.lean index 50590b6..bcbcf0b 100644 --- a/lean4/src/putnam_1991_a3.lean +++ b/lean4/src/putnam_1991_a3.lean @@ -14,9 +14,12 @@ Find all real polynomials $p(x)$ of degree $n \geq 2$ for which there exist real where $p'(x)$ denotes the derivative of $p(x)$. -/ theorem putnam_1991_a3 -(p : Polynomial ℝ) -(n : ℕ) -(hn : n = p.degree) -(hge : n ≥ 2) -: (∃ r : ℕ → ℝ, (∀ i : Fin (n - 1), r i < r (i + 1)) ∧ (∀ i : Fin n, p.eval (r i) = 0) ∧ (∀ i : Fin (n - 1), (Polynomial.derivative p).eval ((r i + r (i + 1)) / 2) = 0)) ↔ p ∈ putnam_1991_a3_solution := + (p : Polynomial ℝ) + (n : ℕ) + (hn : n = p.degree) + (hge : n ≥ 2) : + p ∈ putnam_1991_a3_solution ↔ + (∃ r : ℕ → ℝ, (∀ i : Fin (n - 1), r i < r (i + 1)) ∧ + (∀ i : Fin n, p.eval (r i) = 0) ∧ + (∀ i : Fin (n - 1), (Polynomial.derivative p).eval ((r i + r (i + 1)) / 2) = 0)) := sorry diff --git a/lean4/src/putnam_1991_a5.lean b/lean4/src/putnam_1991_a5.lean index 704e40b..984721a 100644 --- a/lean4/src/putnam_1991_a5.lean +++ b/lean4/src/putnam_1991_a5.lean @@ -8,7 +8,7 @@ noncomputable abbrev putnam_1991_a5_solution : ℝ := sorry Find the maximum value of $\int_0^y \sqrt{x^4+(y-y^2)^2}\,dx$ for $0 \leq y \leq 1$. -/ theorem putnam_1991_a5 -(f : Set.Icc (0 : ℝ) 1 → ℝ) -(hf : ∀ y : Set.Icc 0 1, f y = ∫ x in Set.Ioo 0 y, Real.sqrt (x ^ 4 + (y - y ^ 2) ^ 2)) -: (∃ y : Set.Icc 0 1, f y = putnam_1991_a5_solution) ∧ (∀ y : Set.Icc 0 1, f y ≤ putnam_1991_a5_solution) := + (f : Set.Icc (0 : ℝ) 1 → ℝ) + (hf : ∀ y : Set.Icc 0 1, f y = ∫ x in Set.Ioo 0 y, Real.sqrt (x ^ 4 + (y - y ^ 2) ^ 2)) : + IsGreatest (f '' (Set.Icc 0 1)) putnam_1991_a5_solution := sorry diff --git a/lean4/src/putnam_1991_a6.lean b/lean4/src/putnam_1991_a6.lean index 8f7099e..30928e7 100644 --- a/lean4/src/putnam_1991_a6.lean +++ b/lean4/src/putnam_1991_a6.lean @@ -13,16 +13,18 @@ Let $A(n)$ denote the number of sums of positive integers $a_1+a_2+\cdots+a_r$ w Prove that $A(n)=B(n)$ for each $n \geq 1$. (For example, $A(7)=5$ because the relevant sums are $7,6+1,5+2,4+3,4+2+1$, and $B(7)=5$ because the relevant sums are $4+2+1,2+2+2+1,2+2+1+1+1,2+1+1+1+1+1,1+1+1+1+1+1+1$.) -/ theorem putnam_1991_a6 -(nabsum : ℕ → ℕ × (ℕ → ℕ) → Prop) -(agt bge bg1 bg2 : ℕ × (ℕ → ℕ) → Prop) -(A g B: ℕ → ℕ) -(hnabsum : ∀ n ≥ 1, ∀ ab : ℕ × (ℕ → ℕ), nabsum n ab = (ab.1 ≥ 1 ∧ (∀ i < ab.1, ab.2 i > 0) ∧ (∀ i ≥ ab.1, ab.2 i = 0) ∧ (∑ i : Fin ab.1, ab.2 i) = n)) -(hagt : ∀ a : ℕ × (ℕ → ℕ), agt a ↔ (∀ i : Fin (a.1 - 2), a.2 i > a.2 (i + 1) + a.2 (i + 2)) ∧ 1 < a.1 → a.2 (a.1 - 2) > a.2 (a.1 - 1)) -(hA : ∀ n ≥ 1, A n = {a : ℕ × (ℕ → ℕ) | nabsum n a ∧ agt a}.encard) -(hbge : ∀ b : ℕ × (ℕ → ℕ), bge b = ∀ i : Fin (b.1 - 1), b.2 i ≥ b.2 (i + 1)) -(hg : g 0 = 1 ∧ g 1 = 2 ∧ (∀ j ≥ 2, g j = g (j - 1) + g (j - 2) + 1)) -(hbg1 : ∀ b : ℕ × (ℕ → ℕ), bg1 b ↔ ∀ i : Fin b.1, ∃ j : ℕ, b.2 i = g j) -(hbg2 : ∀ b : ℕ × (ℕ → ℕ), bg2 b ↔ ∃ k : ℕ, b.2 0 = g k ∧ (∀ j ≤ k, ∃ i : Fin b.1, b.2 i = g j)) -(hB : ∀ n ≥ 1, B n = {b : ℕ × (ℕ → ℕ) | nabsum n b ∧ bge b ∧ bg1 b ∧ bg2 b}.encard) -: ∀ n ≥ 1, A n = B n := + (nabsum : ℕ → ℕ × (ℕ → ℕ) → Prop) + (agt bge bg1 bg2 : ℕ × (ℕ → ℕ) → Prop) + (A g B: ℕ → ℕ) + (hnabsum : ∀ᵉ (n ≥ 1) (ab), nabsum n ab ↔ + (ab.1 ≥ 1 ∧ (∀ i < ab.1, ab.2 i > 0) ∧ + (∀ i ≥ ab.1, ab.2 i = 0) ∧ (∑ i : Fin ab.1, ab.2 i) = n)) + (hA : ∀ n ≥ 1, A n = {a | nabsum n a ∧ + (∀ i : Fin (a.1 - 2), a.2 i > a.2 (i + 1) + a.2 (i + 2)) ∧ 1 < a.1 → a.2 (a.1 - 2) > a.2 (a.1 - 1)}.encard) + (hg : g 0 = 1 ∧ g 1 = 2 ∧ (∀ j ≥ 2, g j = g (j - 1) + g (j - 2) + 1)) + (hB : ∀ n ≥ 1, B n = {b | nabsum n b ∧ + (∀ i : Fin (b.1 - 1), b.2 i ≥ b.2 (i + 1)) ∧ + (∀ i : Fin b.1, ∃ j : ℕ, b.2 i = g j) ∧ + (∃ k : ℕ, b.2 0 = g k ∧ (∀ j ≤ k, ∃ i : Fin b.1, b.2 i = g j))}.encard) : + ∀ n ≥ 1, (A n) = (B n) := sorry diff --git a/lean4/src/putnam_1991_b1.lean b/lean4/src/putnam_1991_b1.lean index 2dd2dde..62002c8 100644 --- a/lean4/src/putnam_1991_b1.lean +++ b/lean4/src/putnam_1991_b1.lean @@ -8,13 +8,12 @@ abbrev putnam_1991_b1_solution : Set ℤ := sorry For each integer $n \geq 0$, let $S(n)=n-m^2$, where $m$ is the greatest integer with $m^2 \leq n$. Define a sequence $(a_k)_{k=0}^\infty$ by $a_0=A$ and $a_{k+1}=a_k+S(a_k)$ for $k \geq 0$. For what positive integers $A$ is this sequence eventually constant? -/ theorem putnam_1991_b1 -(m : ℤ → ℤ) -(S : ℤ → ℤ) -(A : ℤ) -(a : ℕ → ℤ) -(hm : ∀ n : ℤ, n ≥ 0 → (m n) ^ 2 ≤ n ∧ (∀ m' : ℤ, m' ^ 2 ≤ n → m' ≤ m n)) -(hS : ∀ n : ℤ, n ≥ 0 → S n = n - (m n) ^ 2) -(ha : a 0 = A ∧ (∀ k : ℕ, a (k + 1) = a k + S (a k))) -(hA : A > 0) -: (∃ (K : ℕ) (c : ℕ), ∀ k ≥ K, a k = c) ↔ A ∈ putnam_1991_b1_solution := + (m S : ℤ → ℤ) + (A : ℤ) + (a : ℕ → ℤ) + (hm : ∀ n, 0 ≤ n → (m n) ^ 2 ≤ n ∧ (∀ m' : ℤ, m' ^ 2 ≤ n → m' ≤ m n)) + (hS : ∀ n, 0 ≤ n → S n = n - (m n) ^ 2) + (ha : a 0 = A ∧ (∀ k, a (k + 1) = a k + S (a k))) + (hA : A > 0) : + (∃ (K : ℕ) (c : ℕ), ∀ k ≥ K, a k = c) ↔ A ∈ putnam_1991_b1_solution := sorry diff --git a/lean4/src/putnam_1991_b2.lean b/lean4/src/putnam_1991_b2.lean index 9243453..83e642c 100644 --- a/lean4/src/putnam_1991_b2.lean +++ b/lean4/src/putnam_1991_b2.lean @@ -11,10 +11,10 @@ g(x+y)&=f(x)g(y)+g(x)f(y). If $f'(0)=0$, prove that $(f(x))^2+(g(x))^2=1$ for all $x$. -/ theorem putnam_1991_b2 -(f g : ℝ → ℝ) -(fgnconst : ¬∃ c : ℝ, f = Function.const ℝ c ∨ g = Function.const ℝ c) -(fgdiff : Differentiable ℝ f ∧ Differentiable ℝ g) -(fadd : ∀ x y : ℝ, f (x + y) = f x * f y - g x * g y) -(gadd : ∀ x y : ℝ, g (x + y) = f x * g y + g x * f y) -: (deriv f 0 = 0) → (∀ x : ℝ, (f x) ^ 2 + (g x) ^ 2 = 1) := + (f g : ℝ → ℝ) + (fgnconst : ¬∃ c, f = Function.const ℝ c ∨ g = Function.const ℝ c) + (fgdiff : Differentiable ℝ f ∧ Differentiable ℝ g) + (fadd : ∀ x y, f (x + y) = f x * f y - g x * g y) + (gadd : ∀ x y, g (x + y) = f x * g y + g x * f y) : + (deriv f 0 = 0) → (∀ x, (f x) ^ 2 + (g x) ^ 2 = 1) := sorry diff --git a/lean4/src/putnam_1991_b6.lean b/lean4/src/putnam_1991_b6.lean index 1754fa9..3e53d84 100644 --- a/lean4/src/putnam_1991_b6.lean +++ b/lean4/src/putnam_1991_b6.lean @@ -8,9 +8,10 @@ noncomputable abbrev putnam_1991_b6_solution : ℝ → ℝ → ℝ := sorry Let $a$ and $b$ be positive numbers. Find the largest number $c$, in terms of $a$ and $b$, such that $a^xb^{1-x} \leq a\frac{\sinh ux}{\sinh u}+b\frac{\sinh u(1-x)}{\sinh u}$ for all $u$ with $0<|u| \leq c$ and for all $x$, $0 0 ∧ b > 0) -(hcle : ∀ c : ℝ, cle c = ∀ u : ℝ, (0 < |u| ∧ |u| ≤ c) → (∀ x ∈ Set.Ioo 0 1, a ^ x * b ^ (1 - x) ≤ a * (Real.sinh (u * x) / Real.sinh u) + b * (Real.sinh (u * (1 - x)) / Real.sinh u))) -: cle (putnam_1991_b6_solution a b) ∧ (∀ c : ℝ, cle c → c ≤ putnam_1991_b6_solution a b) := + (a b : ℝ) + (abpos : a > 0 ∧ b > 0) : + IsGreatest {c | ∀ u, + (0 < |u| ∧ |u| ≤ c) → + (∀ x ∈ Set.Ioo 0 1, a ^ x * b ^ (1 - x) ≤ a * (Real.sinh (u * x) / Real.sinh u) + b * (Real.sinh (u * (1 - x)) / Real.sinh u))} + (putnam_1991_b6_solution a b) := sorry diff --git a/lean4/src/putnam_1992_a1.lean b/lean4/src/putnam_1992_a1.lean index 2b72fa8..f8008c4 100644 --- a/lean4/src/putnam_1992_a1.lean +++ b/lean4/src/putnam_1992_a1.lean @@ -11,6 +11,7 @@ Prove that $f(n) = 1-n$ is the only integer-valued function defined on the integ \end{itemize} -/ theorem putnam_1992_a1 -(f : ℤ → ℤ) -: ((f = fun n ↦ 1 - n) ↔ (∀ n : ℤ, f (f n) = n) ∧ (∀ n : ℤ, f (f (n + 2) + 2) = n) ∧ (f 0 = 1)) := + (f : ℤ → ℤ) : + (f = fun n ↦ 1 - n) ↔ + (∀ n : ℤ, f (f n) = n) ∧ (∀ n : ℤ, f (f (n + 2) + 2) = n) ∧ (f 0 = 1) := sorry diff --git a/lean4/src/putnam_1992_a5.lean b/lean4/src/putnam_1992_a5.lean index fc0a2db..0dc0927 100644 --- a/lean4/src/putnam_1992_a5.lean +++ b/lean4/src/putnam_1992_a5.lean @@ -10,7 +10,8 @@ a_{k+j} = a_{k+m+j} = a_{k+2m+j}, for $0 \leq j \leq m-1$. -/ theorem putnam_1992_a5 -(a : ℕ → ℕ) -(ha : a = fun n ↦ ite (Even {i | (digits 2 n).get i = 1}.ncard) 0 1) -: (¬∃ k > 0, ∃ m > 0, ∀ j ≤ m - 1, a (k + j) = a (k + m + j) ∧ a (k + m + j) = a (k + 2 * m + j)) := + (a : ℕ → ℕ) + (ha : a = fun n ↦ ite (Even {i | (digits 2 n).get i = 1}.ncard) 0 1) : + ¬∃ k > 0, ∃ m > 0, ∀ j ≤ m - 1, + a (k + j) = a (k + m + j) ∧ a (k + m + j) = a (k + 2 * m + j) := sorry diff --git a/lean4/src/putnam_1992_b1.lean b/lean4/src/putnam_1992_b1.lean index 415c456..3bd301b 100644 --- a/lean4/src/putnam_1992_b1.lean +++ b/lean4/src/putnam_1992_b1.lean @@ -8,12 +8,8 @@ abbrev putnam_1992_b1_solution : ℕ → ℤ := sorry Let $S$ be a set of $n$ distinct real numbers. Let $A_S$ be the set of numbers that occur as averages of two distinct elements of $S$. For a given $n \geq 2$, what is the smallest possible number of elements in $A_S$? -/ theorem putnam_1992_b1 -(n : ℕ) -(nge2 : n ≥ 2) -(A : Finset ℝ → Set ℝ) -(hA : A = fun S ↦ {x | ∃ a ∈ S, ∃ b ∈ S, a ≠ b ∧ (a + b) / 2 = x}) -(min : ℤ) -(hmineq : ∃ S : Finset ℝ, S.card = n ∧ min = (A S).ncard) -(hminlb : ∀ S : Finset ℝ, S.card = n → min ≤ (A S).ncard) -: (min = putnam_1992_b1_solution n) := + (n : ℕ) (hn : n ≥ 2) + (A : Finset ℝ → Set ℝ) + (hA : A = fun S ↦ {x | ∃ a ∈ S, ∃ b ∈ S, a ≠ b ∧ (a + b) / 2 = x}) : + IsLeast {k : ℤ | ∃ S : Finset ℝ, S.card = n ∧ k = (A S).ncard} (putnam_1992_b1_solution n) := sorry diff --git a/lean4/src/putnam_1992_b3.lean b/lean4/src/putnam_1992_b3.lean index d3005a4..5a7a24c 100644 --- a/lean4/src/putnam_1992_b3.lean +++ b/lean4/src/putnam_1992_b3.lean @@ -13,7 +13,7 @@ a_{n+1}(x,y)&=\frac{(a_n(x,y))^2+y^2}{2},\text{ for $n \geq 0$.} Find the area of the region $\{(x,y) \mid (a_n(x,y))_{n \geq 0}\text{ converges}\}$. -/ theorem putnam_1992_b3 -(a : (Fin 2 → ℝ) → (ℕ → ℝ)) -(ha : ∀ p : Fin 2 → ℝ, (a p) 0 = p 0 ∧ (∀ n : ℕ, (a p) (n + 1) = (((a p) n) ^ 2 + (p 1) ^ 2) / 2)) -: (MeasureTheory.volume {p : Fin 2 → ℝ | ∃ L : ℝ, Tendsto (a p) atTop (𝓝 L)}).toReal = putnam_1992_b3_solution := + (a : (Fin 2 → ℝ) → (ℕ → ℝ)) + (ha : ∀ p, (a p) 0 = p 0 ∧ (∀ n, (a p) (n + 1) = (((a p) n) ^ 2 + (p 1) ^ 2) / 2)) : + putnam_1992_b3_solution = (MeasureTheory.volume {p | ∃ L, Tendsto (a p) atTop (𝓝 L)}).toReal := sorry diff --git a/lean4/src/putnam_1992_b4.lean b/lean4/src/putnam_1992_b4.lean index b42caf9..37788b6 100644 --- a/lean4/src/putnam_1992_b4.lean +++ b/lean4/src/putnam_1992_b4.lean @@ -12,12 +12,9 @@ Let $p(x)$ be a nonzero polynomial of degree less than $1992$ having no nonconst for polynomials $f(x)$ and $g(x)$. Find the smallest possible degree of $f(x)$. -/ theorem putnam_1992_b4 -(valid : Polynomial ℝ → Prop) -(hvalid : valid = fun p ↦ p ≠ 0 ∧ p.degree < 1992 ∧ IsCoprime p (X ^ 3 - X)) -(pair : Polynomial ℝ → Polynomial ℝ → Prop) -(hpair : pair = fun p f ↦ ∃ g : Polynomial ℝ, iteratedDeriv 1992 (fun x ↦ p.eval x / (x ^ 3 - x)) = fun x ↦ f.eval x / g.eval x) -(min : ℕ) -(hmineq : ∃ p f : Polynomial ℝ, (valid p ∧ pair p f) ∧ min = f.degree) -(hminlb : ∀ p f : Polynomial ℝ, (valid p ∧ pair p f) → min ≤ f.degree) -: (min = putnam_1992_b4_solution) := + (IsValid : Polynomial ℝ → Prop) + (pair : Polynomial ℝ → Polynomial ℝ → Prop) + (IsValid_def : ∀ p, IsValid p ↔ p ≠ 0 ∧ p.degree < 1992 ∧ IsCoprime p (X ^ 3 - X)) + (hpair : ∀ p f, pair p f ↔ ∃ g : Polynomial ℝ, iteratedDeriv 1992 (fun x ↦ p.eval x / (x ^ 3 - x)) = fun x ↦ f.eval x / g.eval x) : + IsLeast {k : ℕ | ∃ p f, IsValid p ∧ pair p f ∧ k = f.degree} putnam_1992_b4_solution := sorry diff --git a/lean4/src/putnam_1992_b5.lean b/lean4/src/putnam_1992_b5.lean index 84a3723..9579273 100644 --- a/lean4/src/putnam_1992_b5.lean +++ b/lean4/src/putnam_1992_b5.lean @@ -21,7 +21,7 @@ Let $D_n$ denote the value of the $(n-1) \times (n-1)$ determinant Is the set $\left\{ \frac{D_n}{n!} \right\}_{n \geq 2}$ bounded? -/ theorem putnam_1992_b5 -(D : ℕ → ℚ) -(hD : D = fun (n : ℕ) ↦ Matrix.det (fun i j : Fin (n - 1) ↦ ite (i = j) ((i : ℕ) + 3 : ℚ) 1)) -: ((Bornology.IsBounded {x | ∃ n ≥ 2, D n / factorial n = x}) ↔ putnam_1992_b5_solution) := + (D : ℕ → ℚ) + (hD : ∀ n, D n = Matrix.det (fun i j : Fin (n - 1) ↦ ite (i = j) ((i : ℕ) + 3 : ℚ) 1)) : + putnam_1992_b5_solution ↔ (Bornology.IsBounded {x | ∃ n ≥ 2, D n / factorial n = x}) := sorry diff --git a/lean4/src/putnam_1992_b6.lean b/lean4/src/putnam_1992_b6.lean index 4d18170..98f4ffd 100644 --- a/lean4/src/putnam_1992_b6.lean +++ b/lean4/src/putnam_1992_b6.lean @@ -13,12 +13,12 @@ Let $M$ be a set of real $n \times n$ matrices such that Prove that $M$ contains at most $n^2$ matrices. -/ theorem putnam_1992_b6 -(n : ℕ) -(npos : n > 0) -(M : Set (Matrix (Fin n) (Fin n) ℝ)) -(h1 : 1 ∈ M) -(h2 : ∀ A ∈ M, ∀ B ∈ M, Xor' (A * B ∈ M) (-A * B ∈ M)) -(h3 : ∀ A ∈ M, ∀ B ∈ M, (A * B = B * A) ∨ (A * B = -B * A)) -(h4 : ∀ A ∈ M, A ≠ 1 → ∃ B ∈ M, A * B = -B * A) -: (M.encard ≤ n ^ 2) := + (n : ℕ) + (npos : 0 < n) + (M : Set (Matrix (Fin n) (Fin n) ℝ)) + (h1 : 1 ∈ M) + (h2 : ∀ A ∈ M, ∀ B ∈ M, Xor' (A * B ∈ M) (-A * B ∈ M)) + (h3 : ∀ A ∈ M, ∀ B ∈ M, (A * B = B * A) ∨ (A * B = -B * A)) + (h4 : ∀ A ∈ M, A ≠ 1 → ∃ B ∈ M, A * B = -B * A) : + M.encard ≤ n ^ 2 := sorry diff --git a/lean4/src/putnam_1993_a4.lean b/lean4/src/putnam_1993_a4.lean index ed0fd3e..eafaf67 100644 --- a/lean4/src/putnam_1993_a4.lean +++ b/lean4/src/putnam_1993_a4.lean @@ -4,9 +4,10 @@ import Mathlib Let $x_1,x_2,\dots,x_{19}$ be positive integers each of which is less than or equal to $93$. Let $y_1,y_2,\dots,y_{93}$ be positive integers each of which is less than or equal to $19$. Prove that there exists a (nonempty) sum of some $x_i$'s equal to a sum of some $y_j$'s. -/ theorem putnam_1993_a4 -(x : Fin 19 → ℤ) -(y : Fin 93 → ℤ) -(hx : ∀ i : Fin 19, x i > 0 ∧ x i ≤ 93) -(hy : ∀ j : Fin 93, y j > 0 ∧ y j ≤ 19) -: ∃ (is : Finset (Fin 19)) (js : Finset (Fin 93)), is ≠ ∅ ∧ (∑ i : is, x i) = (∑ j : js, y j) := + (x : Fin 19 → ℤ) + (y : Fin 93 → ℤ) + (hx : ∀ i, 0 < x i ∧ x i ≤ 93) + (hy : ∀ j, 0 < y j ∧ y j ≤ 19) : + ∃ (is : Finset (Fin 19)) (js : Finset (Fin 93)), + is ≠ ∅ ∧ (∑ i : is, x i) = (∑ j : js, y j) := sorry diff --git a/lean4/src/putnam_1993_a6.lean b/lean4/src/putnam_1993_a6.lean index 3e0886d..414aea0 100644 --- a/lean4/src/putnam_1993_a6.lean +++ b/lean4/src/putnam_1993_a6.lean @@ -5,7 +5,7 @@ The infinite sequence of $2$'s and $3$'s $2,3,3,2,3,3,3,2,3,3,3,2,3,3,2,3,3,3,2, -/ theorem putnam_1993_a6 (seq : ℕ → ℤ) -(hseq23 : ∀ n : ℕ, seq n = 2 ∨ seq n = 3) -(hseq2inds : ∀ n : ℕ, seq n = 2 ↔ (∃ N : ℕ, n = ∑ i : Fin N, (seq i + 1))) -: ∃ r : ℝ, ∀ n : ℕ, seq n = 2 ↔ (∃ m : ℤ, n + 1 = 1 + Int.floor (r * m)) := +(hseq23 : ∀ n, seq n = 2 ∨ seq n = 3) +(hseq2inds : ∀ n, seq n = 2 ↔ (∃ N : ℕ, n = ∑ i : Fin N, (seq i + 1))) +: ∃ r : ℝ, ∀ n, seq n = 2 ↔ (∃ m : ℤ, n + 1 = 1 + Int.floor (r * m)) := sorry diff --git a/lean4/src/putnam_1993_b1.lean b/lean4/src/putnam_1993_b1.lean index 8880750..21dfb84 100644 --- a/lean4/src/putnam_1993_b1.lean +++ b/lean4/src/putnam_1993_b1.lean @@ -5,8 +5,10 @@ abbrev putnam_1993_b1_solution : ℕ := sorry /-- Find the smallest positive integer $n$ such that for every integer $m$ with $0 0 ∧ ∀ m ∈ Set.Ioo (0 : ℤ) 1993, ∃ k : ℤ, (m / 1993 < (k : ℝ) / n) ∧ ((k : ℝ) / n < (m + 1) / 1994))) -: nallmexk putnam_1993_b1_solution ∧ (∀ n : ℕ, nallmexk n → n ≥ putnam_1993_b1_solution) := -sorry +theorem putnam_1993_b1 : + IsLeast + {n : ℕ | 0 < n ∧ + ∀ m ∈ Set.Ioo (0 : ℤ) (1993), ∃ k : ℤ, + (m / 1993 < (k : ℝ) / n) ∧ ((k : ℝ) / n < (m + 1) / 1994) } + putnam_1993_b1_solution := + sorry diff --git a/lean4/src/putnam_1993_b3.lean b/lean4/src/putnam_1993_b3.lean index fba4668..ba9cf38 100644 --- a/lean4/src/putnam_1993_b3.lean +++ b/lean4/src/putnam_1993_b3.lean @@ -5,8 +5,10 @@ abbrev putnam_1993_b3_solution : ℚ × ℚ := sorry /-- Two real numbers $x$ and $y$ are chosen at random in the interval $(0,1)$ with respect to the uniform distribution. What is the probability that the closest integer to $x/y$ is even? Express the answer in the form $r+s\pi$, where $r$ and $s$ are rational numbers. -/ -theorem putnam_1993_b3 -(S : Set (Fin 2 → ℝ)) -(hS : S = {p : Fin 2 → ℝ | 0 < p ∧ p < 1 ∧ Even (round (p 0 / p 1))}) -: let (r, s) := putnam_1993_b3_solution; (MeasureTheory.volume S).toReal / 1 = r + s * Real.pi := +theorem putnam_1993_b3 : + let (r, s) := putnam_1993_b3_solution; + (MeasureTheory.volume + {p : Fin 2 → ℝ | 0 < p ∧ p < 1 ∧ Even (round (p 0 / p 1))} + ).toReal + = r + s * Real.pi := sorry diff --git a/lean4/src/putnam_1993_b5.lean b/lean4/src/putnam_1993_b5.lean index 72ce9d6..282cb77 100644 --- a/lean4/src/putnam_1993_b5.lean +++ b/lean4/src/putnam_1993_b5.lean @@ -3,8 +3,8 @@ import Mathlib /-- Show there do not exist four points in the Euclidean plane such that the pairwise distances between the points are all odd integers. -/ -theorem putnam_1993_b5 -(pdists : (Fin 4 → (EuclideanSpace ℝ (Fin 2))) → Prop) -(hpdists: ∀ p : Fin 4 → (EuclideanSpace ℝ (Fin 2)), pdists p = ∀ i j : Fin 4, i ≠ j → (dist (p i) (p j) = round (dist (p i) (p j)) ∧ Odd (round (dist (p i) (p j))))) -: ¬∃ p : Fin 4 → (EuclideanSpace ℝ (Fin 2)), pdists p := -sorry +theorem putnam_1993_b5: + ¬∃ p : Fin 4 → (EuclideanSpace ℝ (Fin 2)), + ∀ i j, i ≠ j → + (∃ n : ℤ, dist (p i) (p j) = n ∧ Odd n) := + sorry diff --git a/lean4/src/putnam_1994_a4.lean b/lean4/src/putnam_1994_a4.lean index aa58a31..537a0e4 100644 --- a/lean4/src/putnam_1994_a4.lean +++ b/lean4/src/putnam_1994_a4.lean @@ -6,7 +6,11 @@ open Filter Topology Let $A$ and $B$ be $2 \times 2$ matrices with integer entries such that $A$, $A+B$, $A+2B$, $A+3B$, and $A+4B$ are all invertible matrices whose inverses have integer entries. Show that $A+5B$ is invertible and that its inverse has integer entries. -/ theorem putnam_1994_a4 -(A B : Matrix (Fin 2) (Fin 2) ℤ) -(ABinv : Nonempty (Invertible A) ∧ Nonempty (Invertible (A + B)) ∧ Nonempty (Invertible (A + 2 * B)) ∧ Nonempty (Invertible (A + 3 * B)) ∧ Nonempty (Invertible (A + 4 * B))) -: Nonempty (Invertible (A + 5 * B)) := + (A B : Matrix (Fin 2) (Fin 2) ℤ) + (ABinv : Nonempty (Invertible A) ∧ + Nonempty (Invertible (A + B)) ∧ + Nonempty (Invertible (A + 2 * B)) ∧ + Nonempty (Invertible (A + 3 * B)) ∧ + Nonempty (Invertible (A + 4 * B))) + : Nonempty (Invertible (A + 5 * B)) := sorry diff --git a/lean4/src/putnam_1994_a5.lean b/lean4/src/putnam_1994_a5.lean index 5d6bfde..3d7170c 100644 --- a/lean4/src/putnam_1994_a5.lean +++ b/lean4/src/putnam_1994_a5.lean @@ -6,10 +6,10 @@ open Filter Topology Let $(r_n)_{n \geq 0}$ be a sequence of positive real numbers such that $\lim_{n \to \infty} r_n=0$. Let $S$ be the set of numbers representable as a sum $r_{i_1}+r_{i_2}+\cdots+r_{i_{1994}}$, with $i_1 0) -(rlim : Tendsto r atTop (𝓝 0)) -(hS : S = {x : ℝ | ∃ i : Fin 1994 → ℕ, (∀ j k : Fin 1994, j < k → i j < i k) ∧ (x = ∑ j : Fin 1994, r (i j))}) -: ∀ a b : ℝ, a < b → (∃ c d : ℝ, a ≤ c ∧ c < d ∧ d ≤ b ∧ (Set.Ioo c d) ∩ S = ∅) := + (r : ℕ → ℝ) + (S : Set ℝ) + (rpos : ∀ n, r n > 0) + (rlim : Tendsto r atTop (𝓝 0)) + (hS : S = {x | ∃ i : Fin 1994 → ℕ, (∀ j k : Fin 1994, j < k → i j < i k) ∧ (x = ∑ j : Fin 1994, r (i j))}) : + ∀ a b : ℝ, a < b → (∃ c d : ℝ, a ≤ c ∧ c < d ∧ d ≤ b ∧ (Set.Ioo c d) ∩ S = ∅) := sorry diff --git a/lean4/src/putnam_1994_b1.lean b/lean4/src/putnam_1994_b1.lean index 48e64fd..08968e9 100644 --- a/lean4/src/putnam_1994_b1.lean +++ b/lean4/src/putnam_1994_b1.lean @@ -8,8 +8,6 @@ abbrev putnam_1994_b1_solution : Set ℤ := sorry Find all positive integers $n$ that are within $250$ of exactly $15$ perfect squares. -/ theorem putnam_1994_b1 -(n : ℤ) -(nwithin : Prop) -(hnwithin : nwithin = ({m : ℕ | |n - m ^ 2| ≤ 250}.encard = 15)) -: (n > 0 ∧ nwithin) ↔ n ∈ putnam_1994_b1_solution := +(n : ℤ) (hn : 0 < n) +: {m : ℕ | |n - m ^ 2| ≤ 250}.encard = 15 ↔ n ∈ putnam_1994_b1_solution := sorry diff --git a/lean4/src/putnam_1994_b2.lean b/lean4/src/putnam_1994_b2.lean index 9aad8c0..3937555 100644 --- a/lean4/src/putnam_1994_b2.lean +++ b/lean4/src/putnam_1994_b2.lean @@ -8,8 +8,8 @@ abbrev putnam_1994_b2_solution : Set ℝ := sorry For which real numbers $c$ is there a straight line that intersects the curve $x^4+9x^3+cx^2+9x+4$ in four distinct points? -/ theorem putnam_1994_b2 -(c : ℝ) -(inter : Prop) -(hinter : inter = ∃ m b : ℝ, {x : ℝ | m * x + b = x ^ 4 + 9 * x ^ 3 + c * x ^ 2 + 9 * x + 4}.encard = 4) -: inter ↔ c ∈ putnam_1994_b2_solution := + (c : ℝ) : + (∃ m b : ℝ, + {x : ℝ | m * x + b = x ^ 4 + 9 * x ^ 3 + c * x ^ 2 + 9 * x + 4}.encard = 4) + ↔ c ∈ putnam_1994_b2_solution := sorry diff --git a/lean4/src/putnam_1994_b4.lean b/lean4/src/putnam_1994_b4.lean index 5f1f57c..7508f5e 100644 --- a/lean4/src/putnam_1994_b4.lean +++ b/lean4/src/putnam_1994_b4.lean @@ -9,7 +9,7 @@ theorem putnam_1994_b4 (matgcd : Matrix (Fin 2) (Fin 2) ℤ → ℤ) (A : Matrix (Fin 2) (Fin 2) ℤ) (d : ℕ → ℤ) -(hmatgcd : ∀ M : Matrix (Fin 2) (Fin 2) ℤ, matgcd M = Int.gcd (Int.gcd (Int.gcd (M 0 0) (M 0 1)) (M 1 0)) (M 1 1)) +(hmatgcd : ∀ M, matgcd M = Int.gcd (Int.gcd (Int.gcd (M 0 0) (M 0 1)) (M 1 0)) (M 1 1)) (hA : A 0 0 = 3 ∧ A 0 1 = 2 ∧ A 1 0 = 4 ∧ A 1 1 = 3) (hd : ∀ n ≥ 1, d n = matgcd (A ^ n - 1)) : Tendsto d atTop atTop := diff --git a/lean4/src/putnam_1994_b5.lean b/lean4/src/putnam_1994_b5.lean index 626b156..d3435d0 100644 --- a/lean4/src/putnam_1994_b5.lean +++ b/lean4/src/putnam_1994_b5.lean @@ -6,9 +6,11 @@ open Filter Topology For any real number $\alpha$, define the function $f_\alpha(x)=\lfloor \alpha x \rfloor$. Let $n$ be a positive integer. Show that there exists an $\alpha$ such that for $1 \leq k \leq n$, $f_\alpha^k(n^2)=n^2-k=f_{\alpha^k}(n^2)$. -/ theorem putnam_1994_b5 -(f : ℝ → ℤ → ℤ) -(n : ℕ) -(hf : ∀ (α : ℝ) (x : ℤ), f α x = Int.floor (α * x)) -(npos : n > 0) -: ∃ α : ℝ, ∀ k ∈ Set.Icc 1 n, ((f α)^[k] (n ^ 2) = n ^ 2 - k) ∧ (f (α ^ k) (n ^ 2) = n ^ 2 - k) := + (n : ℕ) (hn : n > 0) + (f : ℝ → ℤ → ℤ) + (hf : ∀ (α : ℝ) (x : ℤ), f α x = Int.floor (α * x)) : + ∃ α : ℝ, + ∀ k ∈ Set.Icc 1 n, + ((f α)^[k] (n ^ 2) = n ^ 2 - k) ∧ + (f (α ^ k) (n ^ 2) = n ^ 2 - k) := sorry diff --git a/lean4/src/putnam_1995_a4.lean b/lean4/src/putnam_1995_a4.lean index f5794c8..be8ca7d 100644 --- a/lean4/src/putnam_1995_a4.lean +++ b/lean4/src/putnam_1995_a4.lean @@ -10,5 +10,5 @@ theorem putnam_1995_a4 (hn : n > 0) (necklace : Fin n → ℤ) (hnecklacesum : ∑ i : Fin n, necklace i = n - 1) -: ∃ cut : Fin n, ∀ k : Fin n, ∑ i : {j : Fin n | j.1 ≤ k}, necklace (cut + i) ≤ k := +: ∃ cut, ∀ k, ∑ i : {j : Fin n | j.1 ≤ k}, necklace (cut + i) ≤ k := sorry diff --git a/lean4/src/putnam_1995_a5.lean b/lean4/src/putnam_1995_a5.lean index c995369..5edb9a7 100644 --- a/lean4/src/putnam_1995_a5.lean +++ b/lean4/src/putnam_1995_a5.lean @@ -8,13 +8,13 @@ abbrev putnam_1995_a5_solution : Prop := sorry Let $x_{1},x_{2},\dots,x_{n}$ be differentiable (real-valued) functions of a single variable $f$ which satisfy \begin{align*} \frac{dx_{1}}{dt} &= a_{11}x_{1} + a_{12}x_{2} + \cdots + a_{1n}x_{n} \ \frac{dx_{2}}{dt} &= a_{21}x_{1} + a_{22}x_{2} + \cdots + a_{2n}x_{n} \ \vdots && \vdots \ \frac{dx_{n}}{dt} &= a_{n1}x_{1} + a_{n2}x_{2} + \cdots + a_{nn}x_{n} \end{align*} for some constants $a_{ij}>0$. Suppose that for all $i$, $x_{i}(t) \to 0$ as $t \to \infty$. Are the functions $x_{1},x_{2},\dots,x_{n}$ necessarily linearly dependent? -/ theorem putnam_1995_a5 -(hdiffx : (n : ℕ) → (Fin n → (ℝ → ℝ)) → Prop) -(ha : (n : ℕ) → (Fin n → Fin n → ℝ) → Prop) -(hcomb : (n : ℕ) → (Fin n → (ℝ → ℝ)) → (Fin n → Fin n → ℝ) → Prop) -(hxlim : (n : ℕ) → (Fin n → (ℝ → ℝ)) → Prop) -(hdiffx_def : hdiffx = (fun (n : ℕ) (x : Fin n → (ℝ → ℝ)) => ∀ i : Fin n, Differentiable ℝ (x i))) -(ha_def : ha = (fun (n : ℕ) (a : Fin n → Fin n → ℝ) => ∀ i j : Fin n, a i j > 0)) -(hcomb_def : hcomb = (fun (n : ℕ) (x : Fin n → (ℝ → ℝ)) (a : Fin n → Fin n → ℝ) => ∀ t : ℝ, ∀ i : Fin n, (deriv (x i)) t = ∑ j : Fin n, (a i j) * ((x j) t))) -(hxlim_def : hxlim = (fun (n : ℕ) (x : Fin n → (ℝ → ℝ)) => ∀ i : Fin n, Tendsto (x i) atTop (𝓝 0))) -: putnam_1995_a5_solution ↔ (∀ (n : ℕ) (x : Fin n → (ℝ → ℝ)) (a : Fin n → Fin n → ℝ), (n > 0 ∧ hdiffx n x ∧ ha n a ∧ hcomb n x a ∧ hxlim n x) → ¬(∀ b : Fin n → ℝ, (∀ t : ℝ, ∑ i : Fin n, (b i) * ((x i) t) = 0) → (∀ i : Fin n, b i = 0))) := -sorry + (n : ℕ) (hn : n > 0) : + putnam_1995_a5_solution ↔ + (∀ (x : Fin n → (ℝ → ℝ)) (a : Fin n → Fin n → ℝ), + (∀ i, Differentiable ℝ (x i)) → + (∀ i j, a i j > 0) → + (∀ t i, (deriv (x i)) t = ∑ j : Fin n, (a i j) * ((x j) t)) → + (∀ i, Tendsto (x i) atTop (𝓝 0)) → + ¬(∀ b : Fin n → ℝ, (∀ t : ℝ, ∑ i : Fin n, (b i) * ((x i) t) = 0) → + (∀ i, b i = 0))) := + sorry diff --git a/lean4/src/putnam_1995_a6.lean b/lean4/src/putnam_1995_a6.lean index 96175b8..ebe5838 100644 --- a/lean4/src/putnam_1995_a6.lean +++ b/lean4/src/putnam_1995_a6.lean @@ -7,10 +7,10 @@ Suppose that each of $n$ people writes down the numbers $1,2,3$ in random order -/ theorem putnam_1995_a6 (SM : (n : ℕ) → Set (Matrix (Fin 3) (Fin n) ℤ)) -(hSM : SM = (fun n : ℕ => {M : Matrix (Fin 3) (Fin n) ℤ | ∀ j : Fin n, {M i j | i : Fin 3} = {1, 2, 3}})) (Srowsums : (n : ℕ) → Matrix (Fin 3) (Fin n) ℤ → Set ℤ) -(hSrowsums : Srowsums = (fun (n : ℕ) (M : Matrix (Fin 3) (Fin n) ℤ) => {∑ j : Fin n, M i j | i : Fin 3})) (sumsprog : (Set ℤ) → Prop) -(hsumsprog : sumsprog = (fun sums : Set ℤ => sums.encard = 3 ∧ (∃ a b c : ℤ, {a, b, c} = sums ∧ b = a + 1 ∧ c = a + 2))) +(hSM : ∀ n, SM n = {M : Matrix (Fin 3) (Fin n) ℤ | ∀ j : Fin n, {M i j | i : Fin 3} = {1, 2, 3}}) +(hSrowsums : ∀ n M, Srowsums n M = {∑ j : Fin n, M i j | i : Fin 3}) +(hsumsprog : ∀ sums, sumsprog sums ↔ sums.encard = 3 ∧ (∃ a b c : ℤ, {a, b, c} = sums ∧ b = a + 1 ∧ c = a + 2)) : ∃ n ≥ 1995, {M : (SM n) | sumsprog (Srowsums n M)}.encard ≥ 4 * {M : (SM n) | (Srowsums n M).encard = 1}.encard := sorry diff --git a/lean4/src/putnam_1995_b1.lean b/lean4/src/putnam_1995_b1.lean index f377d2e..ca3161d 100644 --- a/lean4/src/putnam_1995_b1.lean +++ b/lean4/src/putnam_1995_b1.lean @@ -6,7 +6,11 @@ open Filter Topology Real Nat For a partition $\pi$ of $\{1, 2, 3, 4, 5, 6, 7, 8, 9\}$, let $\pi(x)$ be the number of elements in the part containing $x$. Prove that for any two partitions $\pi$ and $\pi'$, there are two distinct numbers $x$ and $y$ in $\{1, 2, 3, 4, 5, 6, 7, 8, 9\}$ such that $\pi(x) = \pi(y)$ and $\pi'(x) = \pi'(y)$. [A {\em partition} of a set $S$ is a collection of disjoint subsets (parts) whose union is $S$.] -/ theorem putnam_1995_b1 -(part_ct : Finpartition (Finset.range 9) → (Finset.range 9) → ℕ) -(hp : ∀ partition k, part_ct partition k = (Exists.choose (Finpartition.exists_mem partition k.2)).card) -: ∀ Pt1 Pt2 : Finpartition (Finset.range 9), ∃ x y : Finset.range 9, x ≠ y ∧ part_ct Pt1 x = part_ct Pt1 y ∧ part_ct Pt2 x = part_ct Pt2 y := + (Pi : Finpartition (Finset.range 9) → (Finset.range 9) → ℕ) + (Pi_def : ∀ partition k, Pi partition k = (Exists.choose (Finpartition.exists_mem partition k.2)).card) : + ∀ Pt1 Pt2, + ∃ x y : Finset.range 9, + x ≠ y ∧ + Pi Pt1 x = Pi Pt1 y ∧ + Pi Pt2 x = Pi Pt2 y := sorry diff --git a/lean4/src/putnam_1995_b3.lean b/lean4/src/putnam_1995_b3.lean index 9d8e5d2..fb8f302 100644 --- a/lean4/src/putnam_1995_b3.lean +++ b/lean4/src/putnam_1995_b3.lean @@ -12,7 +12,7 @@ theorem putnam_1995_b3 (n : ℕ) (hn : n > 0) (digits_set : Set (ℕ → ℕ)) -(hdigits_set : digits_set = {f : ℕ → ℕ | f 0 ≠ 0 ∧ (∀ i : Fin (n ^ 2), f i ≤ 9) ∧ (∀ i ≥ n ^ 2, f i = 0)}) +(hdigits_set : digits_set = {f | f 0 ≠ 0 ∧ (∀ i : Fin (n ^ 2), f i ≤ 9) ∧ (∀ i ≥ n ^ 2, f i = 0)}) (digits_to_matrix : (ℕ → ℕ) → Matrix (Fin n) (Fin n) ℤ) (hdigits_to_matrix : digits_to_matrix = fun f => (fun i j => f (i.1 * n + j.1))) : ∑' f : digits_set, (digits_to_matrix f).det = putnam_1995_b3_solution n := diff --git a/lean4/src/putnam_1995_b6.lean b/lean4/src/putnam_1995_b6.lean index 7e533c9..1017e5f 100644 --- a/lean4/src/putnam_1995_b6.lean +++ b/lean4/src/putnam_1995_b6.lean @@ -6,7 +6,10 @@ open Filter Topology Real Nat For a positive real number $\alpha$, define \[ S(\alpha) = \{ \lfloor n\alpha \rfloor : n = 1,2,3,\dots \}. \] Prove that $\{1,2,3,\dots\}$ cannot be expressed as the disjoint union of three sets $S(\alpha), S(\beta)$ and $S(\gamma)$. [As usual, $\lfloor x \rfloor$ is the greatest integer $\leq x$.] -/ theorem putnam_1995_b6 -(S : ℝ → Set ℕ) -(hS : S = fun (α : ℝ) => {x : ℕ | ∃ n : ℕ, n ≥ 1 ∧ x = floor (n * α)}) -: ¬ ∃ α β γ : ℝ, α > 0 ∧ β > 0 ∧ γ > 0 ∧ (S α) ∩ (S β) = ∅ ∧ (S β) ∩ (S γ) = ∅ ∧ (S α) ∩ (S γ) = ∅ ∧ Set.Ici 1 = (S α) ∪ (S β) ∪ (S γ) := + (S : ℝ → Set ℕ) + (hS : S = fun (α : ℝ) => {x : ℕ | ∃ n : ℕ, n ≥ 1 ∧ x = floor (n * α)}) : + ¬ ∃ α β γ, + α > 0 ∧ β > 0 ∧ γ > 0 ∧ + (S α) ∩ (S β) = ∅ ∧ (S β) ∩ (S γ) = ∅ ∧ (S α) ∩ (S γ) = ∅ ∧ + Set.Ici 1 = (S α) ∪ (S β) ∪ (S γ) := sorry diff --git a/lean4/src/putnam_1996_b3.lean b/lean4/src/putnam_1996_b3.lean index fc06769..310b120 100644 --- a/lean4/src/putnam_1996_b3.lean +++ b/lean4/src/putnam_1996_b3.lean @@ -9,11 +9,10 @@ abbrev putnam_1996_b3_solution : ℕ → ℕ := sorry Given that $\{x_1,x_2,\ldots,x_n\}=\{1,2,\ldots,n\}$, find, with proof, the largest possible value, as a function of $n$ (with $n \geq 2$), of $x_1x_2+x_2x_3+\cdots+x_{n-1}x_n+x_nx_1$. -/ theorem putnam_1996_b3 -(n : ℕ) -(xset : (ℕ → ℤ) → Prop) -(xsum : (ℕ → ℤ) → ℤ) -(nge2 : n ≥ 2) -(hxset : ∀ x : ℕ → ℤ, xset x = (x '' (Finset.range n) = Set.Icc (1 : ℤ) n)) -(hxsum : ∀ x : ℕ → ℤ, xsum x = ∑ i : Fin n, x i * x ((i + 1) % n)) -: (∃ x : ℕ → ℤ, xset x ∧ xsum x = putnam_1996_b3_solution n) ∧ (∀ x : ℕ → ℤ, xset x → xsum x ≤ putnam_1996_b3_solution n) := + (n : ℕ) (hn : n ≥ 2) : + IsGreatest + {k | ∃ x : ℕ → ℤ, + (x '' (Finset.range n) = Set.Icc (1 : ℤ) n) ∧ + ∑ i : Fin n, x i * x ((i + 1) % n) = k} + (putnam_1996_b3_solution n) := sorry diff --git a/lean4/src/putnam_1996_b4.lean b/lean4/src/putnam_1996_b4.lean index e97ceed..6ac76e1 100644 --- a/lean4/src/putnam_1996_b4.lean +++ b/lean4/src/putnam_1996_b4.lean @@ -10,7 +10,7 @@ For any square matrix $A$, we can define $\sin A$ by the usual power series: $\s theorem putnam_1996_b4 (matsin : Matrix (Fin 2) (Fin 2) ℝ → Matrix (Fin 2) (Fin 2) ℝ) (mat1996 : Matrix (Fin 2) (Fin 2) ℝ) -(hmatsin : ∀ A : Matrix (Fin 2) (Fin 2) ℝ, matsin A = ∑' n : ℕ, ((-(1 : ℝ)) ^ n / (2 * n + 1)!) • A ^ (2 * n + 1)) +(hmatsin : ∀ A, matsin A = ∑' n : ℕ, ((-(1 : ℝ)) ^ n / (2 * n + 1)!) • A ^ (2 * n + 1)) (hmat1996 : mat1996 0 0 = 1 ∧ mat1996 0 1 = 1996 ∧ mat1996 1 0 = 0 ∧ mat1996 1 1 = 1) -: (∃ A : Matrix (Fin 2) (Fin 2) ℝ, matsin A = mat1996) ↔ putnam_1996_b4_solution := +: (∃ A, matsin A = mat1996) ↔ putnam_1996_b4_solution := sorry diff --git a/lean4/src/putnam_1997_a6.lean b/lean4/src/putnam_1997_a6.lean index 03dda8c..0eb09e4 100644 --- a/lean4/src/putnam_1997_a6.lean +++ b/lean4/src/putnam_1997_a6.lean @@ -8,15 +8,13 @@ abbrev putnam_1997_a6_solution : ℤ → ℤ → ℝ := sorry For a positive integer $n$ and any real number $c$, define $x_k$ recursively by $x_0=0$, $x_1=1$, and for $k\geq 0$, \[x_{k+2}=\frac{cx_{k+1}-(n-k)x_k}{k+1}.\] Fix $n$ and then take $c$ to be the largest value for which $x_{n+1}=0$. Find $x_k$ in terms of $n$ and $k$, $1\leq k\leq n$. -/ theorem putnam_1997_a6 -(n : ℤ) -(hn : n > 0) -(C : ℝ) -(x : ℝ → (ℤ → ℝ)) -(hx0 : ∀ c : ℝ, x c 0 = 0) -(hx1 : ∀ c : ℝ, x c 1 = 1) -(hxk : ∀ c : ℝ, ∀ k ≥ 0, x c (k + 2) = (c*(x c (k + 1)) - (n - k)*(x c k))/(k + 1)) -(S : Set ℝ) -(hS : S = {c : ℝ | x c (n + 1) = 0}) -(hC : C = sSup S) -: ∀ k : Set.Icc 1 n, x C k = putnam_1997_a6_solution n k := + (n : ℤ) + (hn : n > 0) + (x : ℝ → (ℤ → ℝ)) + (hx0 : ∀ c, x c 0 = 0) + (hx1 : ∀ c, x c 1 = 1) + (hxk : ∀ c, ∀ k ≥ 0, x c (k + 2) = (c*(x c (k + 1)) - (n - k)*(x c k))/(k + 1)) + (S : Set ℝ) + (hS : S = {c : ℝ | x c (n + 1) = 0}) : + ∀ k : Set.Icc 1 n, x (sSup S) k = putnam_1997_a6_solution n k := sorry diff --git a/lean4/src/putnam_1998_a2.lean b/lean4/src/putnam_1998_a2.lean index 2e32cc3..79d3a9d 100644 --- a/lean4/src/putnam_1998_a2.lean +++ b/lean4/src/putnam_1998_a2.lean @@ -5,9 +5,9 @@ Let $s$ be any arc of the unit circle lying entirely in the first quadrant. Let -/ theorem putnam_1998_a2 (quadrant : (EuclideanSpace ℝ (Fin 2)) → Prop) -(hquadrant : quadrant = fun P ↦ P 0 > 0 ∧ P 1 > 0 ∧ dist 0 P = 1) (isarc : (EuclideanSpace ℝ (Fin 2)) → (EuclideanSpace ℝ (Fin 2)) → Prop) -(hisarc : isarc = fun P Q ↦ quadrant P ∧ quadrant Q ∧ P 0 > Q 0) +(hquadrant : ∀ P, quadrant P ↔ P 0 > 0 ∧ P 1 > 0 ∧ dist 0 P = 1) +(hisarc : ∀ P Q, isarc P Q ↔ quadrant P ∧ quadrant Q ∧ P 0 > Q 0) (arc : (EuclideanSpace ℝ (Fin 2)) → (EuclideanSpace ℝ (Fin 2)) → Set (EuclideanSpace ℝ (Fin 2))) (harc : arc = fun P Q ↦ {R : EuclideanSpace ℝ (Fin 2) | quadrant R ∧ P 0 > R 0 ∧ R 0 > Q 0}) (A B : (EuclideanSpace ℝ (Fin 2)) → (EuclideanSpace ℝ (Fin 2)) → ℝ) diff --git a/lean4/src/putnam_1998_a5.lean b/lean4/src/putnam_1998_a5.lean index 228fd21..7afe300 100644 --- a/lean4/src/putnam_1998_a5.lean +++ b/lean4/src/putnam_1998_a5.lean @@ -6,11 +6,13 @@ open Set Function Metric Let $\mathcal F$ be a finite collection of open discs in $\mathbb R^2$ whose union contains a set $E\subseteq \mathbb R^2$. Show that there is a pairwise disjoint subcollection $D_1,\ldots, D_n$ in $\mathcal F$ such that \[E\subseteq \cup_{j=1}^n 3D_j.\] Here, if $D$ is the disc of radius $r$ and center $P$, then $3D$ is the disc of radius $3r$ and center $P$. -/ theorem putnam_1998_a5 -(k : ℕ) -(c : Fin k → (EuclideanSpace ℝ (Fin 2))) -(r : Fin k → ℝ) -(hr : ∀ i : Fin k, r i > 0) -(E : Set (EuclideanSpace ℝ (Fin 2))) -(hE : E ⊆ ⋃ i : Fin k, ball (c i) (r i)) -: ∃ (n : ℕ) (t : Fin n → Fin k), (∀ i j : Fin n, i ≠ j → (ball (c (t i)) (r (t i)) ∩ ball (c (t j)) (r (t j)) = ∅)) ∧ E ⊆ ⋃ i : Fin n, ball (c (t i)) (3 * (r (t i))) := + (k : ℕ) + (c : Fin k → (EuclideanSpace ℝ (Fin 2))) + (r : Fin k → ℝ) + (hr : ∀ i, r i > 0) + (E : Set (EuclideanSpace ℝ (Fin 2))) + (hE : E ⊆ ⋃ i, ball (c i) (r i)) : + ∃ (n : ℕ) (t : Fin n → Fin k), + (∀ i j, i ≠ j → (ball (c (t i)) (r (t i)) ∩ ball (c (t j)) (r (t j)) = ∅)) ∧ + E ⊆ ⋃ i : Fin n, ball (c (t i)) (3 * (r (t i))) := sorry diff --git a/lean4/src/putnam_1998_b4.lean b/lean4/src/putnam_1998_b4.lean index 582f1bd..c493e2e 100644 --- a/lean4/src/putnam_1998_b4.lean +++ b/lean4/src/putnam_1998_b4.lean @@ -8,7 +8,9 @@ abbrev putnam_1998_b4_solution : Set (ℕ × ℕ) := sorry Find necessary and sufficient conditions on positive integers $m$ and $n$ so that \[\sum_{i=0}^{mn-1} (-1)^{\lfloor i/m \rfloor +\lfloor i/n\rfloor}=0.\] -/ theorem putnam_1998_b4 -(quantity : ℕ → ℕ → ℤ) -(hquantity : quantity = fun n m => ∑ i in Finset.range (m * n), (-1)^(i/m + i/n)) -: ∀ n m : ℕ, n > 0 ∧ m > 0 → ((quantity n m) = 0 ↔ ⟨n, m⟩ ∈ putnam_1998_b4_solution) := + (quantity : ℕ → ℕ → ℤ) + (hquantity : quantity = fun n m => ∑ i in Finset.range (m * n), (-1)^(i/m + i/n)) + (n m : ℕ) + (hnm : n > 0 ∧ m > 0) : + quantity n m = 0 ↔ ⟨n, m⟩ ∈ putnam_1998_b4_solution := sorry diff --git a/lean4/src/putnam_1999_a2.lean b/lean4/src/putnam_1999_a2.lean index 87e99c2..ef6b322 100644 --- a/lean4/src/putnam_1999_a2.lean +++ b/lean4/src/putnam_1999_a2.lean @@ -5,7 +5,8 @@ Let $p(x)$ be a polynomial that is nonnegative for all real $x$. Prove that for -/ theorem putnam_1999_a2 (p : Polynomial ℝ) -(hp : ∀ x : ℝ, p.eval x ≥ 0) -: ∃ k : ℕ, ∃ f : Fin k → Polynomial ℝ, k > 0 ∧ -∀ x : ℝ, p.eval x = ∑ j : Fin k, ((f j).eval x) ^ 2 := +(hp : ∀ x, p.eval x ≥ 0) +: ∃ᵉ (k) (f : Fin k → Polynomial ℝ), + k > 0 ∧ + ∀ x : ℝ, p.eval x = ∑ j : Fin k, ((f j).eval x) ^ 2 := sorry diff --git a/lean4/src/putnam_1999_b2.lean b/lean4/src/putnam_1999_b2.lean index 6616a8d..232a09d 100644 --- a/lean4/src/putnam_1999_b2.lean +++ b/lean4/src/putnam_1999_b2.lean @@ -6,8 +6,11 @@ open Filter Topology Metric Let $P(x)$ be a polynomial of degree $n$ such that $P(x)=Q(x)P''(x)$, where $Q(x)$ is a quadratic polynomial and $P''(x)$ is the second derivative of $P(x)$. Show that if $P(x)$ has at least two distinct roots then it must have $n$ distinct roots. -/ theorem putnam_1999_b2 -(P Q : Polynomial ℂ) -(hQ : Q.natDegree = 2) -(hP : ∀ x : ℂ, P.eval x = (Q.eval x) * (Polynomial.derivative^[2] P).eval x) -: (∃ x1 x2 : ℂ, x1 ≠ x2 ∧ P.eval x1 = 0 ∧ P.eval x2 = 0) → (∃ f : Fin (P.natDegree) → ℂ, (∀ i j : Fin (P.natDegree), i ≠ j → f i ≠ f j) ∧ (∀ i : Fin (P.natDegree), P.eval (f i) = 0)) := + (P Q : Polynomial ℂ) + (hQ : Q.natDegree = 2) + (hP : ∀ x : ℂ, P.eval x = (Q.eval x) * (Polynomial.derivative^[2] P).eval x) : + (∃ x1 x2, x1 ≠ x2 ∧ P.eval x1 = 0 ∧ P.eval x2 = 0) → + (∃ f : Fin (P.natDegree) → ℂ, + (∀ i j, i ≠ j → f i ≠ f j) ∧ + (∀ i, P.eval (f i) = 0)) := sorry diff --git a/lean4/src/putnam_1999_b3.lean b/lean4/src/putnam_1999_b3.lean index 3a51656..e12b172 100644 --- a/lean4/src/putnam_1999_b3.lean +++ b/lean4/src/putnam_1999_b3.lean @@ -9,7 +9,7 @@ Let $A=\{(x,y):0\leq x,y<1\}$. For $(x,y)\in A$, let \[S(x,y) = \sum_{\frac{1}{ -/ theorem putnam_1999_b3 (A : Set (ℝ × ℝ)) -(hA : A = {xy : ℝ × ℝ | 0 ≤ xy.1 ∧ xy.1 < 1 ∧ 0 ≤ xy.2 ∧ xy.2 < 1}) +(hA : A = {xy | 0 ≤ xy.1 ∧ xy.1 < 1 ∧ 0 ≤ xy.2 ∧ xy.2 < 1}) (S : ℝ → ℝ → ℝ) (hS : S = fun x y => ∑' m : ℕ, ∑' n : ℕ, if (m > 0 ∧ n > 0 ∧ 1/2 ≤ m/n ∧ m/n ≤ 2) then x^m * y^n else 0) : Tendsto (fun xy : (ℝ × ℝ) => (1 - xy.1 * xy.2^2) * (1 - xy.1^2 * xy.2) * (S xy.1 xy.2)) (𝓝[A] ⟨1,1⟩) (𝓝 putnam_1999_b3_solution) := diff --git a/lean4/src/putnam_1999_b5.lean b/lean4/src/putnam_1999_b5.lean index 1f6668f..81a1c0e 100644 --- a/lean4/src/putnam_1999_b5.lean +++ b/lean4/src/putnam_1999_b5.lean @@ -8,11 +8,11 @@ noncomputable abbrev putnam_1999_b5_solution : ℕ → ℝ := sorry For an integer $n\geq 3$, let $\theta=2\pi/n$. Evaluate the determinant of the $n\times n$ matrix $I+A$, where $I$ is the $n\times n$ identity matrix and $A=(a_{jk})$ has entries $a_{jk}=\cos(j\theta+k\theta)$ for all $j,k$. -/ theorem putnam_1999_b5 -(n : ℕ) -(hn : n ≥ 3) -(theta : ℝ) -(htheta : theta = 2 * Real.pi / n) -(A : Matrix (Fin n) (Fin n) ℝ) -(hA : A = fun j k => Real.cos ((j.1 + 1) * theta + (k.1 + 1) * theta)) -: ((1 : Matrix (Fin n) (Fin n) ℝ) + A).det = putnam_1999_b5_solution n := + (n : ℕ) + (hn : n ≥ 3) + (theta : ℝ) + (htheta : theta = 2 * Real.pi / n) + (A : Matrix (Fin n) (Fin n) ℝ) + (hA : A = fun j k => Real.cos ((j.1 + 1) * theta + (k.1 + 1) * theta)) : + (1 + A).det = putnam_1999_b5_solution n := sorry diff --git a/lean4/src/putnam_1999_b6.lean b/lean4/src/putnam_1999_b6.lean index 4316e53..c3f6c73 100644 --- a/lean4/src/putnam_1999_b6.lean +++ b/lean4/src/putnam_1999_b6.lean @@ -6,8 +6,8 @@ open Filter Topology Metric Let $S$ be a finite set of integers, each greater than 1. Suppose that for each integer $n$ there is some $s\in S$ such that $\gcd(s,n)=1$ or $\gcd(s,n)=s$. Show that there exist $s,t\in S$ such that $\gcd(s,t)$ is prime. -/ theorem putnam_1999_b6 -(S : Finset ℤ) -(hSgt : ∀ s : ℤ, s ∈ S → s > 1) -(hSgcd : ∀ n : ℤ, ∃ s : S, Int.gcd s n = 1 ∨ Int.gcd s n = (s : ℤ)) -: ∃ s t : S, Prime (Int.gcd s t) := + (S : Finset ℤ) + (hSgt : ∀ s, s ∈ S → s > 1) + (hSgcd : ∀ n, ∃ s ∈ S, Int.gcd s n = 1 ∨ Int.gcd s n = s) : + ∃ᵉ (s ∈ S) (t ∈ S), Prime (Int.gcd s t) := sorry diff --git a/lean4/src/putnam_2000_a1.lean b/lean4/src/putnam_2000_a1.lean index 54262a5..47e89f0 100644 --- a/lean4/src/putnam_2000_a1.lean +++ b/lean4/src/putnam_2000_a1.lean @@ -10,5 +10,10 @@ Let $A$ be a positive real number. What are the possible values of $\sum_{j=0}^\ theorem putnam_2000_a1 (A : ℝ) (Apos : A > 0) -: ({S : ℝ | ∃ x : ℕ → ℝ, (∀ j : ℕ, x j > 0) ∧ (∑' j : ℕ, x j) = A ∧ (∑' j : ℕ, (x j) ^ 2) = S} = putnam_2000_a1_solution A) := +: {S : ℝ | + ∃ x : ℕ → ℝ, + (∀ j : ℕ, x j > 0) ∧ + (∑' j : ℕ, x j) = A ∧ + (∑' j : ℕ, (x j) ^ 2) = S} + = putnam_2000_a1_solution A := sorry diff --git a/lean4/src/putnam_2000_a2.lean b/lean4/src/putnam_2000_a2.lean index f22022e..1dd347c 100644 --- a/lean4/src/putnam_2000_a2.lean +++ b/lean4/src/putnam_2000_a2.lean @@ -5,6 +5,12 @@ open Topology Filter /-- Prove that there exist infinitely many integers $n$ such that $n,n+1,n+2$ are each the sum of the squares of two integers. -/ -theorem putnam_2000_a2 -: ∀ n : ℕ, ∃ N : ℤ, ∃ i : Fin 6 → ℕ, N > n ∧ N = (i 0)^2 + (i 1)^2 ∧ N + 1 = (i 2)^2 + (i 3)^2 ∧ N + 2 = (i 4)^2 + (i 5)^2 := -sorry +theorem putnam_2000_a2 : + ∀ n : ℕ, + ∃ N : ℤ, + ∃ i : Fin 6 → ℕ, + N > n ∧ + N = (i 0)^2 + (i 1)^2 ∧ + N + 1 = (i 2)^2 + (i 3)^2 ∧ + N + 2 = (i 4)^2 + (i 5)^2 := + sorry diff --git a/lean4/src/putnam_2000_a5.lean b/lean4/src/putnam_2000_a5.lean index 70b1277..f8ecf2b 100644 --- a/lean4/src/putnam_2000_a5.lean +++ b/lean4/src/putnam_2000_a5.lean @@ -6,12 +6,12 @@ open Topology Filter Three distinct points with integer coordinates lie in the plane on a circle of radius $r>0$. Show that two of these points are separated by a distance of at least $r^{1/3}$. -/ theorem putnam_2000_a5 -(r : ℝ) -(z : EuclideanSpace ℝ (Fin 2)) -(p : Fin 3 → (EuclideanSpace ℝ (Fin 2))) -(rpos : r > 0) -(pdiff : ∀ n m : Fin 3, (n ≠ m) → (p n ≠ p m)) -(pint : ∀ (n : Fin 3) (i : Fin 2), p n i = round (p n i)) -(pcirc : ∀ n : Fin 3, p n ∈ Metric.sphere z r) -: ∃ n m : Fin 3, (n ≠ m) ∧ (dist (p n) (p m) ≥ r ^ ((1 : ℝ) / 3)) := + (r : ℝ) + (z : EuclideanSpace ℝ (Fin 2)) + (p : Fin 3 → (EuclideanSpace ℝ (Fin 2))) + (rpos : r > 0) + (pdiff : ∀ n m, (n ≠ m) → (p n ≠ p m)) + (pint : ∀ n i, p n i = round (p n i)) + (pcirc : ∀ n, p n ∈ Metric.sphere z r) : + ∃ n m, (n ≠ m) ∧ (dist (p n) (p m) ≥ r ^ ((1 : ℝ) / 3)) := sorry diff --git a/lean4/src/putnam_2000_b3.lean b/lean4/src/putnam_2000_b3.lean index eee4241..6c9a40d 100644 --- a/lean4/src/putnam_2000_b3.lean +++ b/lean4/src/putnam_2000_b3.lean @@ -9,15 +9,14 @@ N_0\leq N_1\leq N_2\leq \cdots \mbox{ and } \lim_{k\to\infty} N_k = 2N. \] -/ theorem putnam_2000_b3 -(N : ℕ) -(Npos : N > 0) -(a : Fin (N + 1) → ℝ) -(haN : a N ≠ 0) -(f : ℝ → ℝ) -(hf : f = fun t ↦ ∑ j : Icc 1 N, a j * Real.sin (2 * Real.pi * j * t)) -(mult : (ℝ → ℝ) → ℝ → ℕ) -(hmult : ∀ g : ℝ → ℝ, ∀ t : ℝ, (∃ c : ℕ, iteratedDeriv c g t ≠ 0) → (iteratedDeriv (mult g t) g t ≠ 0 ∧ ∀ k < (mult g t), iteratedDeriv k g t = 0)) -(M : ℕ → ℕ) -(hM : M = fun k ↦ ∑' t : Ico (0 : ℝ) 1, mult (iteratedDeriv k f) t) -: ((∀ i j : ℕ, i ≤ j → M i ≤ M j) ∧ Tendsto M atTop (𝓝 (2 * N))) := + (N : ℕ) (hN : N > 0) + (a : Fin (N + 1) → ℝ) + (f : ℝ → ℝ) + (mult : (ℝ → ℝ) → ℝ → ℕ) + (M : ℕ → ℕ) + (haN : a N ≠ 0) + (hf : ∀ t, f t = ∑ j : Icc 1 N, a j * Real.sin (2 * Real.pi * j * t)) + (hmult : ∀ g : ℝ → ℝ, ∀ t : ℝ, (∃ c : ℕ, iteratedDeriv c g t ≠ 0) → (iteratedDeriv (mult g t) g t ≠ 0 ∧ ∀ k < (mult g t), iteratedDeriv k g t = 0)) + (hM : ∀ k, M k = ∑' t : Ico (0 : ℝ) 1, mult (iteratedDeriv k f) t) : + ((∀ i j : ℕ, i ≤ j → M i ≤ M j) ∧ Tendsto M atTop (𝓝 (2 * N))) := sorry diff --git a/lean4/src/putnam_2000_b4.lean b/lean4/src/putnam_2000_b4.lean index c34fed0..2d94b9f 100644 --- a/lean4/src/putnam_2000_b4.lean +++ b/lean4/src/putnam_2000_b4.lean @@ -8,6 +8,6 @@ Let $f(x)$ be a continuous function such that $f(2x^2-1)=2xf(x)$ for all $x$. S theorem putnam_2000_b4 (f : ℝ → ℝ) (hfcont : Continuous f) -(hf : ∀ x : ℝ, f (2 * x ^ 2 - 1) = 2 * x * f x) -: ∀ x : ℝ, x ∈ Icc (-1) 1 → f x = 0 := +(hf : ∀ x, f (2 * x ^ 2 - 1) = 2 * x * f x) +: ∀ x, x ∈ Icc (-1) 1 → f x = 0 := sorry diff --git a/lean4/src/putnam_2000_b5.lean b/lean4/src/putnam_2000_b5.lean index c4c7225..bc644bb 100644 --- a/lean4/src/putnam_2000_b5.lean +++ b/lean4/src/putnam_2000_b5.lean @@ -7,8 +7,8 @@ Let $S_0$ be a finite set of positive integers. We define finite sets $S_1,S_2,\ -/ theorem putnam_2000_b5 (S : ℕ → Set ℤ) -(hSfin : ∀ n : ℕ, Set.Finite (S n)) -(hSpos : ∀ n : ℕ, ∀ s ∈ S n, s > 0) -(hSdef : ∀ n : ℕ, ∀ a : ℤ, a ∈ S (n + 1) ↔ Xor' (a - 1 ∈ S n) (a ∈ S n)) -: (∀ n : ℕ, ∃ N ≥ n, S N = S 0 ∪ {M : ℤ | M - N ∈ S 0}) := +(hSfin : ∀ n, Set.Finite (S n)) +(hSpos : ∀ n, ∀ s ∈ S n, s > 0) +(hSdef : ∀ n, ∀ a, a ∈ S (n + 1) ↔ Xor' (a - 1 ∈ S n) (a ∈ S n)) +: (∀ n, ∃ N ≥ n, S N = S 0 ∪ {M : ℤ | M - N ∈ S 0}) := sorry diff --git a/lean4/src/putnam_2001_a3.lean b/lean4/src/putnam_2001_a3.lean index f8b4900..fc7e7a7 100644 --- a/lean4/src/putnam_2001_a3.lean +++ b/lean4/src/putnam_2001_a3.lean @@ -12,6 +12,6 @@ the product of two non-constant polynomials with integer coefficients? theorem putnam_2001_a3 (P : ℤ → Polynomial ℤ) (hP : P = fun m : ℤ => (Polynomial.X)^4 - (Polynomial.C (2*m + 4))*(Polynomial.X)^2 + Polynomial.C ((m - 2)^2)) -: {m : ℤ | ∃ a : Polynomial ℤ, ∃ b : Polynomial ℤ, P m = a * b ∧ +: {m : ℤ | ∃ a b, P m = a * b ∧ (∃ n ∈ Ici 1, a.coeff n ≠ 0) ∧ (∃ n ∈ Ici 1, b.coeff n ≠ 0)} = putnam_2001_a3_solution := sorry diff --git a/lean4/src/putnam_2001_b1.lean b/lean4/src/putnam_2001_b1.lean index add7046..756d0c5 100644 --- a/lean4/src/putnam_2001_b1.lean +++ b/lean4/src/putnam_2001_b1.lean @@ -6,13 +6,12 @@ open Topology Filter Polynomial Set Let $n$ be an even positive integer. Write the numbers $1,2,\ldots,n^2$ in the squares of an $n \times n$ grid so that the $k$-th row, from left to right, is $(k-1)n+1,(k-1)n+2,\ldots,(k-1)n+n$. Color the squares of the grid so that half of the squares in each row and in each column are red and the other half are black (a checkerboard coloring is one possibility). Prove that for each coloring, the sum of the numbers on the red squares is equal to the sum of the numbers on the black squares. -/ theorem putnam_2001_b1 -(n : ℕ) -(nums : Fin n → Fin n → ℤ) -(colors : Fin n → Fin n → Fin 2) -(npos : n > 0) -(neven : Even n) -(hnums : ∀ k l : Fin n, nums k l = k * n + l + 1) -(hcolorsrows : ∀ k : Fin n, (∑ l : Fin n, (if (colors k l = 0) then 1 else 0)) = n / 2) -(hcolorscols : ∀ l : Fin n, (∑ k : Fin n, (if (colors k l = 0) then 1 else 0)) = n / 2) -: (∑ k : Fin n, ∑ l : Fin n, (if (colors k l = 0) then (nums k l) else 0)) = (∑ k : Fin n, ∑ l : Fin n, (if (colors k l = 1) then (nums k l) else 0)) := + (n : ℕ) + (hn : n > 0) (hn' : Even n) + (nums : Fin n → Fin n → ℤ) + (colors : Fin n → Fin n → Fin 2) + (hnums : ∀ k l, nums k l = k * n + l + 1) + (hcolorsrows : ∀ k, (∑ l, (if (colors k l = 0) then 1 else 0)) = n / 2) + (hcolorscols : ∀ l, (∑ k, (if (colors k l = 0) then 1 else 0)) = n / 2) : + (∑ k, ∑ l, (if (colors k l = 0) then (nums k l) else 0)) = (∑ k, ∑ l, (if (colors k l = 1) then (nums k l) else 0)) := sorry diff --git a/lean4/src/putnam_2001_b6.lean b/lean4/src/putnam_2001_b6.lean index b62cb76..4fca6a9 100644 --- a/lean4/src/putnam_2001_b6.lean +++ b/lean4/src/putnam_2001_b6.lean @@ -8,9 +8,10 @@ abbrev putnam_2001_b6_solution : Prop := sorry /-- Assume that $(a_n)_{n \geq 1}$ is an increasing sequence of positive real numbers such that $\lim a_n/n=0$. Must there exist infinitely many positive integers $n$ such that $a_{n-i}+a_{n+i}<2a_n$ for $i=1,2,\ldots,n-1$? -/ -theorem putnam_2001_b6 -(aposinc alim : (ℤ → ℝ) → Prop) -(haposinc : ∀ a : ℤ → ℝ, aposinc a ↔ ∀ n ≥ 1, a n > 0 ∧ a n < a (n + 1)) -(halim : ∀ a : ℤ → ℝ, alim a ↔ Tendsto (fun n : ℤ => a (n + 1) / (n + 1)) atTop (𝓝 0)) -: (∀ a : ℤ → ℝ, (aposinc a ∧ alim a) → {n : ℤ | n > 0 ∧ (∀ i ∈ Set.Icc 1 (n - 1), a (n - i) + a (n + i) < 2 * a n)}.Infinite) ↔ putnam_2001_b6_solution := +theorem putnam_2001_b6 : + (∀ a : ℤ → ℝ, + (∀ n ≥ 1, a n > 0 ∧ a n < a (n + 1)) → + (Tendsto (fun n : ℤ => a (n + 1) / (n + 1)) atTop (𝓝 0)) → + {n : ℤ | n > 0 ∧ (∀ i ∈ Set.Icc 1 (n - 1), a (n - i) + a (n + i) < 2 * a n)}.Infinite) + ↔ putnam_2001_b6_solution := sorry diff --git a/lean4/src/putnam_2002_a1.lean b/lean4/src/putnam_2002_a1.lean index 876cc09..0cf1c8a 100644 --- a/lean4/src/putnam_2002_a1.lean +++ b/lean4/src/putnam_2002_a1.lean @@ -11,6 +11,6 @@ theorem putnam_2002_a1 (k : ℕ) (P : ℕ → Polynomial ℝ) (kpos : k > 0) -(Pderiv : ∀ n : ℕ, ∀ x : ℝ, iteratedDeriv n (fun x' : ℝ => 1 / (x' ^ k - 1)) x = ((P n).eval x) / ((x ^ k - 1) ^ (n + 1))) -: ∀ n : ℕ, (P n).eval 1 = putnam_2002_a1_solution k n := +(Pderiv : ∀ n x, iteratedDeriv n (fun x' : ℝ => 1 / (x' ^ k - 1)) x = ((P n).eval x) / ((x ^ k - 1) ^ (n + 1))) +: ∀ n, (P n).eval 1 = putnam_2002_a1_solution k n := sorry diff --git a/lean4/src/putnam_2003_a1.lean b/lean4/src/putnam_2003_a1.lean index 2ddf5ee..f0f5fc5 100644 --- a/lean4/src/putnam_2003_a1.lean +++ b/lean4/src/putnam_2003_a1.lean @@ -9,7 +9,11 @@ abbrev putnam_2003_a1_solution : ℕ → ℕ := sorry Let $n$ be a fixed positive integer. How many ways are there to write $n$ as a sum of positive integers, \[ n = a_1 + a_2 + \dots + a_k, \] with $k$ an arbitrary positive integer and $a_1 \leq a_2 \leq \dots \leq a_k \leq a_1 + 1$? For example, with $n = 4$, there are four ways: $4, 2 + 2, 1 + 1 + 2, 1 + 1 + 1 + 1$ -/ theorem putnam_2003_a1 -(n : ℕ) -(hn : n > 0) -: Set.encard {a : ℕ → ℤ | ∃ k > 0, (∑ i : Fin k, a i = n) ∧ (∀ i : Fin k, a i > 0) ∧ (∀ i : Fin (k - 1), a i ≤ a (i + 1)) ∧ a (k - 1) ≤ a 0 + 1 ∧ (∀ i ≥ k, a i = 0)} = putnam_2003_a1_solution n := +(n : ℕ) (hn : n > 0) +: Set.encard {a : ℕ → ℤ | + ∃ k > 0, (∑ i : Fin k, a i = n) ∧ + (∀ i : Fin k, a i > 0) ∧ + (∀ i : Fin (k - 1), a i ≤ a (i + 1)) ∧ + a (k - 1) ≤ a 0 + 1 ∧ (∀ i ≥ k, a i = 0)} + = putnam_2003_a1_solution n := sorry diff --git a/lean4/src/putnam_2003_a3.lean b/lean4/src/putnam_2003_a3.lean index 8f7776f..9e081b2 100644 --- a/lean4/src/putnam_2003_a3.lean +++ b/lean4/src/putnam_2003_a3.lean @@ -1,6 +1,6 @@ import Mathlib -open MvPolynomial +open Set noncomputable abbrev putnam_2003_a3_solution : ℝ := sorry -- 2 * Real.sqrt 2 - 1 @@ -8,7 +8,7 @@ noncomputable abbrev putnam_2003_a3_solution : ℝ := sorry Find the minimum value of $|\sin x+\cos x+\tan x+\cot x+\sec x+\csc x|$ for real numbers $x$. -/ theorem putnam_2003_a3 -(f : ℝ → ℝ) -(hf : ∀ x : ℝ, f x = |Real.sin x + Real.cos x + Real.tan x + 1 / Real.tan x + 1 / Real.cos x + 1 / Real.sin x|) -: (∃ x : ℝ, f x = putnam_2003_a3_solution) ∧ (∀ x : ℝ, f x ≥ putnam_2003_a3_solution) := -sorry + (f : ℝ → ℝ) + (hf : ∀ x : ℝ, f x = |Real.sin x + Real.cos x + Real.tan x + 1 / Real.tan x + 1 / Real.cos x + 1 / Real.sin x|) : + IsLeast (f '' univ) putnam_2003_a3_solution := + sorry diff --git a/lean4/src/putnam_2003_a5.lean b/lean4/src/putnam_2003_a5.lean index f2fcef2..5392c7b 100644 --- a/lean4/src/putnam_2003_a5.lean +++ b/lean4/src/putnam_2003_a5.lean @@ -6,13 +6,14 @@ open MvPolynomial Set A Dyck $n$-path is a lattice path of $n$ upsteps $(1,1)$ and $n$ downsteps $(1,-1)$ that starts at the origin $O$ and never dips below the $x$-axis. A return is a maximal sequence of contiguous downsteps that terminates on the $x$-axis. Show that there is a one-to-one correspondence between the Dyck $n$-paths with no return of even length and the Dyck $(n-1)$-paths. -/ theorem putnam_2003_a5 -(n : ℕ) -(npos : n > 0) +(n : ℕ) (npos : n > 0) (dyckpath : (m : ℕ) → Set ((Fin (2 * m)) → ℤ)) -(hdyckpath : dyckpath = fun m ↦ {p : Fin (2 * m) → ℤ | - range p ⊆ {-1, 1} ∧ ∑ k : Fin (2 * m), p k = 0 ∧ ∀ j : Fin (2 * m), ∑ k : Fin (2 * m), ite (k ≤ j) (p k) 0 ≥ 0}) +(hdyckpath : dyckpath = fun m ↦ {p | + range p ⊆ {-1, 1} ∧ ∑ k, p k = 0 ∧ ∀ j, ∑ k, ite (k ≤ j) (p k) 0 ≥ 0}) (noevenreturn : (m : ℕ) → Set ((Fin (2 * m)) → ℤ)) -(hnoevenreturn : noevenreturn = fun m ↦ {p : Fin (2 * m) → ℤ | - ¬∃ i j : Fin (2 * m), i < j ∧ p i = 1 ∧ (∀ k ∈ Ioc i j, p i = -1) ∧ Even (j.1 - i.1) ∧ ∑ k : Fin (2 * m), ite (k ≤ j) (p k) 0 = 0}) -: (∃ f : ((Fin (2 * n)) → ℤ) → (Fin (2 * (n - 1)) → ℤ), ∀ y ∈ dyckpath (n - 1), ∃! x, x ∈ dyckpath n ∩ noevenreturn n ∧ f x = y) := +(hnoevenreturn : noevenreturn = fun m ↦ {p | + ¬∃ i j, i < j ∧ p i = 1 ∧ (∀ k ∈ Ioc i j, p i = -1) ∧ + Even (j.1 - i.1) ∧ ∑ k, ite (k ≤ j) (p k) 0 = 0}) + : ∃ f : ((Fin (2 * n)) → ℤ) → (Fin (2 * (n - 1)) → ℤ), + ∀ y ∈ dyckpath (n - 1), ∃! x, x ∈ dyckpath n ∩ noevenreturn n ∧ f x = y := sorry diff --git a/lean4/src/putnam_2003_a6.lean b/lean4/src/putnam_2003_a6.lean index 6432fa2..895b1b1 100644 --- a/lean4/src/putnam_2003_a6.lean +++ b/lean4/src/putnam_2003_a6.lean @@ -9,6 +9,6 @@ For a set $S$ of nonnegative integers, let $r_S(n)$ denote the number of ordered -/ theorem putnam_2003_a6 (r : Set ℕ → ℕ → ℕ) -(hr : ∀ (S : Set ℕ) (n : ℕ), r S n = ∑' s1 : S, ∑' s2 : S, if (s1 ≠ s2 ∧ s1 + s2 = n) then 1 else 0) +(hr : ∀ S n, r S n = ∑' s1 : S, ∑' s2 : S, if (s1 ≠ s2 ∧ s1 + s2 = n) then 1 else 0) : (∃ A B : Set ℕ, A ∪ B = ℕ ∧ A ∩ B = ∅ ∧ (∀ n : ℕ, r A n = r B n)) ↔ putnam_2003_a6_solution := sorry diff --git a/lean4/src/putnam_2003_b2.lean b/lean4/src/putnam_2003_b2.lean index 5aa07b1..9aca305 100644 --- a/lean4/src/putnam_2003_b2.lean +++ b/lean4/src/putnam_2003_b2.lean @@ -7,10 +7,11 @@ open MvPolynomial Set Let $n$ be a positive integer. Starting with the sequence $$1, \frac{1}{2}, \frac{1}{3}, \dots, \frac{1}{n},$$ form a new sequence of $n-1$ entries $$\frac{3}{4}, \frac{5}{12}, \dots, \frac{2n-1}{2n(n-1)}$$ by taking the averages of two consecutive entries in the first sequence. Repeat the averaging of neighbors on the second sequence to obtain a third sequence of $n-2$ entries, and continue until the final sequence produced consists of a single number $x_n$. Show that $x_n < 2/n$. -/ theorem putnam_2003_b2 -(n : ℕ) -(hn : n > 0) -(seq : ℕ → ℕ → ℚ) -(hinit : ∀ j ∈ Icc 1 n, seq 0 j = 1 / j) -(havg : ∀ k ∈ Icc 1 (n - 1), ∀ j ∈ Icc 1 (n - k), seq k j = (seq (k - 1) j + seq (k - 1) (j + 1)) / 2) -: (seq (n - 1) 1 < 2 / n) := + (n : ℕ) + (hn : n > 0) + (seq : ℕ → ℕ → ℚ) + (hinit : ∀ j ∈ Icc 1 n, seq 0 j = 1 / j) + (havg : ∀ᵉ (k ∈ Icc 1 (n - 1)) (j ∈ Icc 1 (n - k)), + seq k j = (seq (k - 1) j + seq (k - 1) (j + 1)) / 2) : + (seq (n - 1) 1 < 2 / n) := sorry diff --git a/lean4/src/putnam_2003_b4.lean b/lean4/src/putnam_2003_b4.lean index 6caffa2..0a14804 100644 --- a/lean4/src/putnam_2003_b4.lean +++ b/lean4/src/putnam_2003_b4.lean @@ -6,11 +6,11 @@ open MvPolynomial Set Nat Let $f(z)=az^4+bz^3+cz^2+dz+e=a(z-r_1)(z-r_2)(z-r_3)(z-r_4)$ where $a,b,c,d,e$ are integers, $a \neq 0$. Show that if $r_1+r_2$ is a rational number and $r_1+r_2 \neq r_3+r_4$, then $r_1r_2$ is a rational number. -/ theorem putnam_2003_b4 -(f : ℝ → ℝ) -(a b c d e : ℤ) -(r1 r2 r3 r4 : ℝ) -(ane0 : a ≠ 0) -(hf1 : ∀ z : ℝ, f z = a * z ^ 4 + b * z ^ 3 + c * z ^ 2 + d * z + e) -(hf2 : ∀ z : ℝ, f z = a * (z - r1) * (z - r2) * (z - r3) * (z - r4)) -: (¬Irrational (r1 + r2) ∧ r1 + r2 ≠ r3 + r4) → ¬Irrational (r1 * r2) := + (f : ℝ → ℝ) + (a b c d e : ℤ) + (r1 r2 r3 r4 : ℝ) + (ane0 : a ≠ 0) + (hf1 : ∀ z, f z = a * z ^ 4 + b * z ^ 3 + c * z ^ 2 + d * z + e) + (hf2 : ∀ z, f z = a * (z - r1) * (z - r2) * (z - r3) * (z - r4)) : + (¬Irrational (r1 + r2) ∧ r1 + r2 ≠ r3 + r4) → ¬Irrational (r1 * r2) := sorry diff --git a/lean4/src/putnam_2004_a1.lean b/lean4/src/putnam_2004_a1.lean index 9d0fb18..c5eaf01 100644 --- a/lean4/src/putnam_2004_a1.lean +++ b/lean4/src/putnam_2004_a1.lean @@ -8,7 +8,10 @@ abbrev putnam_2004_a1_solution : Prop := sorry Basketball star Shanille O'Keal's team statistician keeps track of the number, $S(N)$, of successful free throws she has made in her first $N$ attempts of the season. Early in the season, $S(N)$ was less than $80\%$ of $N$, but by the end of the season, $S(N)$ was more than $80\%$ of $N$. Was there necessarily a moment in between when $S(N)$ was exactly $80\%$ of $N$? -/ theorem putnam_2004_a1 -(S : (ℕ → Fin 2) → ℕ → ℝ) -(hS : ∀ attempts : ℕ → Fin 2, ∀ N ≥ 1, S attempts N = (∑ i : Fin N, (attempts i).1) / N) -: (∀ (attempts : ℕ → Fin 2) (a b : ℕ), (1 ≤ a ∧ a < b ∧ S attempts a < 0.8 ∧ S attempts b > 0.8) → (∃ c : ℕ, a < c ∧ c < b ∧ S attempts c = 0.8)) ↔ putnam_2004_a1_solution := + (S : (ℕ → Fin 2) → ℕ → ℝ) + (hS : ∀ attempts, ∀ N ≥ 1, S attempts N = (∑ i : Fin N, (attempts i).1) / N) : + (∀ attempts a b, + (1 ≤ a ∧ a < b ∧ S attempts a < 0.8 ∧ S attempts b > 0.8) → + (∃ c : ℕ, a < c ∧ c < b ∧ S attempts c = 0.8)) + ↔ putnam_2004_a1_solution := sorry diff --git a/lean4/src/putnam_2004_a4.lean b/lean4/src/putnam_2004_a4.lean index e389f2b..7d2d8c1 100644 --- a/lean4/src/putnam_2004_a4.lean +++ b/lean4/src/putnam_2004_a4.lean @@ -7,10 +7,11 @@ open Nat Topology Filter Show that for any positive integer $n$ there is an integer $N$ such that the product $x_1x_2 \cdots x_n$ can be expressed identically in the form $x_1x_2 \cdots x_n=\sum_{i=1}^Nc_i(a_{i1}x_1+a_{i2}x_2+\cdots+a_{in}x_n)^n$ where the $c_i$ are rational numbers and each $a_{ij}$ is one of the numbers $-1,0,1$. -/ theorem putnam_2004_a4 -(n : ℕ) -(x : Fin n → ℝ) -(avals : ℕ → (ℕ → Fin n → ℝ) → Prop) -(npos : n > 0) -(havals : ∀ (N : ℕ) (a : (ℕ → Fin n → ℝ)), avals N a = ∀ (i : Fin N) (j : Fin n), (a i j = -1 ∨ a i j = 0 ∨ a i j = 1)) -: ∃ (N : ℕ) (c : Fin N → ℚ) (a : ℕ → Fin n → ℝ), avals N a ∧ ((∏ i : Fin n, x i) = ∑ i : Fin N, c i * (∑ j : Fin n, a i j * x j) ^ n) := + (n : ℕ) (npos : n > 0) + (x : Fin n → ℝ) + (avals : ℕ → (ℕ → Fin n → ℝ) → Prop) + (havals : ∀ N a, avals N a ↔ ∀ (i : Fin N) (j : Fin n), (a i j = -1 ∨ a i j = 0 ∨ a i j = 1)) : + ∃ (N : ℕ) (c : Fin N → ℚ) (a : ℕ → Fin n → ℝ), + avals N a ∧ + (∏ i : Fin n, x i) = ∑ i : Fin N, c i * (∑ j : Fin n, a i j * x j) ^ n := sorry diff --git a/lean4/src/putnam_2004_b6.lean b/lean4/src/putnam_2004_b6.lean index e23032b..088c3d7 100644 --- a/lean4/src/putnam_2004_b6.lean +++ b/lean4/src/putnam_2004_b6.lean @@ -6,14 +6,12 @@ open Nat Topology Filter Let $\mathcal{A}$ be a non-empty set of positive integers, and let $N(x)$ denote the number of elements of $\mathcal{A}$ not exceeding $x$. Let $\mathcal{B}$ denote the set of positive integers $b$ that can be written in the form $b=a-a'$ with $a \in \mathcal{A}$ and $a' \in \mathcal{A}$. Let $b_1 0) -(hN : ∀ x : ℝ, N x = Set.encard {a : A | a ≤ x}) -(hB : B = {b' > 0 | ∃ a ∈ A, ∃ a' ∈ A, b' = a - a'}) -(hbB : Set.range b = B ∧ ∀ i : ℕ, b i < b (i + 1)) -: (∀ r : ℕ, ∃ i : ℕ, (b (i + 1) - b i) ≥ r) → Tendsto (fun x => N x / x) atTop (𝓝 0) := + (A B : Set ℕ) + (N b : ℝ → ℕ) + (Anempty : A.Nonempty) + (Apos : ∀ a ∈ A, a > 0) + (hN : ∀ x : ℝ, N x = Set.encard {a : A | a ≤ x}) + (hB : B = {b' > 0 | ∃ a ∈ A, ∃ a' ∈ A, b' = a - a'}) + (hbB : Set.range b = B ∧ ∀ i : ℕ, b i < b (i + 1)) : + (∀ r : ℕ, ∃ i : ℕ, (b (i + 1) - b i) ≥ r) → Tendsto (fun x => N x / x) atTop (𝓝 0) := sorry diff --git a/lean4/src/putnam_2005_a5.lean b/lean4/src/putnam_2005_a5.lean index 64aa9ce..ca7c538 100644 --- a/lean4/src/putnam_2005_a5.lean +++ b/lean4/src/putnam_2005_a5.lean @@ -7,5 +7,6 @@ noncomputable abbrev putnam_2005_a5_solution : ℝ := sorry /-- Evaluate $\int_0^1 \frac{\ln(x+1)}{x^2+1}\,dx$. -/ -theorem putnam_2005_a5 : ∫ x in (0:ℝ)..1, (Real.log (x+1))/(x^2 + 1) = putnam_2005_a5_solution := +theorem putnam_2005_a5 : + ∫ x in (0:ℝ)..1, (Real.log (x+1))/(x^2 + 1) = putnam_2005_a5_solution := sorry diff --git a/lean4/src/putnam_2006_b1.lean b/lean4/src/putnam_2006_b1.lean index 04fccb5..2ea5dd6 100644 --- a/lean4/src/putnam_2006_b1.lean +++ b/lean4/src/putnam_2006_b1.lean @@ -6,9 +6,12 @@ noncomputable abbrev putnam_2006_b1_solution : ℝ := sorry Show that the curve $x^3 + 3xy + y^3 = 1$ contains only one set of three distinct points, $A$, $B$, and $C$, which are vertices of an equilateral triangle, and find its area. -/ theorem putnam_2006_b1 -(curve : Set (ℝ × ℝ)) -(hcurve : curve = {(x, y) | x ^ 3 + 3 * x * y + y ^ 3 = 1}) -(equilateral : Set (ℝ × ℝ) → Prop) -(hequilateral : equilateral = fun S ↦ S.encard = 3 ∧ ∃ d : ℝ, ∀ P ∈ S, ∀ Q ∈ S, P ≠ Q → Real.sqrt ((P.1 - Q.1)^2 + (P.2 - Q.2)^2) = d) -: ((∃! S : Set (ℝ × ℝ), S ⊆ curve ∧ equilateral S) ∧ (∃ S : Set (ℝ × ℝ), S ⊆ curve ∧ equilateral S ∧ (MeasureTheory.volume (convexHull ℝ S)).toReal = putnam_2006_b1_solution)) := + (curve : Set (ℝ × ℝ)) + (hcurve : curve = {(x, y) | x ^ 3 + 3 * x * y + y ^ 3 = 1}) + (equilateral : Set (ℝ × ℝ) → Prop) + (hequilateral : ∀ S, equilateral S ↔ S.encard = 3 ∧ + ∃ d : ℝ, ∀ P ∈ S, ∀ Q ∈ S, P ≠ Q → + Real.sqrt ((P.1 - Q.1)^2 + (P.2 - Q.2)^2) = d) : + (∃! S : Set (ℝ × ℝ), S ⊆ curve ∧ equilateral S) ∧ + (∃ S : Set (ℝ × ℝ), S ⊆ curve ∧ equilateral S ∧ (MeasureTheory.volume (convexHull ℝ S)).toReal = putnam_2006_b1_solution) := sorry diff --git a/lean4/src/putnam_2006_b4.lean b/lean4/src/putnam_2006_b4.lean index f2249b6..bb91cbc 100644 --- a/lean4/src/putnam_2006_b4.lean +++ b/lean4/src/putnam_2006_b4.lean @@ -6,12 +6,12 @@ noncomputable abbrev putnam_2006_b4_solution : ℕ → ℕ := sorry Let $Z$ denote the set of points in $\mathbb{R}^n$ whose coordinates are $0$ or $1$. (Thus $Z$ has $2^n$ elements, which are the vertices of a unit hypercube in $\mathbb{R}^n$.) Given a vector subspace $V$ of $\mathbb{R}^n$, let $Z(V)$ denote the number of members of $Z$ that lie in $V$. Let $k$ be given, $0 \leq k \leq n$. Find the maximum, over all vector subspaces $V \subseteq \mathbb{R}^n$ of dimension $k$, of the number of points in $V \cap Z$. -/ theorem putnam_2006_b4 -(n k max : ℕ) -(npos : n > 0) -(hk : k ≤ n) -(Z : Set (Fin n → ℝ)) -(hZ : Z = {P : Fin n → ℝ | ∀ j : Fin n, P j = 0 ∨ P j = 1}) -(hmaxeq : ∃ V : Subspace ℝ (Fin n → ℝ), Module.rank ℝ V = k ∧ (Z ∩ V).ncard = max) -(hmaxub : ∀ V : Subspace ℝ (Fin n → ℝ), Module.rank ℝ V = k → (Z ∩ V).ncard ≤ max) -: (max = putnam_2006_b4_solution k) := -sorry + (n k : ℕ) + (npos : n > 0) + (hk : k ≤ n) + (Z : Set (Fin n → ℝ)) + (hZ : Z = {P : Fin n → ℝ | ∀ j : Fin n, P j = 0 ∨ P j = 1}) : + IsGreatest + {y | ∃ V : Subspace ℝ (Fin n → ℝ), Module.rank ℝ V = k ∧ (Z ∩ V).ncard = y} + (putnam_2006_b4_solution k) := + sorry diff --git a/lean4/src/putnam_2006_b5.lean b/lean4/src/putnam_2006_b5.lean index 06c6a31..e323756 100644 --- a/lean4/src/putnam_2006_b5.lean +++ b/lean4/src/putnam_2006_b5.lean @@ -8,11 +8,10 @@ noncomputable abbrev putnam_2006_b5_solution : ℝ := sorry For each continuous function $f: [0,1] \to \mathbb{R}$, let $I(f) = \int_0^1 x^2 f(x)\,dx$ and $J(x) = \int_0^1 x \left(f(x)\right)^2\,dx$. Find the maximum value of $I(f) - J(f)$ over all such functions $f$. -/ theorem putnam_2006_b5 -(I J : (ℝ → ℝ) → ℝ) -(hI : I = fun f ↦ ∫ x in (0)..1, x ^ 2 * (f x)) -(hJ : J = fun f ↦ ∫ x in (0)..1, x * (f x) ^ 2) -(max : ℝ) -(heqmax : ∃ f : ℝ → ℝ, ContinuousOn f (Icc 0 1) ∧ I f - J f = max) -(hmaxub : ∀ f : ℝ → ℝ, ContinuousOn f (Icc 0 1) → I f - J f ≤ max) -: (max = putnam_2006_b5_solution) := + (I J : (ℝ → ℝ) → ℝ) + (hI : I = fun f ↦ ∫ x in (0)..1, x ^ 2 * (f x)) + (hJ : J = fun f ↦ ∫ x in (0)..1, x * (f x) ^ 2) : + IsGreatest + {y | ∃ f : ℝ → ℝ, ContinuousOn f (Icc 0 1) ∧ I f - J f = y} + putnam_2006_b5_solution := sorry diff --git a/lean4/src/putnam_2007_a2.lean b/lean4/src/putnam_2007_a2.lean index 77c634f..3c5be29 100644 --- a/lean4/src/putnam_2007_a2.lean +++ b/lean4/src/putnam_2007_a2.lean @@ -1,14 +1,19 @@ import Mathlib +open MeasureTheory + noncomputable abbrev putnam_2007_a2_solution : ENNReal := sorry -- 4 /-- Find the least possible area of a convex set in the plane that intersects both branches of the hyperbola $xy=1$ and both branches of the hyperbola $xy=-1$. (A set $S$ in the plane is called \emph{convex} if for any two points in $S$ the line segment connecting them is contained in $S$.) -/ -theorem putnam_2007_a2 -(Sinterpos Sinterneg Sconv : Set (Fin 2 → ℝ) → Prop) -(hSinterpos : ∀ S : Set (Fin 2 → ℝ), Sinterpos S = ((∃ p ∈ S, p 0 > 0 ∧ p 1 > 0 ∧ p 0 * p 1 = 1) ∧ (∃ p ∈ S, p 0 < 0 ∧ p 1 < 0 ∧ p 0 * p 1 = 1))) -(hSinterneg : ∀ S : Set (Fin 2 → ℝ), Sinterneg S = ((∃ p ∈ S, p 0 < 0 ∧ p 1 > 0 ∧ p 0 * p 1 = -1) ∧ (∃ p ∈ S, p 0 > 0 ∧ p 1 < 0 ∧ p 0 * p 1 = -1))) -(hSconv : ∀ S : Set (Fin 2 → ℝ), Sconv S = (Convex ℝ S ∧ Sinterpos S ∧ Sinterneg S)) -: (∃ S : Set (Fin 2 → ℝ), Sconv S ∧ MeasureTheory.volume S = putnam_2007_a2_solution) ∧ (∀ S : Set (Fin 2 → ℝ), Sconv S → MeasureTheory.volume S ≥ putnam_2007_a2_solution) := +theorem putnam_2007_a2 : + IsLeast + {y | ∃ S : Set (Fin 2 → ℝ), + Convex ℝ S ∧ + (∃ p ∈ S, p 0 > 0 ∧ p 1 > 0 ∧ p 0 * p 1 = 1) ∧ + (∃ p ∈ S, p 0 < 0 ∧ p 1 < 0 ∧ p 0 * p 1 = 1) ∧ + (∃ p ∈ S, p 0 < 0 ∧ p 1 > 0 ∧ p 0 * p 1 = -1) ∧ + (∃ p ∈ S, p 0 > 0 ∧ p 1 < 0 ∧ p 0 * p 1 = -1) ∧ + volume S = y} putnam_2007_a2_solution := sorry diff --git a/lean4/src/putnam_2007_a4.lean b/lean4/src/putnam_2007_a4.lean index 10b77c9..ab3b632 100644 --- a/lean4/src/putnam_2007_a4.lean +++ b/lean4/src/putnam_2007_a4.lean @@ -10,7 +10,7 @@ A \emph{repunit} is a positive integer whose digits in base 10 are all ones. Fin theorem putnam_2007_a4 (S : Set (Polynomial ℝ)) (repunit : ℝ → Prop) -(hrepunit : repunit = fun (x : ℝ) ↦ x > 0 ∧ x = floor x ∧ ∀ d ∈ (digits 10 (floor x)), d = 1) -(hS : ∀ f : Polynomial ℝ, f ∈ S ↔ (∀ n : ℝ, repunit n → repunit (f.eval n))) +(hrepunit : ∀ x, repunit x ↔ x > 0 ∧ x = floor x ∧ ∀ d ∈ (digits 10 (floor x)), d = 1) +(hS : ∀ f, f ∈ S ↔ (∀ n : ℝ, repunit n → repunit (f.eval n))) : (S = putnam_2007_a4_solution) := sorry diff --git a/lean4/src/putnam_2007_b4.lean b/lean4/src/putnam_2007_b4.lean index 2b6b160..362b270 100644 --- a/lean4/src/putnam_2007_b4.lean +++ b/lean4/src/putnam_2007_b4.lean @@ -12,7 +12,6 @@ Let $n$ be a positive integer. Find the number of pairs $P, Q$ of polynomials wi and $\deg P > \deg Q$. -/ theorem putnam_2007_b4 -(n : ℕ) -(npos : n > 0) -: ({(P, Q) : (Polynomial ℝ) × (Polynomial ℝ) | P ^ 2 + Q ^ 2 = Polynomial.X ^ (2 * n) + 1 ∧ P.degree > Q.degree}.ncard = putnam_2007_b4_solution n) := +(n : ℕ) (npos : n > 0) : +({(P, Q) : (Polynomial ℝ) × (Polynomial ℝ) | P ^ 2 + Q ^ 2 = Polynomial.X ^ (2 * n) + 1 ∧ P.degree > Q.degree}.ncard = putnam_2007_b4_solution n) := sorry diff --git a/lean4/src/putnam_2007_b5.lean b/lean4/src/putnam_2007_b5.lean index 22ce05b..d52e10d 100644 --- a/lean4/src/putnam_2007_b5.lean +++ b/lean4/src/putnam_2007_b5.lean @@ -12,7 +12,6 @@ Let $k$ be a positive integer. Prove that there exist polynomials $P_0(n), P_1(n ($\lfloor a \rfloor$ means the largest integer $\leq a$.) -/ theorem putnam_2007_b5 -(k : ℕ) -(kpos : k > 0) -: (∃ P : Finset.range k → Polynomial ℝ, ∀ n : ℤ, ⌊(n : ℝ) / k⌋ ^ k = ∑ i : Finset.range k, (P i).eval (n : ℝ) * ⌊(n : ℝ) / k⌋ ^ (i : ℕ)) := +(k : ℕ) (kpos : k > 0) : +(∃ P : Finset.range k → Polynomial ℝ, ∀ n : ℤ, ⌊(n : ℝ) / k⌋ ^ k = ∑ i : Finset.range k, (P i).eval (n : ℝ) * ⌊(n : ℝ) / k⌋ ^ (i : ℕ)) := sorry diff --git a/lean4/src/putnam_2008_a3.lean b/lean4/src/putnam_2008_a3.lean index a5a2df7..5151001 100644 --- a/lean4/src/putnam_2008_a3.lean +++ b/lean4/src/putnam_2008_a3.lean @@ -4,17 +4,15 @@ import Mathlib Start with a finite sequence $a_1, a_2, \dots, a_n$ of positive integers. If possible, choose two indices $j < k$ such that $a_j$ does not divide $a_k$, and replace $a_j$ and $a_k$ by $\mathrm{gcd}(a_j, a_k)$ and $\mathrm{lcm}(a_j, a_k)$, respectively. Prove that if this process is repeated, it must eventually stop and the final sequence does not depend on the choices made. -/ theorem putnam_2008_a3 -(n : ℕ) -(npos : n > 0) -(a : Fin n → ℕ) -(apos : ∀ i : Fin n, a i > 0) -(cont : (Fin n → ℕ) → Prop) -(hcont : cont = fun s ↦ ∃ j k : Fin n, j < k ∧ ¬(s j ∣ s k)) -(init : (ℕ → Fin n → ℕ) → Prop) -(hinit : init = fun P ↦ P 0 = a) -(trans : (ℕ → Fin n → ℕ) → Prop) -(htrans : trans = fun P ↦ ∀ t : ℕ, cont (P t) → - ∃ j k : Fin n, j < k ∧ ¬(P t j ∣ P t k) ∧ P (t + 1) j = Nat.gcd (P t j) (P t k) ∧ P (t + 1) k = Nat.lcm (P t j) (P t k) ∧ - ∀ i : Fin n, i ≠ j → i ≠ k → P (t + 1) i = P t i) -: (∃ f : Fin n → ℕ, ∀ P : ℕ → Fin n → ℕ, init P → trans P → ∃ t : ℕ, ¬cont (P t) ∧ P t = f) := + (n : ℕ) (npos : n > 0) + (a : Fin n → ℕ) + (apos : ∀ i : Fin n, a i > 0) + (cont : (Fin n → ℕ) → Prop) + (hcont : ∀ s, cont s ↔ ∃ j k : Fin n, j < k ∧ ¬(s j ∣ s k)) + (init trans : (ℕ → Fin n → ℕ) → Prop) + (hinit : ∀ P, init P ↔ P 0 = a) + (htrans : ∀ P, trans P ↔ ∀ t : ℕ, cont (P t) → + ∃ j k : Fin n, j < k ∧ ¬(P t j ∣ P t k) ∧ P (t + 1) j = Nat.gcd (P t j) (P t k) ∧ P (t + 1) k = Nat.lcm (P t j) (P t k) ∧ + ∀ i : Fin n, i ≠ j → i ≠ k → P (t + 1) i = P t i) : + (∃ f, ∀ P, init P → trans P → ∃ t : ℕ, ¬cont (P t) ∧ P t = f) := sorry diff --git a/lean4/src/putnam_2008_b1.lean b/lean4/src/putnam_2008_b1.lean index c7bde71..2a41cec 100644 --- a/lean4/src/putnam_2008_b1.lean +++ b/lean4/src/putnam_2008_b1.lean @@ -10,6 +10,5 @@ def real_circle (c : EuclideanSpace ℝ (Fin 2)) (r : ℝ) : Set (EuclideanSpace What is the maximum number of rational points that can lie on a circle in $\mathbb{R}^2$ whose center is not a rational point? (A \emph{rational point} is a point both of whose coordinates are rational numbers.) -/ theorem putnam_2008_b1 : -∀ (c : EuclideanSpace ℝ (Fin 2)) (r : ℝ), ¬ is_rational_point c → (Set.ncard {p : EuclideanSpace ℝ (Fin 2) | p ∈ real_circle c r ∧ is_rational_point p} ≤ putnam_2008_b1_solution) -∧ ∃ (c : EuclideanSpace ℝ (Fin 2)) (r : ℝ), ¬ is_rational_point c ∧ (Set.ncard {p : EuclideanSpace ℝ (Fin 2) | p ∈ real_circle c r ∧ is_rational_point p} = putnam_2008_b1_solution) := + IsGreatest {n : ℕ | ∃ (c : EuclideanSpace ℝ (Fin 2)) (r : ℝ), ¬ is_rational_point c ∧ (Set.ncard {p : EuclideanSpace ℝ (Fin 2) | p ∈ real_circle c r ∧ is_rational_point p} = n)} putnam_2008_b1_solution := sorry diff --git a/lean4/src/putnam_2008_b5.lean b/lean4/src/putnam_2008_b5.lean index edb1427..08ec097 100644 --- a/lean4/src/putnam_2008_b5.lean +++ b/lean4/src/putnam_2008_b5.lean @@ -9,6 +9,6 @@ Find all continuously differentiable functions f : \mathbb{R} \to \mathbb{R} suc -/ theorem putnam_2008_b5 (fqsat : (ℝ → ℝ) → ℚ → Prop) -(hfqsat : fqsat = fun (f : ℝ → ℝ) (q : ℚ) => ContDiff ℝ 1 f ∧ (∃ p : ℚ, p = f q ∧ p.den = q.den)) +(hfqsat : ∀ f q, fqsat f q ↔ ContDiff ℝ 1 f ∧ (∃ p : ℚ, p = f q ∧ p.den = q.den)) : ∀ f : (ℝ → ℝ), (∀ q : ℚ, fqsat f q) ↔ f ∈ putnam_2008_b5_solution := sorry diff --git a/lean4/src/putnam_2009_b1.lean b/lean4/src/putnam_2009_b1.lean index 09a54fb..df714dd 100644 --- a/lean4/src/putnam_2009_b1.lean +++ b/lean4/src/putnam_2009_b1.lean @@ -7,9 +7,9 @@ Show that every positive rational number can be written as a quotient of product -/ theorem putnam_2009_b1 (isquotprodprimefact : ℚ → Prop) -(hisquotprodprimefact : isquotprodprimefact = -fun (q : ℚ) => (∃ (k m : ℕ) (a : Fin k → ℕ) (b : Fin m → ℕ), -(∀ i : Fin k, Nat.Prime (a i)) ∧ (∀ j : Fin m, Nat.Prime (b j)) -∧ (q = (∏ i : Fin k, Nat.factorial (a i))/(∏ j : Fin m, Nat.factorial (b j))))) +(hisquotprodprimefact : ∀ q, isquotprodprimefact q ↔ + (∃ (k m : ℕ) (a : Fin k → ℕ) (b : Fin m → ℕ), + (∀ i : Fin k, Nat.Prime (a i)) ∧ (∀ j : Fin m, Nat.Prime (b j)) + ∧ (q = (∏ i : Fin k, Nat.factorial (a i))/(∏ j : Fin m, Nat.factorial (b j))))) : ∀ q : ℚ, q > 0 → isquotprodprimefact q := sorry diff --git a/lean4/src/putnam_2009_b3.lean b/lean4/src/putnam_2009_b3.lean index b749515..e197909 100644 --- a/lean4/src/putnam_2009_b3.lean +++ b/lean4/src/putnam_2009_b3.lean @@ -9,7 +9,7 @@ Call a subset $S$ of $\{1, 2, \dots, n\}$ \emph{mediocre} if it has the followin -/ theorem putnam_2009_b3 (mediocre : ℤ → Set ℤ → Prop) -(hmediocre : mediocre = fun n S ↦ (S ⊆ Icc 1 n) ∧ ∀ a ∈ S, ∀ b ∈ S, 2 ∣ a + b → (a + b) / 2 ∈ S) +(hmediocre : ∀ n S, mediocre n S ↔ (S ⊆ Icc 1 n) ∧ ∀ a ∈ S, ∀ b ∈ S, 2 ∣ a + b → (a + b) / 2 ∈ S) (A : ℤ → ℤ) (hA : A = fun n ↦ ({S : Set ℤ | mediocre n S}.ncard : ℤ)) : ({n : ℤ | n > 0 ∧ A (n + 2) - 2 * A (n + 1) + A n = 1} = putnam_2009_b3_solution) := diff --git a/lean4/src/putnam_2009_b6.lean b/lean4/src/putnam_2009_b6.lean index 995296d..2b2e5fb 100644 --- a/lean4/src/putnam_2009_b6.lean +++ b/lean4/src/putnam_2009_b6.lean @@ -6,7 +6,10 @@ open Topology MvPolynomial Filter Set Metric Prove that for every positive integer $n$, there is a sequence of integers $a_0, a_1, \dots, a_{2009}$ with $a_0 = 0$ and $a_{2009} = n$ such that each term after $a_0$ is either an earlier term plus $2^k$ for some nonnegative integer $k$, or of the form $b\,\mathrm{mod}\,c$ for some earlier positive terms $b$ and $c$. [Here $b\,\mathrm{mod}\,c$ denotes the remainder when $b$ is divided by $c$, so $0 \leq (b\,\mathrm{mod}\,c) < c$.] -/ theorem putnam_2009_b6 -(n : ℕ) -(npos : n > 0) -: (∃ a : ℕ → ℤ, a 0 = 0 ∧ a 2009 = n ∧ ∀ i : Icc 1 2009, ((∃ j k : ℕ, j < i ∧ a i = a j + 2 ^ k) ∨ ∃ b c : ℕ, b < i ∧ c < i ∧ a b > 0 ∧ a c > 0 ∧ a i = (a b) % (a c))) := + (n : ℕ) (npos : n > 0) : + (∃ a : ℕ → ℤ, + a 0 = 0 ∧ a 2009 = n ∧ + ∀ i : Icc 1 2009, + ((∃ j k : ℕ, j < i ∧ a i = a j + 2 ^ k) ∨ + ∃ b c : ℕ, b < i ∧ c < i ∧ a b > 0 ∧ a c > 0 ∧ a i = (a b) % (a c))) := sorry diff --git a/lean4/src/putnam_2010_b2.lean b/lean4/src/putnam_2010_b2.lean index c230b92..e3669f5 100644 --- a/lean4/src/putnam_2010_b2.lean +++ b/lean4/src/putnam_2010_b2.lean @@ -7,11 +7,9 @@ abbrev putnam_2010_b2_solution : ℕ := sorry Given that $A$, $B$, and $C$ are noncollinear points in the plane with integer coordinates such that the distances $AB$, $AC$, and $BC$ are integers, what is the smallest possible value of $AB$? -/ theorem putnam_2010_b2 -(ABCintcoords : EuclideanSpace ℝ (Fin 2) → EuclideanSpace ℝ (Fin 2) → EuclideanSpace ℝ (Fin 2) → Prop) -(ABCintdists : EuclideanSpace ℝ (Fin 2) → EuclideanSpace ℝ (Fin 2) → EuclideanSpace ℝ (Fin 2) → Prop) -(ABCall : EuclideanSpace ℝ (Fin 2) → EuclideanSpace ℝ (Fin 2) → EuclideanSpace ℝ (Fin 2) → Prop) -(hABCintcoords : ∀ A B C : EuclideanSpace ℝ (Fin 2), ABCintcoords A B C = (∀ i : Fin 2, A i = round (A i) ∧ B i = round (B i) ∧ C i = round (C i))) -(hABCintdists : ∀ A B C : EuclideanSpace ℝ (Fin 2), ABCintdists A B C = (dist A B = round (dist A B) ∧ dist A C = round (dist A C) ∧ dist B C = round (dist B C))) -(hABCall : ∀ A B C : EuclideanSpace ℝ (Fin 2), ABCall A B C = (¬Collinear ℝ {A, B, C} ∧ ABCintcoords A B C ∧ ABCintdists A B C)) -: (∃ A B C : EuclideanSpace ℝ (Fin 2), ABCall A B C ∧ dist A B = putnam_2010_b2_solution) ∧ (∀ A B C : EuclideanSpace ℝ (Fin 2), ABCall A B C → dist A B ≥ putnam_2010_b2_solution) := + (ABCintcoords ABCintdists ABCall: EuclideanSpace ℝ (Fin 2) → EuclideanSpace ℝ (Fin 2) → EuclideanSpace ℝ (Fin 2) → Prop) + (hABCintcoords : ∀ A B C, ABCintcoords A B C ↔ (∀ i : Fin 2, A i = round (A i) ∧ B i = round (B i) ∧ C i = round (C i))) + (hABCintdists : ∀ A B C, ABCintdists A B C ↔ (dist A B = round (dist A B) ∧ dist A C = round (dist A C) ∧ dist B C = round (dist B C))) + (hABCall : ∀ A B C, ABCall A B C ↔ (¬Collinear ℝ {A, B, C} ∧ ABCintcoords A B C ∧ ABCintdists A B C)) : + IsLeast {y | ∃ A B C, ABCall A B C ∧ y = dist A B} putnam_2010_b2_solution := sorry diff --git a/lean4/src/putnam_2010_b3.lean b/lean4/src/putnam_2010_b3.lean index 72158cb..cf1623b 100644 --- a/lean4/src/putnam_2010_b3.lean +++ b/lean4/src/putnam_2010_b3.lean @@ -8,10 +8,12 @@ abbrev putnam_2010_b3_solution : Set ℕ := sorry There are $2010$ boxes labeled $B_1, B_2, \dots, B_{2010}$, and $2010n$ balls have been distributed among them, for some positive integer $n$. You may redistribute the balls by a sequence of moves, each of which consists of choosing an $i$ and moving \emph{exactly} $i$ balls from box $B_i$ into any one other box. For which values of $n$ is it possible to reach the distribution with exactly $n$ balls in each box, regardless of the initial distribution of balls? -/ theorem putnam_2010_b3 -(n : ℕ) -(npos : n > 0) -(trans : (ℕ → Fin 2010 → ℕ) → ℕ → Prop) -(htrans : trans = fun P T ↦ ∀ t : ℕ, t < T → - ∃ i j : Fin 2010, i ≠ j ∧ P t i ≥ i.1 + 1 ∧ P (t + 1) i = P t i - (i.1 + 1) ∧ P (t + 1) j = P t j + (i.1 + 1) ∧ ∀ k : Fin 2010, k ≠ i → k ≠ j → P (t + 1) k = P t k) -: ((∀ B : Fin 2010 → ℕ, ∑ i : Fin 2010, B i = 2010 * n → ∃ (P : ℕ → Fin 2010 → ℕ) (T : ℕ), P 0 = B ∧ trans P T ∧ ∀ i : Fin 2010, P T i = n) ↔ n ∈ putnam_2010_b3_solution) := + (n : ℕ) (hn : n > 0) + (trans : (ℕ → Fin 2010 → ℕ) → ℕ → Prop) + (htrans : ∀ P T, trans P T ↔ ∀ t : ℕ, t < T → ∃ i j, + i ≠ j ∧ + P t i ≥ i.1 + 1 ∧ P (t + 1) i = P t i - (i.1 + 1) ∧ P (t + 1) j = P t j + (i.1 + 1) ∧ + ∀ k : Fin 2010, k ≠ i → k ≠ j → P (t + 1) k = P t k) : + (∀ B, ∑ i, B i = 2010 * n → ∃ᵉ (P) (T), P 0 = B ∧ trans P T ∧ ∀ i, P T i = n) + ↔ n ∈ putnam_2010_b3_solution := sorry diff --git a/lean4/src/putnam_2010_b5.lean b/lean4/src/putnam_2010_b5.lean index d5d6f26..1bb7f22 100644 --- a/lean4/src/putnam_2010_b5.lean +++ b/lean4/src/putnam_2010_b5.lean @@ -7,6 +7,6 @@ abbrev putnam_2010_b5_solution : Prop := sorry /-- Is there a strictly increasing function $f: \mathbb{R} \to \mathbb{R}$ such that $f'(x) = f(f(x))$ for all $x$? -/ -theorem putnam_2010_b5 -: (∃ f : ℝ → ℝ, StrictMono f ∧ Differentiable ℝ f ∧ (∀ x : ℝ, deriv f x = f (f x))) ↔ putnam_2010_b5_solution := +theorem putnam_2010_b5 : + (∃ f : ℝ → ℝ, StrictMono f ∧ Differentiable ℝ f ∧ (∀ x : ℝ, deriv f x = f (f x))) ↔ putnam_2010_b5_solution := sorry diff --git a/lean4/src/putnam_2011_a1.lean b/lean4/src/putnam_2011_a1.lean index f7ed031..76e6d4f 100644 --- a/lean4/src/putnam_2011_a1.lean +++ b/lean4/src/putnam_2011_a1.lean @@ -11,12 +11,12 @@ Define a \emph{growing spiral} in the plane to be a sequence of points with inte How many of the points $(x,y)$ with integer coordinates $0 \leq x \leq 2011,0 \leq y \leq 2011$ \emph{cannot} be the last point, $P_n$ of any growing spiral? -/ theorem putnam_2011_a1 -(isspiral : List (Fin 2 → ℤ) → Prop) -(hisspiral : isspiral = (fun P : List (Fin 2 → ℤ) => P.length ≥ 3 ∧ P[0]! = 0 ∧ -(∃ l : Fin (P.length - 1) → ℕ, l > 0 ∧ StrictMono l ∧ (∀ i : Fin (P.length - 1), -(i.1 % 4 = 0 → (P[i] 0 + l i = P[i.1 + 1]! 0 ∧ P[i] 1 = P[i.1 + 1]! 1)) ∧ -(i.1 % 4 = 1 → (P[i] 0 = P[i.1 + 1]! 0 ∧ P[i] 1 + l i = P[i.1 + 1]! 1)) ∧ -(i.1 % 4 = 2 → (P[i] 0 - l i = P[i.1 + 1]! 0 ∧ P[i] 1 = P[i.1 + 1]! 1)) ∧ -(i.1 % 4 = 3 → (P[i] 0 = P[i.1 + 1]! 0 ∧ P[i] 1 - l i = P[i.1 + 1]! 1)))))) -: {p : Fin 2 → ℤ | 0 ≤ p 0 ∧ p 0 ≤ 2011 ∧ 0 ≤ p 1 ∧ p 1 ≤ 2011 ∧ ¬∃ spiral : List (Fin 2 → ℤ), isspiral spiral ∧ spiral.getLast! = p}.encard = putnam_2011_a1_solution := + (IsSpiral : List (Fin 2 → ℤ) → Prop) + (IsSpiral_def : ∀ P, IsSpiral P ↔ P.length ≥ 3 ∧ P[0]! = 0 ∧ + (∃ l : Fin (P.length - 1) → ℕ, l > 0 ∧ StrictMono l ∧ (∀ i : Fin (P.length - 1), + (i.1 % 4 = 0 → (P[i] 0 + l i = P[i.1 + 1]! 0 ∧ P[i] 1 = P[i.1 + 1]! 1)) ∧ + (i.1 % 4 = 1 → (P[i] 0 = P[i.1 + 1]! 0 ∧ P[i] 1 + l i = P[i.1 + 1]! 1)) ∧ + (i.1 % 4 = 2 → (P[i] 0 - l i = P[i.1 + 1]! 0 ∧ P[i] 1 = P[i.1 + 1]! 1)) ∧ + (i.1 % 4 = 3 → (P[i] 0 = P[i.1 + 1]! 0 ∧ P[i] 1 - l i = P[i.1 + 1]! 1))))) : + {p | 0 ≤ p 0 ∧ p 0 ≤ 2011 ∧ 0 ≤ p 1 ∧ p 1 ≤ 2011 ∧ ¬∃ spiral, IsSpiral spiral ∧ spiral.getLast! = p}.encard = putnam_2011_a1_solution := sorry diff --git a/lean4/src/putnam_2011_b2.lean b/lean4/src/putnam_2011_b2.lean index 711bfa7..6595c29 100644 --- a/lean4/src/putnam_2011_b2.lean +++ b/lean4/src/putnam_2011_b2.lean @@ -8,10 +8,9 @@ abbrev putnam_2011_b2_solution : Set ℕ := sorry Let $S$ be the set of all ordered triples $(p,q,r)$ of prime numbers for which at least one rational number $x$ satisfies $px^2+qx+r=0$. Which primes appear in seven or more elements of $S$? -/ theorem putnam_2011_b2 -(S : Set (Fin 3 → ℕ)) -(t : ℕ) -(t7inS : Prop) -(hS : S = {s : Fin 3 → ℕ | (s 0).Prime ∧ (s 1).Prime ∧ (s 2).Prime ∧ ∃ x : ℚ, (s 0) * x ^ 2 + (s 1) * x + (s 2) = 0}) -(ht7inS : t7inS = ({s ∈ S | ∃ i : Fin 3, s i = t}.encard ≥ 7)) -: (t.Prime ∧ t7inS) ↔ t ∈ putnam_2011_b2_solution := + (S : Set (Fin 3 → ℕ)) + (t : ℕ) + (hS : S = {s : Fin 3 → ℕ | (s 0).Prime ∧ (s 1).Prime ∧ (s 2).Prime ∧ ∃ x : ℚ, (s 0) * x ^ 2 + (s 1) * x + (s 2) = 0}) : + (t.Prime ∧ ({s ∈ S | ∃ i : Fin 3, s i = t}.encard ≥ 7)) + ↔ t ∈ putnam_2011_b2_solution := sorry diff --git a/lean4/src/putnam_2011_b3.lean b/lean4/src/putnam_2011_b3.lean index 22e94b5..5403c66 100644 --- a/lean4/src/putnam_2011_b3.lean +++ b/lean4/src/putnam_2011_b3.lean @@ -7,6 +7,11 @@ abbrev putnam_2011_b3_solution : Prop := sorry /-- Let $f$ and $g$ be (real-valued) functions defined on an open interval containing $0$, with $g$ nonzero and continuous at $0$. If $fg$ and $f/g$ are differentiable at $0$, must $f$ be differentiable at $0$? -/ -theorem putnam_2011_b3 -: ((∀ f g : ℝ → ℝ, g 0 ≠ 0 → ContinuousAt g 0 → DifferentiableAt ℝ (fun x ↦ f x * g x) 0 → DifferentiableAt ℝ (fun x ↦ f x / g x) 0 → (DifferentiableAt ℝ f 0)) ↔ putnam_2011_b3_solution) := +theorem putnam_2011_b3 : + ((∀ f g : ℝ → ℝ, + g 0 ≠ 0 → ContinuousAt g 0 → + DifferentiableAt ℝ (fun x ↦ f x * g x) 0 → + DifferentiableAt ℝ (fun x ↦ f x / g x) 0 → + (DifferentiableAt ℝ f 0)) + ↔ putnam_2011_b3_solution) := sorry diff --git a/lean4/src/putnam_2011_b4.lean b/lean4/src/putnam_2011_b4.lean index 3bfdc22..f0c5d03 100644 --- a/lean4/src/putnam_2011_b4.lean +++ b/lean4/src/putnam_2011_b4.lean @@ -7,9 +7,8 @@ In a tournament, $2011$ players meet $2011$ times to play a multiplayer game. Ev -/ theorem putnam_2011_b4 (games : Fin 2011 → Fin 2011 → Bool) -(T : Matrix (Fin 2011) (Fin 2011) ℂ) -(hT : T = (fun p1 p2 : Fin 2011 => ({g : Fin 2011 | games g p1 = games g p2}.ncard : ℂ))) -(W : Matrix (Fin 2011) (Fin 2011) ℂ) -(hW : W = (fun p1 p2 : Fin 2011 => ({g : Fin 2011 | games g p1 ∧ !games g p2}.ncard - {g : Fin 2011 | !games g p1 ∧ games g p2}.ncard : ℂ))) +(T W : Matrix (Fin 2011) (Fin 2011) ℂ) +(hT : T = (fun p1 p2 => ({g | games g p1 = games g p2}.ncard : ℂ))) +(hW : W = (fun p1 p2 => ({g | games g p1 ∧ !games g p2}.ncard - {g | !games g p1 ∧ games g p2}.ncard : ℂ))) : ∃ n : ℕ, (T + Complex.I • W).det = n ∧ (2 ^ 2010) ∣ n := sorry diff --git a/lean4/src/putnam_2011_b5.lean b/lean4/src/putnam_2011_b5.lean index 6dd47c9..9f744b0 100644 --- a/lean4/src/putnam_2011_b5.lean +++ b/lean4/src/putnam_2011_b5.lean @@ -13,7 +13,8 @@ Prove there is a constant $B>0$ such that for all $n$, \] -/ theorem putnam_2011_b5 -(a : ℕ → ℝ) -(h : ∃ A : ℝ, ∀ n : ℕ, ∫ x : ℝ, ((∑ i : Finset.range n, 1 / (1 + (x - a i) ^ 2)) ^ 2) ≤ A * n) -: (∃ B : ℝ, B > 0 ∧ ∀ n : ℕ, ∑' i : Finset.range n, ∑' j : Finset.range n, (1 + (a i - a j) ^ 2) ≥ B * n ^ 3) := + (a : ℕ → ℝ) + (h : ∃ A : ℝ, ∀ n : ℕ, ∫ x : ℝ, ((∑ i : Finset.range n, 1 / (1 + (x - a i) ^ 2)) ^ 2) ≤ A * n) : + ∃ B : ℝ, B > 0 ∧ + ∀ n : ℕ, ∑' i : Finset.range n, ∑' j : Finset.range n, (1 + (a i - a j) ^ 2) ≥ B * n ^ 3 := sorry diff --git a/lean4/src/putnam_2012_a1.lean b/lean4/src/putnam_2012_a1.lean index 0c48a46..d85dfa4 100644 --- a/lean4/src/putnam_2012_a1.lean +++ b/lean4/src/putnam_2012_a1.lean @@ -8,6 +8,6 @@ Let $d_1,d_2,\dots,d_{12}$ be real numbers in the open interval $(1,12)$. Show t -/ theorem putnam_2012_a1 (d : Fin 12 → ℝ) -(hd : ∀ i : Fin 12, d i ∈ Set.Ioo 1 12) -: ∃ i j k : Fin 12, i ≠ j ∧ i ≠ k ∧ j ≠ k ∧ d k ≥ d i ∧ d k ≥ d j ∧ (d i) ^ 2 + (d j) ^ 2 > (d k) ^ 2 := +(hd : ∀ i, d i ∈ Set.Ioo 1 12) +: ∃ i j k, i ≠ j ∧ i ≠ k ∧ j ≠ k ∧ d k ≥ d i ∧ d k ≥ d j ∧ (d i) ^ 2 + (d j) ^ 2 > (d k) ^ 2 := sorry diff --git a/lean4/src/putnam_2012_a5.lean b/lean4/src/putnam_2012_a5.lean index 8c20607..f770b97 100644 --- a/lean4/src/putnam_2012_a5.lean +++ b/lean4/src/putnam_2012_a5.lean @@ -8,15 +8,14 @@ abbrev putnam_2012_a5_solution : Set (ℕ × ℕ) := sorry Let $\FF_p$ denote the field of integers modulo a prime $p$, and let $n$ be a positive integer. Let $v$ be a fixed vector in $\FF_p^n$, let $M$ be an $n \times n$ matrix with entries of $\FF_p$, and define $G: \FF_p^n \to \FF_p^n$ by $G(x) = v + Mx$. Let $G^{(k)}$ denote the $k$-fold composition of $G$ with itself, that is, $G^{(1)}(x) = G(x)$ and $G^{(k+1)}(x) = G(G^{(k)}(x))$. Determine all pairs $p, n$ for which there exist $v$ and $M$ such that the $p^n$ vectors $G^{(k)}(0)$, $k=1,2,\dots,p^n$ are distinct. -/ theorem putnam_2012_a5 -(n p : ℕ) -(hn : n > 0) -(hp : Nat.Prime p) -{F : Type*} [Field F] [Fintype F] -(hK : Fintype.card F = p) -(G : Matrix (Fin n) (Fin n) F → (Fin n → F) → (Fin n → F) → (Fin n → F)) -(hG : ∀ M : Matrix (Fin n) (Fin n) F, ∀ v x : (Fin n → F), G M v x = v + mulVec M x) -: (n, p) ∈ putnam_2012_a5_solution ↔ -∃ M : Matrix (Fin n) (Fin n) F, -∃ v : (Fin n → F), -¬(∃ i j : Finset.range (p^n), i ≠ j ∧ (G M v)^[i + 1] 0 = (G M v)^[j + 1] 0) := -sorry + (n p : ℕ) + (hn : n > 0) + (hp : Nat.Prime p) + {F : Type*} [Field F] [Fintype F] + (hK : Fintype.card F = p) + (G : Matrix (Fin n) (Fin n) F → (Fin n → F) → (Fin n → F) → (Fin n → F)) + (hG : ∀ M v x, G M v x = v + mulVec M x) : + (n, p) ∈ putnam_2012_a5_solution ↔ + ∃ᵉ (M : Matrix (Fin n) (Fin n) F) (v : (Fin n → F)), + ¬(∃ i j : Finset.range (p^n), i ≠ j ∧ (G M v)^[i + 1] 0 = (G M v)^[j + 1] 0) := + sorry diff --git a/lean4/src/putnam_2012_a6.lean b/lean4/src/putnam_2012_a6.lean index ee9cdfe..2b6bdfa 100644 --- a/lean4/src/putnam_2012_a6.lean +++ b/lean4/src/putnam_2012_a6.lean @@ -11,6 +11,9 @@ Must $f(x,y)$ be identically $0$? -/ theorem putnam_2012_a6 (p : ((ℝ × ℝ) → ℝ) → Prop) -(hp : p = fun f ↦ Continuous f ∧ ∀ x1 x2 y1 y2 : ℝ, x2 > x1 → y2 > y1 → (x2 - x1) * (y2 - y1) = 1 → ∫ x in x1..x2, ∫ y in y1..y2, f (x, y) = 0) -: ((∀ f : (ℝ × ℝ) → ℝ, ∀ x y : ℝ, p f → f (x, y) = 0) ↔ putnam_2012_a6_solution) := +(hp : ∀ f, p f ↔ + Continuous f ∧ + ∀ x1 x2 y1 y2 : ℝ, x2 > x1 → y2 > y1 + → (x2 - x1) * (y2 - y1) = 1 → ∫ x in x1..x2, ∫ y in y1..y2, f (x, y) = 0) +: ((∀ f x y, p f → f (x, y) = 0) ↔ putnam_2012_a6_solution) := sorry diff --git a/lean4/src/putnam_2012_b3.lean b/lean4/src/putnam_2012_b3.lean index 21a610d..a43ace4 100644 --- a/lean4/src/putnam_2012_b3.lean +++ b/lean4/src/putnam_2012_b3.lean @@ -9,10 +9,13 @@ A round-robin tournament of $2n$ teams lasted for $2n-1$ days, as follows. On ea -/ theorem putnam_2012_b3 (nmatchupsgames nmatchupsall : (n : ℕ) → (Fin (2 * n - 1) → (Fin (2 * n) → Fin (2 * n))) → Prop) -(hnmatchupsall : nmatchupsall = (fun (n : ℕ) (matchups : Fin (2 * n - 1) → (Fin (2 * n) → Fin (2 * n))) => ∀ t1 t2 : Fin (2 * n), t1 ≠ t2 → (∃ d : Fin (2 * n - 1), matchups d t1 = t2))) -(hnmatchupsgames : nmatchupsgames = (fun (n : ℕ) (matchups : Fin (2 * n - 1) → (Fin (2 * n) → Fin (2 * n))) => ∀ d : Fin (2 * n - 1), ∀ t : Fin (2 * n), matchups d t ≠ t ∧ matchups d (matchups d t) = t)) (nmatchupswins nmatchupswinschoices: (n : ℕ) → (Fin (2 * n - 1) → (Fin (2 * n) → Fin (2 * n))) → (Fin (2 * n - 1) → (Fin (2 * n) → Bool)) → Prop) -(hnmatchupswins : nmatchupswins = (fun (n : ℕ) (matchups : Fin (2 * n - 1) → (Fin (2 * n) → Fin (2 * n))) (wins : Fin (2 * n - 1) → (Fin (2 * n) → Bool)) => ∀ d : Fin (2 * n - 1), ∀ t : Fin (2 * n), wins d t = !(wins d (matchups d t)))) -(hnmatchupswinschoices : nmatchupswinschoices = (fun (n : ℕ) (matchups : Fin (2 * n - 1) → (Fin (2 * n) → Fin (2 * n))) (wins : Fin (2 * n - 1) → (Fin (2 * n) → Bool)) => ∃ choices : Fin (2 * n - 1) → Fin (2 * n), (∀ d : Fin (2 * n - 1), wins d (choices d)) ∧ Function.Injective choices)) -: (∀ n ≥ 1, ∀ (matchups : Fin (2 * n - 1) → (Fin (2 * n) → Fin (2 * n))) (wins : Fin (2 * n - 1) → (Fin (2 * n) → Bool)), (nmatchupsgames n matchups ∧ nmatchupsall n matchups ∧ nmatchupswins n matchups wins) → nmatchupswinschoices n matchups wins) ↔ putnam_2012_b3_solution := +(hnmatchupsall : ∀ n matchups, nmatchupsall n matchups ↔ ∀ t1 t2, t1 ≠ t2 → (∃ d, matchups d t1 = t2)) +(hnmatchupsgames : ∀ n matchups, nmatchupsgames n matchups ↔ ∀ d, ∀ t, matchups d t ≠ t ∧ matchups d (matchups d t) = t) +(hnmatchupswins : ∀ n matchups wins, nmatchupswins n matchups wins ↔ ∀ d t, wins d t = !(wins d (matchups d t))) +(hnmatchupswinschoices : ∀ n matchups wins, nmatchupswinschoices n matchups wins ↔ ∃ choices, (∀ d, wins d (choices d)) ∧ Function.Injective choices) : + (∀ n ≥ 1, ∀ matchups wins, + (nmatchupsgames n matchups ∧ nmatchupsall n matchups ∧ nmatchupswins n matchups wins) → + nmatchupswinschoices n matchups wins) + ↔ putnam_2012_b3_solution := sorry diff --git a/lean4/src/putnam_2013_a2.lean b/lean4/src/putnam_2013_a2.lean index 66cd872..521c168 100644 --- a/lean4/src/putnam_2013_a2.lean +++ b/lean4/src/putnam_2013_a2.lean @@ -12,14 +12,16 @@ $2 \cdot 5$, $2 \cdot 3 \cdot 4$, $2 \cdot 3 \cdot 5$, $2 \cdot 4 \cdot 5$, and Show that the function $f$ from $S$ to the integers is one-to-one. -/ theorem putnam_2013_a2 -(S : Set ℤ) -(hS : S = {n : ℤ | n > 0 ∧ ¬∃ m : ℤ, m ^ 2 = n}) -(P : ℤ → List ℤ → Prop) -(hP : P = fun n : ℤ => fun a : List ℤ => a.length > 0 ∧ n < a[0]! ∧ -(∃ m : ℤ, m ^ 2 = n * a.prod) ∧ (∀ i : Fin (a.length - 1), a[i] < a[i+(1:ℕ)])) -(T : ℤ → Set ℤ) -(hT : T = fun n : ℤ => {m : ℤ | ∃ a : List ℤ, P n a ∧ a[a.length - 1]! = m}) -(f : ℤ → ℤ) -(hf : ∀ n ∈ S, ((∃ r ∈ T n, f n = r) ∧ ∀ r ∈ T n, f n ≤ r)) -: InjOn f S := + (S : Set ℤ) + (hS : S = {n : ℤ | n > 0 ∧ ¬∃ m : ℤ, m ^ 2 = n}) + (P : ℤ → List ℤ → Prop) + (hP : ∀ n a, P n a ↔ + a.length > 0 ∧ n < a[0]! ∧ + (∃ m : ℤ, m ^ 2 = n * a.prod) ∧ + (∀ i : Fin (a.length - 1), a[i] < a[i+(1:ℕ)])) + (T : ℤ → Set ℤ) + (hT : T = fun n : ℤ => {m : ℤ | ∃ a : List ℤ, P n a ∧ a[a.length - 1]! = m}) + (f : ℤ → ℤ) + (hf : ∀ n ∈ S, ((∃ r ∈ T n, f n = r) ∧ ∀ r ∈ T n, f n ≤ r)) : + InjOn f S := sorry diff --git a/lean4/src/putnam_2013_a4.lean b/lean4/src/putnam_2013_a4.lean index 3867ce5..8c50fcf 100644 --- a/lean4/src/putnam_2013_a4.lean +++ b/lean4/src/putnam_2013_a4.lean @@ -6,18 +6,16 @@ open Function Set A finite collection of digits $0$ and $1$ is written around a circle. An \emph{arc} of length $L \geq 0$ consists of $L$ consecutive digits around the circle. For each arc $w$, let $Z(w)$ and $N(w)$ denote the number of $0$'s in $w$ and the number of $1$'s in $w$, respectively. Assume that $|Z(w)-Z(w')| \leq 1$ for any two arcs $w,w'$ of the same length. Suppose that some arcs $w_1,\dots,w_k$ have the property that $Z=\frac{1}{k}\sum_{j=1}^k Z(w_j)$ and $N=\frac{1}{k}\sum_{j=1}^k N(w_j)$ are both integers. Prove that there exists an arc $w$ with $Z(w)=Z$ and $N(w)=N$. -/ theorem putnam_2013_a4 -(n : ℕ) +(n k : ℕ) (circle : Fin n → Fin 2) (Z N : Fin n × Fin (n + 1) → ℤ) -(k : ℕ) (ws : Fin k → Fin n × Fin (n + 1)) (Zsum Nsum : ℤ) -(npos : n ≥ 1) -(hZ : ∀ w : Fin n × Fin (n + 1), Z w = ∑ l : {x : Fin n | x < w.2}, if (circle (w.1 + l) = 0) then 1 else 0) -(hN : ∀ w : Fin n × Fin (n + 1), N w = ∑ l : {x : Fin n | x < w.2}, if (circle (w.1 + l) = 1) then 1 else 0) -(Zle1 : ∀ w w' : Fin n × Fin (n + 1), w.2 = w'.2 → |(Z w : ℤ) - Z w'| ≤ 1) -(kpos : k ≥ 1) +(npos : n ≥ 1) (kpos : k ≥ 1) +(hZ : ∀ w, Z w = ∑ l : {x : Fin n | x < w.2}, if (circle (w.1 + l) = 0) then 1 else 0) +(hN : ∀ w, N w = ∑ l : {x : Fin n | x < w.2}, if (circle (w.1 + l) = 1) then 1 else 0) +(Zle1 : ∀ w w', w.2 = w'.2 → |(Z w : ℤ) - Z w'| ≤ 1) (hZsum : Zsum = ((1 : ℝ) / k) * ∑ j : Fin k, Z (ws j)) -(hNsum : Nsum = ((1 : ℝ) / k) * ∑ j : Fin k, N (ws j)) -: ∃ w : Fin n × Fin (n + 1), Z w = Zsum ∧ N w = Nsum := +(hNsum : Nsum = ((1 : ℝ) / k) * ∑ j : Fin k, N (ws j)) : +∃ w, Z w = Zsum ∧ N w = Nsum := sorry diff --git a/lean4/src/putnam_2013_a5.lean b/lean4/src/putnam_2013_a5.lean index e3d2baa..a47ab89 100644 --- a/lean4/src/putnam_2013_a5.lean +++ b/lean4/src/putnam_2013_a5.lean @@ -13,9 +13,9 @@ theorem putnam_2013_a5 (areadef2 : (Fin m → Fin m → Fin m → ℝ) → Prop) (areadef3 : (Fin m → Fin m → Fin m → ℝ) → Prop) (mge3 : m ≥ 3) -(harea2 : ∀ a b c : Fin 2 → ℝ, area2 a b c = (volume (convexHull ℝ {a, b, c})).toReal) -(harea3 : ∀ a b c : Fin 3 → ℝ, area3 a b c = (μH[2] (convexHull ℝ {a, b, c})).toReal) -(hareadef2 : ∀ a : Fin m → Fin m → Fin m → ℝ, areadef2 a = ∀ A : Fin m → (Fin 2 → ℝ), (∑ i : Fin m, ∑ j : Fin m, ∑ k : Fin m, if (i < j ∧ j < k) then (a i j k * area2 (A i) (A j) (A k)) else 0) ≥ 0) -(hareadef3 : ∀ a : Fin m → Fin m → Fin m → ℝ, areadef3 a = ∀ A : Fin m → (Fin 3 → ℝ), (∑ i : Fin m, ∑ j : Fin m, ∑ k : Fin m, if (i < j ∧ j < k) then (a i j k * area3 (A i) (A j) (A k)) else 0) ≥ 0) -: ∀ a : Fin m → Fin m → Fin m → ℝ, areadef2 a → areadef3 a := +(harea2 : ∀ a b c, area2 a b c = (volume (convexHull ℝ {a, b, c})).toReal) +(harea3 : ∀ a b c, area3 a b c = (μH[2] (convexHull ℝ {a, b, c})).toReal) +(hareadef2 : ∀ a, areadef2 a ↔ ∀ A : Fin m → (Fin 2 → ℝ), (∑ i : Fin m, ∑ j : Fin m, ∑ k : Fin m, if (i < j ∧ j < k) then (a i j k * area2 (A i) (A j) (A k)) else 0) ≥ 0) +(hareadef3 : ∀ a, areadef3 a ↔ ∀ A : Fin m → (Fin 3 → ℝ), (∑ i : Fin m, ∑ j : Fin m, ∑ k : Fin m, if (i < j ∧ j < k) then (a i j k * area3 (A i) (A j) (A k)) else 0) ≥ 0) +: ∀ a, areadef2 a → areadef3 a := sorry diff --git a/lean4/src/putnam_2013_a6.lean b/lean4/src/putnam_2013_a6.lean index 2213510..5702789 100644 --- a/lean4/src/putnam_2013_a6.lean +++ b/lean4/src/putnam_2013_a6.lean @@ -21,15 +21,15 @@ $a$ & 0 & 2 & -4 & 12 & -4 & 2 \\ For every finite subset $S$ of $\mathbb{Z} \times \mathbb{Z}$, define $A(S)=\sum_{(\mathbf{s},\mathbf{s}') \in S \times S} w(\mathbf{s}-\mathbf{s}')$. Prove that if $S$ is any finite nonempty subset of $\mathbb{Z} \times \mathbb{Z}$, then $A(S)>0$. (For example, if $S=\{(0,1),(0,2),(2,0),(3,1)\}$, then the terms in $A(S)$ are $12,12,12,12,4,4,0,0,0,0,-1,-1,-2,-2,-4,-4$.) -/ theorem putnam_2013_a6 -(w : ℤ → ℤ → ℤ) -(A : Finset (ℤ × ℤ) → ℤ) -(hwn1 : w (-2) (-2) = -1 ∧ w 2 (-2) = -1 ∧ w (-2) 2 = -1 ∧ w 2 2 = -1) -(hwn2 : w (-1) (-2) = -2 ∧ w 1 (-2) = -2 ∧ w (-2) (-1) = -2 ∧ w 2 (-1) = -2 ∧ w (-2) 1 = -2 ∧ w 2 1 = -2 ∧ w (-1) 2 = -2 ∧ w 1 2 = -2) -(hw2 : w 0 (-2) = 2 ∧ w (-2) 0 = 2 ∧ w 2 0 = 2 ∧ w 0 2 = 2) -(hw4 : w (-1) (-1) = 4 ∧ w 1 (-1) = 4 ∧ w (-1) 1 = 4 ∧ w 1 1 = 4) -(hwn4 : w 0 (-1) = -4 ∧ w (-1) 0 = -4 ∧ w 1 0 = -4 ∧ w 0 1 = -4) -(hw12 : w 0 0 = 12) -(hw0 : ∀ a b : ℤ, (|a| > 2 ∨ |b| > 2) → w a b = 0) -(hA : ∀ S : Finset (ℤ × ℤ), A S = ∑ s in S, ∑ s' in S, w (s - s').1 (s - s').2) -: ∀ S : Finset (ℤ × ℤ), Nonempty S → A S > 0 := + (w : ℤ → ℤ → ℤ) + (A : Finset (ℤ × ℤ) → ℤ) + (hwn1 : w (-2) (-2) = -1 ∧ w 2 (-2) = -1 ∧ w (-2) 2 = -1 ∧ w 2 2 = -1) + (hwn2 : w (-1) (-2) = -2 ∧ w 1 (-2) = -2 ∧ w (-2) (-1) = -2 ∧ w 2 (-1) = -2 ∧ w (-2) 1 = -2 ∧ w 2 1 = -2 ∧ w (-1) 2 = -2 ∧ w 1 2 = -2) + (hw2 : w 0 (-2) = 2 ∧ w (-2) 0 = 2 ∧ w 2 0 = 2 ∧ w 0 2 = 2) + (hw4 : w (-1) (-1) = 4 ∧ w 1 (-1) = 4 ∧ w (-1) 1 = 4 ∧ w 1 1 = 4) + (hwn4 : w 0 (-1) = -4 ∧ w (-1) 0 = -4 ∧ w 1 0 = -4 ∧ w 0 1 = -4) + (hw12 : w 0 0 = 12) + (hw0 : ∀ a b : ℤ, (|a| > 2 ∨ |b| > 2) → w a b = 0) + (hA : ∀ S, A S = ∑ s in S, ∑ s' in S, w (s - s').1 (s - s').2) : + ∀ S : Finset (ℤ × ℤ), Nonempty S → A S > 0 := sorry diff --git a/lean4/src/putnam_2013_b2.lean b/lean4/src/putnam_2013_b2.lean index 878e99d..b3fa49b 100644 --- a/lean4/src/putnam_2013_b2.lean +++ b/lean4/src/putnam_2013_b2.lean @@ -20,11 +20,11 @@ Determine the maximum value of $f(0)$ as $f$ ranges through $C$, and prove that this maximum is attained. -/ theorem putnam_2013_b2 -(CN : ℕ → Set (ℝ → ℝ)) -(C : Set (ℝ → ℝ)) -(hCN : CN = fun N : ℕ => {f : ℝ → ℝ | (∀ x : ℝ, f x ≥ 0) ∧ -∃ a : List ℝ, a.length = N + 1 ∧ (∀ n : Fin (N + 1), 3 ∣ (n : ℕ) → a[n]! = 0) ∧ -∀ x : ℝ, f x = 1 + ∑ n in Finset.Icc 1 N, a[(n : ℕ)]! * Real.cos (2*Real.pi*n*x)}) -(hC : C = ⋃ N ∈ Ici 1, CN N) -: (∀ f ∈ C, f 0 ≤ putnam_2013_b2_solution) ∧ ∃ f ∈ C, f 0 = putnam_2013_b2_solution := + (CN : ℕ → Set (ℝ → ℝ)) + (hCN : ∀ N : ℕ, CN N = + {f : ℝ → ℝ | + (∀ x : ℝ, f x ≥ 0) ∧ + ∃ a : List ℝ, a.length = N + 1 ∧ (∀ n : Fin (N + 1), 3 ∣ (n : ℕ) → a[n]! = 0) ∧ + ∀ x : ℝ, f x = 1 + ∑ n in Finset.Icc 1 N, a[(n : ℕ)]! * Real.cos (2*Real.pi*n*x)}) : + IsGreatest {f 0 | f ∈ ⋃ N ∈ Ici 1, CN N} putnam_2013_b2_solution := sorry diff --git a/lean4/src/putnam_2013_b5.lean b/lean4/src/putnam_2013_b5.lean index 3b07ae19..20f9cc9 100644 --- a/lean4/src/putnam_2013_b5.lean +++ b/lean4/src/putnam_2013_b5.lean @@ -6,10 +6,9 @@ open Function Set Let $X=\{1,2,\dots,n\}$, and let $k \in X$. Show that there are exactly $k \cdot n^{n-1}$ functions $f:X \to X$ such that for every $x \in X$ there is a $j \geq 0$ such that $f^{(j)}(x) \leq k$. [Here $f^{(j)}$ denotes the $j$\textsuperscript{th} iterate of $f$, so that $f^{(0)}(x)=x$ and $f^{(j+1)}(x)=f(f^{(j)}(x))$.] -/ theorem putnam_2013_b5 -(n : ℕ) -(k : Set.Icc 1 n) -(fiter : (Set.Icc 1 n → Set.Icc 1 n) → Prop) -(npos : n ≥ 1) -(hfiter : ∀ f : Set.Icc 1 n → Set.Icc 1 n, fiter f = ∀ x : Set.Icc 1 n, ∃ j : ℕ, f^[j] x ≤ k) -: {f : Set.Icc 1 n → Set.Icc 1 n | fiter f}.encard = k * n ^ (n - 1) := + (n : ℕ) (hn : n ≥ 1) + (k : Set.Icc 1 n) + (fiter : (Set.Icc 1 n → Set.Icc 1 n) → Prop) + (hfiter : ∀ f, fiter f ↔ ∀ x : Set.Icc 1 n, ∃ j : ℕ, f^[j] x ≤ k) : + {f | fiter f}.encard = k * n ^ (n - 1) := sorry diff --git a/lean4/src/putnam_2014_b2.lean b/lean4/src/putnam_2014_b2.lean index 4ae0bd7..7d2c42e 100644 --- a/lean4/src/putnam_2014_b2.lean +++ b/lean4/src/putnam_2014_b2.lean @@ -8,11 +8,10 @@ noncomputable abbrev putnam_2014_b2_solution : ℝ := sorry /-- Suppose that \( f \) is a function on the interval \([1,3]\) such that \(-1 \leq f(x) \leq 1\) for all \( x \) and \( \int_{1}^{3} f(x) \, dx = 0 \). How large can \(\int_{1}^{3} \frac{f(x)}{x} \, dx \) be? -/ -theorem putnam_2014_b2 -(fbound finteq0 : (ℝ → ℝ) → Prop) -(fint : (ℝ → ℝ) → ℝ) -(hfbound : fbound = (fun f : ℝ → ℝ => ∀ x : Set.Icc (1 : ℝ) 3, -1 ≤ f x ∧ f x ≤ 1)) -(hfinteq0 : finteq0 = (fun f : ℝ → ℝ => (∫ x in Set.Ioo 1 3, f x) = 0)) -(hfint : fint = (fun f : ℝ → ℝ => ∫ x in Set.Ioo 1 3, (f x) / x)) -: (∃ f : ℝ → ℝ, fbound f ∧ finteq0 f ∧ fint f = putnam_2014_b2_solution) ∧ (∀ f : ℝ → ℝ, (fbound f ∧ finteq0 f) → fint f ≤ putnam_2014_b2_solution) := +theorem putnam_2014_b2 : + IsGreatest {t | ∃ f : ℝ → ℝ, + (∀ x : Set.Icc (1 : ℝ) 3, -1 ≤ f x ∧ f x ≤ 1) ∧ + (∫ x in Set.Ioo 1 3, f x = 0) ∧ + (∫ x in Set.Ioo 1 3, (f x) / x) = t} + putnam_2014_b2_solution := sorry diff --git a/lean4/src/putnam_2014_b3.lean b/lean4/src/putnam_2014_b3.lean index bccddf9..acbbe14 100644 --- a/lean4/src/putnam_2014_b3.lean +++ b/lean4/src/putnam_2014_b3.lean @@ -8,7 +8,7 @@ Let $A$ be an $m \times n$ matrix with rational entries. Suppose that there are theorem putnam_2014_b3 (m n : ℕ) (A : Matrix (Fin m) (Fin n) ℚ) -(mnpos : m > 0 ∧ n > 0) +(mnpos : 0 < m ∧ 0 < n) (Aprime : {p : ℕ | p.Prime ∧ ∃ (i : Fin m) (j : Fin n), |A i j| = p}.encard ≥ m + n) : A.rank ≥ 2 := sorry diff --git a/lean4/src/putnam_2015_a1.lean b/lean4/src/putnam_2015_a1.lean index 039d283..b65acde 100644 --- a/lean4/src/putnam_2015_a1.lean +++ b/lean4/src/putnam_2015_a1.lean @@ -4,13 +4,12 @@ import Mathlib Let $A$ and $B$ be points on the same branch of the hyperbola $xy=1$. Suppose that $P$ is a point lying between $A$ and $B$ on this hyperbola, such that the area of the triangle $APB$ is as large as possible. Show that the region bounded by the hyperbola and the chord $AP$ has the same area as the region bounded by the hyperbola and the chord $PB$. -/ theorem putnam_2015_a1 -(hyperbola : Set (Fin 2 → ℝ)) -(hhyperbola : hyperbola = {p : Fin 2 → ℝ | p 1 = 1 / p 0 ∧ p 0 > 0}) -(A B : Fin 2 → ℝ) -(P : Fin 2 → ℝ) -(PPline : (Fin 2 → ℝ) → (Fin 2 → ℝ) → (ℝ → ℝ)) -(hAB : A ∈ hyperbola ∧ B ∈ hyperbola ∧ A 0 < B 0) -(hP : P ∈ hyperbola ∧ A 0 < P 0 ∧ P 0 < B 0 ∧ (∀ P' : Fin 2 → ℝ, (P' ∈ hyperbola ∧ A 0 < P' 0 ∧ P' 0 < B 0) → MeasureTheory.volume (convexHull ℝ {A, P', B}) ≤ MeasureTheory.volume (convexHull ℝ {A, P, B}))) -(hPPline : ∀ P1 P2 : Fin 2 → ℝ, P1 0 ≠ P2 0 → PPline P1 P2 = (fun x : ℝ => ((P2 1 - P1 1) / (P2 0 - P1 0)) * (x - P1 0) + P1 1)) -: ∫ x in Set.Ioo (A 0) (P 0), (PPline A P) x - 1 / x = ∫ x in Set.Ioo (P 0) (B 0), (PPline P B) x - 1 / x := + (hyperbola : Set (Fin 2 → ℝ)) + (hhyperbola : hyperbola = {p | p 1 = 1 / p 0 ∧ p 0 > 0}) + (A B P : Fin 2 → ℝ) + (PPline : (Fin 2 → ℝ) → (Fin 2 → ℝ) → (ℝ → ℝ)) + (hAB : A ∈ hyperbola ∧ B ∈ hyperbola ∧ A 0 < B 0) + (hP : P ∈ hyperbola ∧ A 0 < P 0 ∧ P 0 < B 0 ∧ (∀ P', (P' ∈ hyperbola ∧ A 0 < P' 0 ∧ P' 0 < B 0) → MeasureTheory.volume (convexHull ℝ {A, P', B}) ≤ MeasureTheory.volume (convexHull ℝ {A, P, B}))) + (hPPline : ∀ P1 P2, P1 0 ≠ P2 0 → PPline P1 P2 = (fun x : ℝ => ((P2 1 - P1 1) / (P2 0 - P1 0)) * (x - P1 0) + P1 1)) : + ∫ x in Set.Ioo (A 0) (P 0), (PPline A P) x - 1 / x = ∫ x in Set.Ioo (P 0) (B 0), (PPline P B) x - 1 / x := sorry diff --git a/lean4/src/putnam_2015_a3.lean b/lean4/src/putnam_2015_a3.lean index d0e4739..135c75d 100644 --- a/lean4/src/putnam_2015_a3.lean +++ b/lean4/src/putnam_2015_a3.lean @@ -5,6 +5,6 @@ abbrev putnam_2015_a3_solution : ℂ := sorry /-- Compute $\log_2 \left( \prod_{a=1}^{2015}\prod_{b=1}^{2015}(1+e^{2\pi iab/2015}) \right)$. Here $i$ is the imaginary unit (that is, $i^2=-1$). -/ -theorem putnam_2015_a3 -: Complex.log (∏ a : Fin 2015, ∏ b : Fin 2015, (1 + Complex.exp (2 * Real.pi * Complex.I * (a.1 + 1) * (b.1 + 1) / 2015))) / Complex.log 2 = putnam_2015_a3_solution := +theorem putnam_2015_a3 : + Complex.log (∏ a : Fin 2015, ∏ b : Fin 2015, (1 + Complex.exp (2 * Real.pi * Complex.I * (a.1 + 1) * (b.1 + 1) / 2015))) / Complex.log 2 = putnam_2015_a3_solution := sorry diff --git a/lean4/src/putnam_2015_a4.lean b/lean4/src/putnam_2015_a4.lean index 939ed0a..7744628 100644 --- a/lean4/src/putnam_2015_a4.lean +++ b/lean4/src/putnam_2015_a4.lean @@ -15,9 +15,6 @@ theorem putnam_2015_a4 (p : ℝ → Prop) (hS : S = fun (x : ℝ) ↦ {n : ℤ | n > 0 ∧ Even ⌊n * x⌋}) (hf : f = fun (x : ℝ) ↦ ∑' n : S x, 1 / 2 ^ (n : ℤ)) -(hp : p = fun (l : ℝ) ↦ ∀ x ∈ Set.Ico 0 1, f x ≥ l) -(L : ℝ) -(hpL : p L) -(hLub : ∀ l : ℝ, p l → l ≤ L) -: (L = putnam_2015_a4_solution) := +(hp : ∀ l, p l ↔ ∀ x ∈ Set.Ico 0 1, f x ≥ l) +: IsGreatest p putnam_2015_a4_solution := sorry diff --git a/lean4/src/putnam_2015_b3.lean b/lean4/src/putnam_2015_b3.lean index af89ce7..d94fbc1 100644 --- a/lean4/src/putnam_2015_b3.lean +++ b/lean4/src/putnam_2015_b3.lean @@ -6,10 +6,8 @@ abbrev putnam_2015_b3_solution : Set (Matrix (Fin 2) (Fin 2) ℝ) := sorry Let $S$ be the set of all $2 \times 2$ real matrices $M=\begin{pmatrix} a & b \\ c & d \end{pmatrix}$ whose entries $a,b,c,d$ (in that order) form an arithmetic progression. Find all matrices $M$ in $S$ for which there is some integer $k>1$ such that $M^k$ is also in $S$. -/ theorem putnam_2015_b3 -(S : Set (Matrix (Fin 2) (Fin 2) ℝ)) -(M : Matrix (Fin 2) (Fin 2) ℝ) -(MinS : Prop) -(hS : S = {M' : Matrix (Fin 2) (Fin 2) ℝ | (M' 0 1 - M' 0 0 = M' 1 0 - M' 0 1) ∧ (M' 1 0 - M' 0 1 = M' 1 1 - M' 1 0)}) -(hMinS : MinS = (M ∈ S ∧ (∃ k > 1, M ^ k ∈ S))) -: MinS ↔ M ∈ putnam_2015_b3_solution := + (M : Matrix (Fin 2) (Fin 2) ℝ) + (S : Set (Matrix (Fin 2) (Fin 2) ℝ)) + (hS : S = {M' | (M' 0 1 - M' 0 0 = M' 1 0 - M' 0 1) ∧ (M' 1 0 - M' 0 1 = M' 1 1 - M' 1 0)}) : + (M ∈ S ∧ (∃ k > 1, M ^ k ∈ S)) ↔ M ∈ putnam_2015_b3_solution := sorry diff --git a/lean4/src/putnam_2016_a1.lean b/lean4/src/putnam_2016_a1.lean index c834268..74a53e5 100644 --- a/lean4/src/putnam_2016_a1.lean +++ b/lean4/src/putnam_2016_a1.lean @@ -7,6 +7,6 @@ abbrev putnam_2016_a1_solution : ℕ := sorry /-- Find the smallest positive integer $j$ such that for every polynomial $p(x)$ with integer coefficients and for every integer $k$, the integer \[ p^{(j)}(k) = \left. \frac{d^j}{dx^j} p(x) \right|_{x=k} \] (the $j$-th derivative of $p(x)$ at $k$) is divisible by 2016. -/ -theorem putnam_2016_a1 -: (∀ j : ℕ+, (∀ P : ℤ[X], ∀ k : ℤ, 2016 ∣ (derivative^[j] P).eval k) → j ≥ putnam_2016_a1_solution) ∧ (∀ P : ℤ[X], ∀ k : ℤ, 2016 ∣ (derivative^[putnam_2016_a1_solution] P).eval k) := +theorem putnam_2016_a1 : + IsLeast {j : ℕ | 0 < j ∧ ∀ P : ℤ[X], ∀ k : ℤ, 2016 ∣ (derivative^[j] P).eval k} putnam_2016_a1_solution := sorry diff --git a/lean4/src/putnam_2016_a6.lean b/lean4/src/putnam_2016_a6.lean index fa9606b..6781fc8 100644 --- a/lean4/src/putnam_2016_a6.lean +++ b/lean4/src/putnam_2016_a6.lean @@ -11,13 +11,10 @@ Find the smallest constant $C$ such that for every real polynomial $P(x)$ of deg \] -/ theorem putnam_2016_a6 -(C : ℝ) -(max : Polynomial ℝ → ℝ) -(hmax : ∀ P : Polynomial ℝ, ∃ x ∈ Icc 0 1, |P.eval x| = max P) -(hmaxub : ∀ P : Polynomial ℝ, ∀ x ∈ Icc 0 1, |P.eval x| ≤ max P) -(p : ℝ → Prop) -(hp : p = fun c ↦ ∀ P : Polynomial ℝ, P.degree = 3 → (∃ x ∈ Icc 0 1, P.eval x = 0) → ∫ x in (0)..1, |P.eval x| ≤ c * max P) -(hpC : p C) -(hClb : ∀ c : ℝ, p c → C ≤ c) -: (C = putnam_2016_a6_solution) := + (p : ℝ → Prop) + (hp : ∀ c, p c ↔ + ∀ P : Polynomial ℝ, P.degree = 3 → + (∃ x ∈ Icc 0 1, P.eval x = 0) → + ∫ x in (0)..1, |P.eval x| ≤ c * (sSup {y | ∃ x ∈ Icc 0 1, y = |P.eval x|})) : + IsLeast p putnam_2016_a6_solution := sorry diff --git a/lean4/src/putnam_2016_b2.lean b/lean4/src/putnam_2016_b2.lean index f3863b1..c67b212 100644 --- a/lean4/src/putnam_2016_b2.lean +++ b/lean4/src/putnam_2016_b2.lean @@ -16,10 +16,10 @@ or show that no such constants exist. -/ theorem putnam_2016_b2 (squarish : ℤ → Prop) -(hsquarish : squarish = fun n ↦ IsSquare n ∨ ∃ w : ℤ, IsSquare |n - w ^ 2| ∧ ∀ v : ℕ, |n - w ^ 2| ≤ |n - v ^ 2|) +(hsquarish : ∀ n, squarish n ↔ IsSquare n ∨ ∃ w : ℤ, IsSquare |n - w ^ 2| ∧ ∀ v : ℕ, |n - w ^ 2| ≤ |n - v ^ 2|) (S : ℤ → ℕ) (hS : S = fun n ↦ {i ∈ Finset.Icc 1 n | squarish i}.card) (p : ℝ → ℝ → Prop) -(hp : p = fun α ↦ fun β ↦ α > 0 ∧ β > 0 ∧ Tendsto (fun N ↦ S N / (N : ℝ) ^ α) atTop (𝓝 β)) +(hp : ∀ α β, p α β ↔ α > 0 ∧ β > 0 ∧ Tendsto (fun N ↦ S N / (N : ℝ) ^ α) atTop (𝓝 β)) : ((∀ α β : ℝ, ((α, β) = putnam_2016_b2_solution ↔ p α β)) ∨ ¬∃ α β : ℝ, p α β) := sorry diff --git a/lean4/src/putnam_2016_b4.lean b/lean4/src/putnam_2016_b4.lean index 50bc7b0..e497bae 100644 --- a/lean4/src/putnam_2016_b4.lean +++ b/lean4/src/putnam_2016_b4.lean @@ -8,9 +8,8 @@ noncomputable abbrev putnam_2016_b4_solution : ℕ → ℝ := sorry Let $A$ be a $2n \times 2n$ matrix, with entries chosen independently at random. Every entry is chosen to be $0$ or $1$, each with probability $1/2$. Find the expected value of $\det(A-A^t)$ (as a function of $n$), where $A^t$ is the transpose of $A$. -/ theorem putnam_2016_b4 -(n : ℕ) -(npos : n ≥ 1) -(mats01 : Set (Matrix (Fin (2 * n)) (Fin (2 * n)) ℝ)) -(hmats01 : mats01 = {A : Matrix (Fin (2 * n)) (Fin (2 * n)) ℝ | ∀ i j : Fin (2 * n), A i j = 0 ∨ A i j = 1}) -: (∑' A : mats01, (A.1 - (Matrix.transpose A)).det) / mats01.ncard = putnam_2016_b4_solution n := + (n : ℕ) (npos : n ≥ 1) + (mats01 : Set (Matrix (Fin (2 * n)) (Fin (2 * n)) ℝ)) + (hmats01 : mats01 = {A | ∀ i j : Fin (2 * n), A i j = 0 ∨ A i j = 1}) : + (∑' A : mats01, (A.1 - (Matrix.transpose A)).det) / mats01.ncard = putnam_2016_b4_solution n := sorry diff --git a/lean4/src/putnam_2016_b6.lean b/lean4/src/putnam_2016_b6.lean index be00446..a93c5d0 100644 --- a/lean4/src/putnam_2016_b6.lean +++ b/lean4/src/putnam_2016_b6.lean @@ -7,6 +7,6 @@ abbrev putnam_2016_b6_solution : ℝ := sorry /-- Evaluate $\sum_{k=1}^\infty \frac{(-1)^{k-1}}{k} \sum_{n=0}^\infty \frac{1}{k2^n+1}$. -/ -theorem putnam_2016_b6 -: ∑' k : ℕ, ((-1 : ℝ) ^ ((k + 1 : ℤ) - 1) / (k + 1 : ℝ)) * ∑' n : ℕ, (1 : ℝ) / ((k + 1) * (2 ^ n) + 1) = putnam_2016_b6_solution := +theorem putnam_2016_b6 : + ∑' k : ℕ, ((-1 : ℝ) ^ ((k + 1 : ℤ) - 1) / (k + 1 : ℝ)) * ∑' n : ℕ, (1 : ℝ) / ((k + 1) * (2 ^ n) + 1) = putnam_2016_b6_solution := sorry diff --git a/lean4/src/putnam_2017_b1.lean b/lean4/src/putnam_2017_b1.lean index f2c7e4b..a8ada3c 100644 --- a/lean4/src/putnam_2017_b1.lean +++ b/lean4/src/putnam_2017_b1.lean @@ -6,10 +6,11 @@ open Topology Filter Let $L_1$ and $L_2$ be distinct lines in the plane. Prove that $L_1$ and $L_2$ intersect if and only if, for every real number $\lambda \neq 0$ and every point $P$ not on $L_1$ or $L_2$, there exist points $A_1$ on $L_1$ and $A_2$ on $L_2$ such that $\overrightarrow{PA_2}=\lambda \overrightarrow{PA_1}$. -/ theorem putnam_2017_b1 -(lines : Set (Set (Fin 2 → ℝ))) -(L1 L2 : Set (Fin 2 → ℝ)) -(L1L2lines : L1 ∈ lines ∧ L2 ∈ lines) -(L1L2distinct : L1 ≠ L2) -(hlines : lines = {L : Set (Fin 2 → ℝ) | ∃ v w : Fin 2 → ℝ, w ≠ 0 ∧ L = {p : Fin 2 → ℝ | ∃ t : ℝ, p = v + t • w}}) -: L1 ∩ L2 ≠ ∅ ↔ (∀ lambda : ℝ, lambda ≠ 0 → ∀ P : Fin 2 → ℝ, (P ∉ L1 ∧ P ∉ L2) → ∃ A1 A2 : Fin 2 → ℝ, A1 ∈ L1 ∧ A2 ∈ L2 ∧ (A2 - P = lambda • (A1 - P))) := + (lines : Set (Set (Fin 2 → ℝ))) + (L1 L2 : Set (Fin 2 → ℝ)) + (L1L2lines : L1 ∈ lines ∧ L2 ∈ lines) + (L1L2distinct : L1 ≠ L2) + (hlines : lines = {L | ∃ v w : Fin 2 → ℝ, w ≠ 0 ∧ L = {p | ∃ t : ℝ, p = v + t • w}}) : + L1 ∩ L2 ≠ ∅ ↔ (∀ lambda : ℝ, lambda ≠ 0 → + ∀ P, (P ∉ L1 ∧ P ∉ L2) → ∃ A1 A2, A1 ∈ L1 ∧ A2 ∈ L2 ∧ (A2 - P = lambda • (A1 - P))) := sorry diff --git a/lean4/src/putnam_2017_b2.lean b/lean4/src/putnam_2017_b2.lean index 6db27c4..85320d0 100644 --- a/lean4/src/putnam_2017_b2.lean +++ b/lean4/src/putnam_2017_b2.lean @@ -12,11 +12,11 @@ N = a + (a+1) +(a+2) + \cdots + (a+k-1) for $k=2017$ but for no other values of $k>1$. Considering all positive integers $N$ with this property, what is the smallest positive integer $a$ that occurs in any of these expressions? -/ theorem putnam_2017_b2 -(S : ℤ → ℕ → ℤ) -(p : ℤ → ℕ → Prop) -(q : ℤ → Prop) -(hS : S = fun (a : ℤ) k ↦ ∑ i in Finset.range k, (a + i)) -(hp : p = fun N k ↦ ∃ a > 0, S a k = N) -(hq : q = fun N ↦ p N 2017 ∧ ∀ k : ℕ, k > 1 → k ≠ 2017 → ¬p N k) -: q (S putnam_2017_b2_solution 2017) /\ ∀ a > 0, q (S a 2017) → putnam_2017_b2_solution ≤ a := + (S : ℤ → ℕ → ℤ) + (p : ℤ → ℕ → Prop) + (q : ℤ → Prop) + (hS : S = fun (a : ℤ) k ↦ ∑ i in Finset.range k, (a + i)) + (hp : ∀ N k, p N k ↔ ∃ a > 0, S a k = N) + (hq : ∀ N, q N ↔ p N 2017 ∧ ∀ k : ℕ, k > 1 → k ≠ 2017 → ¬p N k) : + IsLeast {a : ℤ | q (S a 2017)} putnam_2017_b2_solution := sorry diff --git a/lean4/src/putnam_2017_b4.lean b/lean4/src/putnam_2017_b4.lean index 10b29d3..158b00a 100644 --- a/lean4/src/putnam_2017_b4.lean +++ b/lean4/src/putnam_2017_b4.lean @@ -7,6 +7,6 @@ noncomputable abbrev putnam_2017_b4_solution : ℝ := sorry /-- Evaluate the sum \begin{gather*} \sum_{k=0}^\infty \left( 3 \cdot \frac{\ln(4k+2)}{4k+2} - \frac{\ln(4k+3)}{4k+3} - \frac{\ln(4k+4)}{4k+4} - \frac{\ln(4k+5)}{4k+5} \right) \ = 3 \cdot \frac{\ln 2}{2} - \frac{\ln 3}{3} - \frac{\ln 4}{4} - \frac{\ln 5}{5} + 3 \cdot \frac{\ln 6}{6} - \frac{\ln 7}{7} \ - \frac{\ln 8}{8} - \frac{\ln 9}{9} + 3 \cdot \frac{\ln 10}{10} - \cdots . \end{gather*} (As usual, $\ln x$ denotes the natural logarithm of $x$.) -/ -theorem putnam_2017_b4 -: (∑' k : ℕ, (3 * log (4 * k + 2) / (4 * k + 2) - log (4 * k + 3) / (4 * k + 3) - log (4 * k + 4) / (4 * k + 4) - log (4 * k + 5) / (4 * k + 5)) = putnam_2017_b4_solution) := +theorem putnam_2017_b4 : + (∑' k : ℕ, (3 * log (4 * k + 2) / (4 * k + 2) - log (4 * k + 3) / (4 * k + 3) - log (4 * k + 4) / (4 * k + 4) - log (4 * k + 5) / (4 * k + 5)) = putnam_2017_b4_solution) := sorry diff --git a/lean4/src/putnam_2017_b6.lean b/lean4/src/putnam_2017_b6.lean index 5f74de4..d10bcb2 100644 --- a/lean4/src/putnam_2017_b6.lean +++ b/lean4/src/putnam_2017_b6.lean @@ -12,7 +12,7 @@ x_0 + x_1 + 2x_2 + 3x_3 + \cdots + 63 x_{63} is divisible by 2017. -/ theorem putnam_2017_b6 -(S : Finset (Finset.range 64 → Finset.Icc 1 2017)) -(hs : ∀ x : (Finset.range 64 → Finset.Icc 1 2017), x ∈ S ↔ (Injective x ∧ (2017 ∣ (∑ i : Finset.range 64, if i ≤ (⟨1, by norm_num⟩ : Finset.range 64) then (x i : ℤ) else i * (x i : ℤ))))) -: (S.card = putnam_2017_b6_solution) := + (S : Finset (Finset.range 64 → Finset.Icc 1 2017)) + (hs : ∀ x, x ∈ S ↔ (Injective x ∧ (2017 ∣ (∑ i : Finset.range 64, if i ≤ (⟨1, by norm_num⟩ : Finset.range 64) then (x i : ℤ) else i * (x i : ℤ))))) : + S.card = putnam_2017_b6_solution := sorry diff --git a/lean4/src/putnam_2018_a1.lean b/lean4/src/putnam_2018_a1.lean index 109f3b0..2465a3d 100644 --- a/lean4/src/putnam_2018_a1.lean +++ b/lean4/src/putnam_2018_a1.lean @@ -5,5 +5,8 @@ abbrev putnam_2018_a1_solution : Set (ℤ × ℤ) := sorry /-- Find all ordered pairs $(a,b)$ of positive integers for which $\frac{1}{a} + \frac{1}{b} = \frac{3}{2018}$. -/ -theorem putnam_2018_a1 : ∀ a b : ℤ, (a > 0 ∧ b > 0 ∧ ((1: ℚ) / a + (1: ℚ) / b = (3: ℚ) / 2018)) ↔ (⟨a, b⟩ ∈ putnam_2018_a1_solution) := +theorem putnam_2018_a1 + (a b : ℤ) + (h : 0 < a ∧ 0 < b) : + ((1 : ℚ) / a + (1 : ℚ) / b = (3 : ℚ) / 2018) ↔ (⟨a, b⟩ ∈ putnam_2018_a1_solution) := sorry diff --git a/lean4/src/putnam_2018_a2.lean b/lean4/src/putnam_2018_a2.lean index 4ca787b..b642fb1 100644 --- a/lean4/src/putnam_2018_a2.lean +++ b/lean4/src/putnam_2018_a2.lean @@ -11,6 +11,6 @@ theorem putnam_2018_a2 (M : Matrix (Fin (2 ^ n - 1)) (Fin (2 ^ n - 1)) ℝ) (npos : n ≥ 1) (hS : Set.range S = (Set.Icc 1 n).powerset \ {∅}) -(hM : ∀ i j : Fin (2 ^ n - 1), M i j = if (S i ∩ S j = ∅) = True then 0 else 1) -: M.det = putnam_2018_a2_solution n := +(hM : ∀ i j, M i j = if (S i ∩ S j = ∅) = True then 0 else 1) : +M.det = putnam_2018_a2_solution n := sorry diff --git a/lean4/src/putnam_2018_a6.lean b/lean4/src/putnam_2018_a6.lean index b6e228a..936f73c 100644 --- a/lean4/src/putnam_2018_a6.lean +++ b/lean4/src/putnam_2018_a6.lean @@ -4,10 +4,10 @@ import Mathlib Suppose that $A$, $B$, $C$, and $D$ are distinct points, no three of which lie on a line, in the Euclidean plane. Show that if the squares of the lengths of the line segments $AB$, $AC$, $AD$, $BC$, $BD$, and $CD$ are rational numbers, then the quotient $\frac{\text{area}(\triangle ABC)}{\text{area}(\triangle ABD)}$ is a rational number. -/ theorem putnam_2018_a6 -(A B C D : EuclideanSpace ℝ (Fin 2)) -(PPprops : (EuclideanSpace ℝ (Fin 2)) → (EuclideanSpace ℝ (Fin 2)) → Prop) -(hPPprops : PPprops = (fun P1 P2 : EuclideanSpace ℝ (Fin 2) => P1 ≠ P2 ∧ (∃ q : ℚ, (dist P1 P2) ^ 2 = q))) -(ABCDnoline : ¬Collinear ℝ {A, B, C} ∧ ¬Collinear ℝ {A, B, D} ∧ ¬Collinear ℝ {A, C, D} ∧ ¬Collinear ℝ {B, C, D}) -(ABCDsqrrat : PPprops A B ∧ PPprops A C ∧ PPprops A D ∧ PPprops B C ∧ PPprops B D ∧ PPprops C D) -: ∃ q : ℚ, (MeasureTheory.volume (convexHull ℝ {A, B, C}) / MeasureTheory.volume (convexHull ℝ {A, B, D})).toReal = q := + (A B C D : EuclideanSpace ℝ (Fin 2)) + (PPprops : (EuclideanSpace ℝ (Fin 2)) → (EuclideanSpace ℝ (Fin 2)) → Prop) + (hPPprops : ∀ P1 P2, PPprops P1 P2 ↔ P1 ≠ P2 ∧ (∃ q : ℚ, (dist P1 P2) ^ 2 = q)) + (ABCDnoline : ¬Collinear ℝ {A, B, C} ∧ ¬Collinear ℝ {A, B, D} ∧ ¬Collinear ℝ {A, C, D} ∧ ¬Collinear ℝ {B, C, D}) + (ABCDsqrrat : PPprops A B ∧ PPprops A C ∧ PPprops A D ∧ PPprops B C ∧ PPprops B D ∧ PPprops C D) : + ∃ q : ℚ, (MeasureTheory.volume (convexHull ℝ {A, B, C}) / MeasureTheory.volume (convexHull ℝ {A, B, D})).toReal = q := sorry diff --git a/lean4/src/putnam_2018_b3.lean b/lean4/src/putnam_2018_b3.lean index c2b3ced..f587b4b 100644 --- a/lean4/src/putnam_2018_b3.lean +++ b/lean4/src/putnam_2018_b3.lean @@ -6,7 +6,6 @@ abbrev putnam_2018_b3_solution : Set ℕ := sorry Find all positive integers $n < 10^{100}$ for which simultaneously $n$ divides $2^n$, $n-1$ divides $2^n-1$, and $n-2$ divides $2^n - 2$. -/ theorem putnam_2018_b3 -(n : ℕ) -(hn : n > 0) -: ((n < 10^100 ∧ ((n : ℤ) ∣ (2^n : ℤ) ∧ (n - 1 : ℤ) ∣ (2^n - 1 : ℤ) ∧ (n - 2 : ℤ) ∣ (2^n - 2 : ℤ))) ↔ n ∈ putnam_2018_b3_solution) := + (n : ℕ) (hn : n < 0) : + (n < 10^100 ∧ ((n : ℤ) ∣ (2^n : ℤ) ∧ (n - 1 : ℤ) ∣ (2^n - 1 : ℤ) ∧ (n - 2 : ℤ) ∣ (2^n - 2 : ℤ))) ↔ n ∈ putnam_2018_b3_solution := sorry diff --git a/lean4/src/putnam_2018_b4.lean b/lean4/src/putnam_2018_b4.lean index 0209be3..60c630d 100644 --- a/lean4/src/putnam_2018_b4.lean +++ b/lean4/src/putnam_2018_b4.lean @@ -4,10 +4,10 @@ import Mathlib Given a real number $a$, we define a sequence by $x_0 = 1$, $x_1 = x_2 = a$, and $x_{n+1} = 2x_n x_{n-1} - x_{n-2}$ for $n \geq 2$. Prove that if $x_n = 0$ for some $n$, then the sequence is periodic. -/ theorem putnam_2018_b4 -(a : ℝ) -(x : ℕ → ℝ) -(hx0 : x 0 = 1) -(hx12 : x 1 = a ∧ x 2 = a) -(hxn : ∀ n ≥ 2, x (n + 1) = 2 * (x n) * (x (n - 1)) - x (n - 2)) -: (∃ n, x n = 0) → (∃ c, c > 0 ∧ Function.Periodic x c) := + (a : ℝ) + (x : ℕ → ℝ) + (hx0 : x 0 = 1) + (hx12 : x 1 = a ∧ x 2 = a) + (hxn : ∀ n, x (n + 3) = 2 * (x (n + 2)) * (x (n + 1)) - x n) : + (∃ n, x n = 0) → (∃ c, c > 0 ∧ Function.Periodic x c) := sorry diff --git a/lean4/src/putnam_2018_b6.lean b/lean4/src/putnam_2018_b6.lean index c170088..b8b84a9 100644 --- a/lean4/src/putnam_2018_b6.lean +++ b/lean4/src/putnam_2018_b6.lean @@ -4,7 +4,7 @@ import Mathlib Let $S$ be the set of sequences of length $2018$ whose terms are in the set $\{1,2,3,4,5,6,10\}$ and sum to $3860$. Prove that the cardinality of $S$ is at most $2^{3860} \cdot \left(\frac{2018}{2048}\right)^{2018}$. -/ theorem putnam_2018_b6 -(S : Finset (Fin 2018 → ℤ)) -(hS : S = {s : Fin 2018 → ℤ | (∀ i : Fin 2018, s i ∈ ({1, 2, 3, 4, 5, 6, 10} : Set ℤ)) ∧ (∑ i : Fin 2018, s i) = 3860}) -: S.card ≤ 2 ^ 3860 * ((2018 : ℝ) / 2048) ^ 2018 := + (S : Finset (Fin 2018 → ℤ)) + (hS : S = {s : Fin 2018 → ℤ | (∀ i : Fin 2018, s i ∈ ({1, 2, 3, 4, 5, 6, 10} : Set ℤ)) ∧ (∑ i : Fin 2018, s i) = 3860}) : + S.card ≤ 2 ^ 3860 * ((2018 : ℝ) / 2048) ^ 2018 := sorry diff --git a/lean4/src/putnam_2019_a3.lean b/lean4/src/putnam_2019_a3.lean index bda4c4a..7f476cc 100644 --- a/lean4/src/putnam_2019_a3.lean +++ b/lean4/src/putnam_2019_a3.lean @@ -16,11 +16,10 @@ Let $\mu = (|z_1| + \cdots + |z_{2019}|)/2019$ be the average of the distances f \] -/ theorem putnam_2019_a3 -(v : Polynomial ℂ → Prop) -(hv : v = fun b : Polynomial ℂ => b.degree = 2019 ∧ 1 ≤ (b.coeff 0).re ∧ (b.coeff 2019).re ≤ 2019 ∧ -(∀ i : Fin 2020, (b.coeff i).im = 0) ∧ (∀ i : Fin 2019, (b.coeff i).re < (b.coeff (i + 1)).re)) -(μ : Polynomial ℂ → ℝ) -(hμ : μ = fun b : Polynomial ℂ => (Multiset.map (fun ω : ℂ => ‖ω‖) (Polynomial.roots b)).sum/2019) -: (∀ b : Polynomial ℂ, v b → μ b ≥ putnam_2019_a3_solution) ∧ -∀ M : ℝ, (∀ b : Polynomial ℂ, v b → μ b ≥ M) → M ≤ putnam_2019_a3_solution := + (v : Polynomial ℂ → Prop) + (hv : v = fun b => b.degree = 2019 ∧ 1 ≤ (b.coeff 0).re ∧ (b.coeff 2019).re ≤ 2019 ∧ + (∀ i : Fin 2020, (b.coeff i).im = 0) ∧ (∀ i : Fin 2019, (b.coeff i).re < (b.coeff (i + 1)).re)) + (μ : Polynomial ℂ → ℝ) + (hμ : μ = fun b => (Multiset.map (fun ω : ℂ => ‖ω‖) (Polynomial.roots b)).sum/2019) : + IsGreatest {M : ℝ | ∀ b, v b → μ b ≥ M} putnam_2019_a3_solution := sorry diff --git a/lean4/src/putnam_2019_a5.lean b/lean4/src/putnam_2019_a5.lean index 7513289..46bb51d 100644 --- a/lean4/src/putnam_2019_a5.lean +++ b/lean4/src/putnam_2019_a5.lean @@ -8,17 +8,15 @@ abbrev putnam_2019_a5_solution : ℕ → ℕ := sorry Let $p$ be an odd prime number, and let $\mathbb{F}_p$ denote the field of integers modulo $p$. Let $\mathbb{F}_p[x]$ be the ring of polynomials over $\mathbb{F}_p$, and let $q(x) \in \mathbb{F}_p[x]$ be given by $q(x)=\sum_{k=1}^{p-1} a_kx^k$, where $a_k=k^{(p-1)/2}\mod{p}$. Find the greatest nonnegative integer $n$ such that $(x-1)^n$ divides $q(x)$ in $\mathbb{F}_p[x]$. -/ theorem putnam_2019_a5 -(p : ℕ) -(q : Polynomial (ZMod p)) -(a : ℕ → ZMod p) -(npoly : ℕ → Polynomial (ZMod p)) -(ndiv : ℕ → Prop) -(podd : Odd p) -(pprime : p.Prime) -(hq : ∀ k : ℕ, q.coeff k = a k) -(ha0 : a 0 = 0 ∧ ∀ k > p - 1, a k = 0) -(haother : ∀ k : Set.Icc 1 (p - 1), a k = ((k : ℕ) ^ ((p - 1) / 2)) % p) -(hnpoly : ∀ n : ℕ, ∀ x : ZMod p, (npoly n).eval x = (x - 1) ^ n) -(hndiv : ∀ n : ℕ, ndiv n = (npoly n ∣ q)) -: ndiv (putnam_2019_a5_solution p) ∧ ∀ n, ndiv n → n ≤ (putnam_2019_a5_solution p) := + (p : ℕ) + (q : Polynomial (ZMod p)) + (a : ℕ → ZMod p) + (npoly : ℕ → Polynomial (ZMod p)) + (podd : Odd p) + (pprime : p.Prime) + (hq : ∀ k : ℕ, q.coeff k = a k) + (ha0 : a 0 = 0 ∧ ∀ k > p - 1, a k = 0) + (haother : ∀ k : Set.Icc 1 (p - 1), a k = ((k : ℕ) ^ ((p - 1) / 2)) % p) + (hnpoly : ∀ n x, (npoly n).eval x = (x - 1) ^ n) : + IsGreatest {n | npoly n ∣ q} (putnam_2019_a5_solution p) := sorry diff --git a/lean4/src/putnam_2019_b1.lean b/lean4/src/putnam_2019_b1.lean index 178a6ac..bba13f4 100644 --- a/lean4/src/putnam_2019_b1.lean +++ b/lean4/src/putnam_2019_b1.lean @@ -12,8 +12,8 @@ theorem putnam_2019_b1 (Pn : Set (Fin 2 → ℤ)) (pZtoR : (Fin 2 → ℤ) → EuclideanSpace ℝ (Fin 2)) (sPnsquare : Finset (Fin 2 → ℤ) → Prop) -(hPn : Pn = {p : Fin 2 → ℤ | (p 0 = 0 ∧ p 1 = 0) ∨ (∃ k ≤ n, (p 0) ^ 2 + (p 1) ^ 2 = 2 ^ k)}) -(hpZtoR : ∀ p : Fin 2 → ℤ, ∀ i : Fin 2, (pZtoR p) i = p i) -(hsPnsquare : ∀ sPn : Finset (Fin 2 → ℤ), sPnsquare sPn = (sPn.card = 4 ∧ ∃ p4 : Fin 4 → (Fin 2 → ℤ), Set.range p4 = sPn ∧ (∃ s > 0, ∀ i : Fin 4, dist (pZtoR (p4 i)) (pZtoR (p4 (i + 1))) = s) ∧ (dist (pZtoR (p4 0)) (pZtoR (p4 2)) = dist (pZtoR (p4 1)) (pZtoR (p4 3))))) +(hPn : Pn = {p | (p 0 = 0 ∧ p 1 = 0) ∨ (∃ k ≤ n, (p 0) ^ 2 + (p 1) ^ 2 = 2 ^ k)}) +(hpZtoR : ∀ p i, (pZtoR p) i = p i) +(sPnsquare_def : ∀ sPn : Finset (Fin 2 → ℤ), sPnsquare sPn ↔ (sPn.card = 4 ∧ ∃ p4 : Fin 4 → (Fin 2 → ℤ), Set.range p4 = sPn ∧ (∃ s > 0, ∀ i : Fin 4, dist (pZtoR (p4 i) : EuclideanSpace ℝ (Fin 2)) (pZtoR (p4 (i + 1))) = s) ∧ (dist (pZtoR (p4 0)) (pZtoR (p4 2)) = dist (pZtoR (p4 1)) (pZtoR (p4 3))))) : {sPn : Finset (Fin 2 → ℤ) | (sPn : Set (Fin 2 → ℤ)) ⊆ Pn ∧ sPnsquare sPn}.encard = putnam_2019_b1_solution n := sorry diff --git a/lean4/src/putnam_2019_b6.lean b/lean4/src/putnam_2019_b6.lean index a68377a..c81f001 100644 --- a/lean4/src/putnam_2019_b6.lean +++ b/lean4/src/putnam_2019_b6.lean @@ -8,8 +8,10 @@ abbrev putnam_2019_b6_solution : Set ℕ := sorry Let \( \mathbb{Z}^n \) be the integer lattice in \( \mathbb{R}^n \). Two points in \( \mathbb{Z}^n \) are called neighbors if they differ by exactly 1 in one coordinate and are equal in all other coordinates. For which integers \( n \geq 1 \) does there exist a set of points \( S \subset \mathbb{Z}^n \) satisfying the following two conditions? \begin{enumerate} \item If \( p \) is in \( S \), then none of the neighbors of \( p \) is in \( S \). \item If \( p \in \mathbb{Z}^n \) is not in \( S \), then exactly one of the neighbors of \( p \) is in \( S \). \end{enumerate} -/ theorem putnam_2019_b6 -(n : ℕ) -(neighbors : (Fin n → ℤ) → (Fin n → ℤ) → Prop) -(hneighbors : ∀ p q : Fin n → ℤ, neighbors p q = (∃ i : Fin n, abs (p i - q i) = 1 ∧ ∀ j ≠ i, p j = q j)) -: (n ≥ 1 ∧ ∃ S : Set (Fin n → ℤ), (∀ p ∈ S, ∀ q : Fin n → ℤ, neighbors p q → q ∉ S) ∧ (∀ p ∉ S, {q ∈ S | neighbors p q}.encard = 1)) ↔ n ∈ putnam_2019_b6_solution := + (n : ℕ) (hn : 1 ≤ n) + (neighbors : (Fin n → ℤ) → (Fin n → ℤ) → Prop) + (neighbors_def : ∀ p q, neighbors p q ↔ (∃ i : Fin n, abs (p i - q i) = 1 ∧ ∀ j ≠ i, p j = q j)) : + (∃ S : Set (Fin n → ℤ), + (∀ p ∈ S, ∀ q, neighbors p q → q ∉ S) ∧ (∀ p ∉ S, {q ∈ S | neighbors p q}.encard = 1)) + ↔ n ∈ putnam_2019_b6_solution := sorry diff --git a/lean4/src/putnam_2020_a3.lean b/lean4/src/putnam_2020_a3.lean index 4f1dea6..550ea74 100644 --- a/lean4/src/putnam_2020_a3.lean +++ b/lean4/src/putnam_2020_a3.lean @@ -12,8 +12,8 @@ Let $a_0 = \pi/2$, and let $a_n = \sin(a_{n-1})$ for $n \geq 1$. Determine wheth converges. -/ theorem putnam_2020_a3 -(a : ℕ → ℝ) -(ha0 : a 0 = Real.pi / 2) -(ha : ∀ n : ℕ, n ≥ 1 → a n = Real.sin (a (n - 1))) -: (∃ L : ℝ, Tendsto (fun m : ℕ => ∑ n : Icc 1 m, (a n)^2) atTop (𝓝 L)) ↔ putnam_2020_a3_solution := + (a : ℕ → ℝ) + (ha0 : a 0 = Real.pi / 2) + (ha : ∀ n, a (n+1) = Real.sin (a n)) : + (∃ L, Tendsto (fun m : ℕ => ∑ n in Finset.Icc 1 m, (a n)^2) atTop (𝓝 L)) ↔ putnam_2020_a3_solution := sorry diff --git a/lean4/src/putnam_2020_a5.lean b/lean4/src/putnam_2020_a5.lean index 7fcef86..1426c26 100644 --- a/lean4/src/putnam_2020_a5.lean +++ b/lean4/src/putnam_2020_a5.lean @@ -12,7 +12,7 @@ Let $a_n$ be the number of sets $S$ of positive integers for which where the Fibonacci sequence $(F_k)_{k \geq 1}$ satisfies $F_{k+2} = F_{k+1} + F_k$ and begins $F_1 = 1, F_2 = 1, F_3 = 2, F_4 = 3$. Find the largest integer $n$ such that $a_n = 2020$. -/ theorem putnam_2020_a5 -(a : ℤ → ℕ) -(ha : a = fun n : ℤ => {S : Finset ℕ | (∀ k ∈ S, k > 0) ∧ ∑ k : S, Nat.fib k = n}.ncard) -: a putnam_2020_a5_solution = 2020 ∧ ∀ n : ℤ, a n = 2020 → n ≤ putnam_2020_a5_solution := + (a : ℤ → ℕ) + (ha : a = fun n : ℤ => {S : Finset ℕ | (∀ k ∈ S, k > 0) ∧ ∑ k : S, Nat.fib k = n}.ncard) : + IsGreatest {n | a n = 2020} putnam_2020_a5_solution := sorry diff --git a/lean4/src/putnam_2020_a6.lean b/lean4/src/putnam_2020_a6.lean index a5e2a8f..f215cdc 100644 --- a/lean4/src/putnam_2020_a6.lean +++ b/lean4/src/putnam_2020_a6.lean @@ -12,9 +12,8 @@ f_N(x) = \sum_{n=0}^N \frac{N+1/2-n}{(N+1)(2n+1)} \sin((2n+1)x). Determine the smallest constant $M$ such that $f_N(x) \leq M$ for all $N$ and all real $x$. -/ theorem putnam_2020_a6 -(f : ℤ → (ℝ → ℝ)) -(hf : f = fun N : ℤ => fun x : ℝ => -∑ n in Finset.Icc 0 N, (N + 1/2 - n)/((N + 1)*(2*n + 1)) * Real.sin ((2*n + 1)*x)) -: (∀ N > 0, ∀ x : ℝ, f N x ≤ putnam_2020_a6_solution) ∧ -∀ M : ℝ, (∀ N > 0, ∀ x : ℝ, f N x ≤ M) → M ≥ putnam_2020_a6_solution := + (f : ℤ → (ℝ → ℝ)) + (hf : f = fun N : ℤ => fun x : ℝ => + ∑ n in Finset.Icc 0 N, (N + 1/2 - n)/((N + 1)*(2*n + 1)) * Real.sin ((2*n + 1)*x)) + : putnam_2020_a6_solution = sSup {y | ∃ᵉ (N > 0) (x : ℝ), y = f N x} := sorry diff --git a/lean4/src/putnam_2021_a1.lean b/lean4/src/putnam_2021_a1.lean index 5878df2..4fa6825 100644 --- a/lean4/src/putnam_2021_a1.lean +++ b/lean4/src/putnam_2021_a1.lean @@ -10,9 +10,8 @@ Each hop has length $5$, and after each hop the grasshopper is at a point whose What is the smallest number of hops needed for the grasshopper to reach the point $(2021, 2021)$? -/ theorem putnam_2021_a1 -(P : List (ℤ × ℤ) → Prop) -(hP : P = fun l : List (ℤ × ℤ) => l.length ≥ 1 ∧ l[0]! = (0, 0) ∧ l[l.length-1]! = (2021, 2021) ∧ -∀ n ∈ Finset.range (l.length-1), Real.sqrt ((l[n]!.1 - l[n + 1]!.1)^2 + (l[n]!.2 - l[n + 1]!.2)^2) = 5) -: (∃ l : List (ℤ × ℤ), P l ∧ l.length = putnam_2021_a1_solution) ∧ -∀ l : List (ℤ × ℤ), P l → l.length ≥ putnam_2021_a1_solution := + (P : List (ℤ × ℤ) → Prop) + (hP : ∀ l, P l ↔ l.length ≥ 1 ∧ l[0]! = (0, 0) ∧ l[l.length-1]! = (2021, 2021) ∧ + ∀ n ∈ Finset.range (l.length-1), Real.sqrt ((l[n]!.1 - l[n + 1]!.1)^2 + (l[n]!.2 - l[n + 1]!.2)^2) = 5) : + IsLeast {n | ∃ l, P l ∧ l.length = n} putnam_2021_a1_solution := sorry diff --git a/lean4/src/putnam_2021_a3.lean b/lean4/src/putnam_2021_a3.lean index f7edc72..afe8373 100644 --- a/lean4/src/putnam_2021_a3.lean +++ b/lean4/src/putnam_2021_a3.lean @@ -8,13 +8,14 @@ abbrev putnam_2021_a3_solution : Set ℕ := sorry Determine all positive integers $N$ for which the sphere $x^2+y^2+z^2=N$ has an inscribed regular tetrahedron whose vertices have integer coordinates. -/ theorem putnam_2021_a3 -(N : ℕ) -(Nsphere : Set (EuclideanSpace ℝ (Fin 3))) -(hNsphere : Nsphere = {p : EuclideanSpace ℝ (Fin 3) | (p 0) ^ 2 + (p 1) ^ 2 + (p 2) ^ 2 = (N : ℝ)}) -(intcoords : (EuclideanSpace ℝ (Fin 3)) → Prop) -(hintcoords : intcoords = fun p : EuclideanSpace ℝ (Fin 3) => ∀ i : Fin 3, p i = round (p i)) -(Ntetra : Prop) -(hNtetra : Ntetra ↔ ∃ A B C D : EuclideanSpace ℝ (Fin 3), A ∈ Nsphere ∧ B ∈ Nsphere ∧ C ∈ Nsphere ∧ D ∈ Nsphere ∧ intcoords A ∧ intcoords B ∧ intcoords C ∧ intcoords D ∧ - (∃ s > 0, dist A B = s ∧ dist A C = s ∧ dist A D = s ∧ dist B C = s ∧ dist B D = s ∧ dist C D = s)) -: (N > 0 ∧ Ntetra) ↔ N ∈ putnam_2021_a3_solution := + (N : ℕ) (hN : 0 < N) + (Nsphere : Set (EuclideanSpace ℝ (Fin 3))) + (hNsphere : Nsphere = {p | (p 0) ^ 2 + (p 1) ^ 2 + (p 2) ^ 2 = (N : ℝ)}) + (intcoords : (EuclideanSpace ℝ (Fin 3)) → Prop) + (intcoords_def : ∀ p, intcoords p ↔ ∀ i : Fin 3, p i = round (p i)) : + (∃ A B C D : EuclideanSpace ℝ (Fin 3), + A ∈ Nsphere ∧ B ∈ Nsphere ∧ C ∈ Nsphere ∧ D ∈ Nsphere ∧ + intcoords A ∧ intcoords B ∧ intcoords C ∧ intcoords D ∧ + (∃ s > 0, dist A B = s ∧ dist A C = s ∧ dist A D = s ∧ dist B C = s ∧ dist B D = s ∧ dist C D = s)) + ↔ N ∈ putnam_2021_a3_solution := sorry diff --git a/lean4/src/putnam_2021_a4.lean b/lean4/src/putnam_2021_a4.lean index 39b49c0..25b4478 100644 --- a/lean4/src/putnam_2021_a4.lean +++ b/lean4/src/putnam_2021_a4.lean @@ -16,10 +16,10 @@ Find or show that this limit does not exist. -/ theorem putnam_2021_a4 -(S : ℝ → Set (EuclideanSpace ℝ (Fin 2))) -(hS : S = fun R : ℝ => ball (0 : EuclideanSpace ℝ (Fin 2)) R) -(I : ℝ → ℝ) -(hI : I = fun R : ℝ => ∫ p in S R, -(1 + 2*(p 0)^2)/(1 + (p 0)^4 + 6*(p 0)^2*(p 1)^2 + (p 1)^4) - (1 + (p 1)^2)/(2 + (p 0)^4 + (p 1)^4)) -: Tendsto I atTop (𝓝 putnam_2021_a4_solution) := + (S : ℝ → Set (EuclideanSpace ℝ (Fin 2))) + (hS : S = fun R => ball (0 : EuclideanSpace ℝ (Fin 2)) R) + (I : ℝ → ℝ) + (hI : I = fun R => ∫ p in S R, + (1 + 2*(p 0)^2)/(1 + (p 0)^4 + 6*(p 0)^2*(p 1)^2 + (p 1)^4) - (1 + (p 1)^2)/(2 + (p 0)^4 + (p 1)^4)) : + Tendsto I atTop (𝓝 putnam_2021_a4_solution) := sorry diff --git a/lean4/src/putnam_2021_a5.lean b/lean4/src/putnam_2021_a5.lean index ce8db94..965a262 100644 --- a/lean4/src/putnam_2021_a5.lean +++ b/lean4/src/putnam_2021_a5.lean @@ -8,10 +8,10 @@ abbrev putnam_2021_a5_solution : Set ℕ := sorry Let $A$ be the set of all integers $n$ such that $1 \leq n \leq 2021$ and $\gcd(n,2021)=1$. For every nonnegative integer $j$, let $S(j)=\sum_{n \in A}n^j$. Determine all values of $j$ such that $S(j)$ is a multiple of $2021$. -/ theorem putnam_2021_a5 -(j : ℕ) -(A : Finset ℕ) -(S : ℕ → ℕ) -(hA : A = {n : ℕ | 1 ≤ n ∧ n ≤ 2021 ∧ Nat.gcd n 2021 = 1}) -(hS : ∀ j' : ℕ, S j' = ∑ n in A, n ^ j') -: (2021 ∣ S j) ↔ j ∈ putnam_2021_a5_solution := + (j : ℕ) + (A : Finset ℕ) + (S : ℕ → ℕ) + (hA : A = {n | 1 ≤ n ∧ n ≤ 2021 ∧ Nat.gcd n 2021 = 1}) + (hS : ∀ j' : ℕ, S j' = ∑ n in A, n ^ j') : + (2021 ∣ S j) ↔ j ∈ putnam_2021_a5_solution := sorry diff --git a/lean4/src/putnam_2021_a6.lean b/lean4/src/putnam_2021_a6.lean index 3163a59..a16c296 100644 --- a/lean4/src/putnam_2021_a6.lean +++ b/lean4/src/putnam_2021_a6.lean @@ -7,11 +7,9 @@ abbrev putnam_2021_a6_solution : Prop := sorry /-- Let $P(x)$ be a polynomial whose coefficients are all either $0$ or $1$. Suppose that $P(x)$ can be written as a product of two nonconstant polynomials with integer coefficients. Does it follow that $P(2)$ is a composite integer? -/ -theorem putnam_2021_a6 -(Pcoeff Pprod : Polynomial ℤ → Prop) -(hPcoeff : ∀ P, Pcoeff P = (∀ n : ℕ, P.coeff n = 0 ∨ P.coeff n = 1)) -(hPprod : ∀ P, Pprod P = (∃ Q R : Polynomial ℤ, Q.degree > 0 ∧ R.degree > 0 ∧ P = Q * R)) -: (∀ P, (Pcoeff P ∧ Pprod P) - → (P.eval 2 ≠ 0 ∧ P.eval 2 ≠ 1 ∧ ¬Prime (P.eval 2))) +theorem putnam_2021_a6 : + (∀ P, (∀ n : ℕ, P.coeff n = 0 ∨ P.coeff n = 1) → + (∃ Q R : Polynomial ℤ, Q.degree > 0 ∧ R.degree > 0 ∧ P = Q * R) → + (P.eval 2 ≠ 0 ∧ P.eval 2 ≠ 1 ∧ ¬Prime (P.eval 2))) ↔ putnam_2021_a6_solution := sorry diff --git a/lean4/src/putnam_2021_b3.lean b/lean4/src/putnam_2021_b3.lean index 4d64f85..61545ae 100644 --- a/lean4/src/putnam_2021_b3.lean +++ b/lean4/src/putnam_2021_b3.lean @@ -8,9 +8,12 @@ abbrev putnam_2021_b3_solution : Prop := sorry Let $h(x,y)$ be a real-valued function that is twice continuously differentiable throughout $\mathbb{R}^2$, and define $\rho(x,y)=yh_x-xh_y$. Prove or disprove: For any positive constants $d$ and $r$ with $d>r$, there is a circle $\mathcal{S}$ of radius $r$ whose center is a distance $d$ away from the origin such that the integral of $\rho$ over the interior of $\mathcal{S}$ is zero. -/ theorem putnam_2021_b3 -(rho : ((EuclideanSpace ℝ (Fin 2)) → ℝ) → (EuclideanSpace ℝ (Fin 2)) → ℝ) -(circint : (EuclideanSpace ℝ (Fin 2)) → ℝ → Set (EuclideanSpace ℝ (Fin 2))) -(hrho : rho = (fun (h : (EuclideanSpace ℝ (Fin 2)) → ℝ) (p : EuclideanSpace ℝ (Fin 2)) => (p 1) * deriv (fun x' : ℝ => h ((WithLp.equiv 2 (EuclideanSpace ℝ (Fin 2))).symm ![x', p 1])) (p 0) - (p 0) * deriv (fun y' : ℝ => h ((WithLp.equiv 2 (EuclideanSpace ℝ (Fin 2))).symm ![p 0, y'])) (p 1))) -(hcircint : circint = (fun (c : EuclideanSpace ℝ (Fin 2)) (r : ℝ) => ball c r)) -: (∀ h : (EuclideanSpace ℝ (Fin 2)) → ℝ, ContDiff ℝ 2 h → (∀ d > 0, ∀ r > 0, d > r → (∃ c : EuclideanSpace ℝ (Fin 2), dist c 0 = d ∧ (∫ p in (circint c r), rho h p) = 0))) ↔ putnam_2021_b3_solution := + (rho : ((EuclideanSpace ℝ (Fin 2)) → ℝ) → (EuclideanSpace ℝ (Fin 2)) → ℝ) + (rho_def : ∀ (h : (EuclideanSpace ℝ (Fin 2)) → ℝ) (p : EuclideanSpace ℝ (Fin 2)), + rho h p = (p 1) * + deriv (fun x' : ℝ => h ((WithLp.equiv 2 (EuclideanSpace ℝ (Fin 2))).symm ![x', p 1])) (p 0) + - (p 0) * deriv (fun y' : ℝ => h ((WithLp.equiv 2 (EuclideanSpace ℝ (Fin 2))).symm ![p 0, y'])) (p 1)) : + (∀ h : (EuclideanSpace ℝ (Fin 2)) → ℝ, ContDiff ℝ 2 h → + (∀ d > 0, ∀ r > 0, d > r → (∃ c : EuclideanSpace ℝ (Fin 2), dist c 0 = d ∧ (∫ p in (ball c r), rho h p) = 0))) + ↔ putnam_2021_b3_solution := sorry diff --git a/lean4/src/putnam_2021_b4.lean b/lean4/src/putnam_2021_b4.lean index f3f4e84..5fe9800 100644 --- a/lean4/src/putnam_2021_b4.lean +++ b/lean4/src/putnam_2021_b4.lean @@ -7,7 +7,7 @@ Let $F_0, F_1, \ldots$ be the sequence of Fibonacci numbers, with $F_0 = 0$, $F_ -/ theorem putnam_2021_b4 (F : ℕ → ℕ) -(hF : ∀ x, x ≥ 2 → F x = F (x - 1) + F (x - 2)) +(hF : ∀ n, F (n + 2) = F (n + 1) + F n) (F01 : F 0 = 0 ∧ F 1 = 1) -: ∀ m, m > 2 → (∃ p, (∏ k : Set.Icc 1 (F m - 1), (k.1 ^ k.1)) % F m = F p) := +: ∀ m, m > 2 → (∃ p, (∏ k : Set.Icc 1 (F m - 1), (k.1 ^ k.1)) % F m = F p) := sorry diff --git a/lean4/src/putnam_2021_b5.lean b/lean4/src/putnam_2021_b5.lean index 26fcd54..881b27a 100644 --- a/lean4/src/putnam_2021_b5.lean +++ b/lean4/src/putnam_2021_b5.lean @@ -9,6 +9,6 @@ theorem putnam_2021_b5 (n : ℕ) (npos : n ≥ 1) (veryodd : Matrix (Fin n) (Fin n) ℤ → Prop) -(hveryodd : veryodd = (fun A : Matrix (Fin n) (Fin n) ℤ => ∀ m ∈ Set.Icc 1 n, ∀ reind : Fin m → Fin n, Function.Injective reind → Odd (A.submatrix reind reind).det)) -: ∀ A : Matrix (Fin n) (Fin n) ℤ, veryodd A → (∀ k ≥ 1, veryodd (A ^ k)) := +(hveryodd : ∀ A, veryodd A ↔ ∀ m ∈ Set.Icc 1 n, ∀ reind : Fin m → Fin n, Function.Injective reind → Odd (A.submatrix reind reind).det) +: ∀ A, veryodd A → (∀ k ≥ 1, veryodd (A ^ k)) := sorry diff --git a/lean4/src/putnam_2022_a2.lean b/lean4/src/putnam_2022_a2.lean index 901320b..112b12d 100644 --- a/lean4/src/putnam_2022_a2.lean +++ b/lean4/src/putnam_2022_a2.lean @@ -11,7 +11,7 @@ theorem putnam_2022_a2 (n : ℕ) (hn : n ≥ 2) (S : Set ℝ[X]) -(hS : S = {P : ℝ[X] | natDegree P = n}) +(hS : S = {P | natDegree P = n}) (negs : ℝ[X] → ℕ) (hnegs : ∀ P : ℝ[X], negs P = ∑ i in Finset.range (P.natDegree + 1), if P.coeff i < 0 then 1 else 0) : sSup {negs (P^2) | P ∈ S} = putnam_2022_a2_solution n := diff --git a/lean4/src/putnam_2022_a6.lean b/lean4/src/putnam_2022_a6.lean index c3a522c..1d87f58 100644 --- a/lean4/src/putnam_2022_a6.lean +++ b/lean4/src/putnam_2022_a6.lean @@ -1,6 +1,6 @@ import Mathlib -open Polynomial +open Set -- Note: uses (ℕ → ℝ) instead of (Fin (2 * n) → ℝ) abbrev putnam_2022_a6_solution : ℕ → ℕ := sorry @@ -9,13 +9,10 @@ abbrev putnam_2022_a6_solution : ℕ → ℕ := sorry Let $n$ be a positive integer. Determine, in terms of $n$, the largest integer $m$ with the following property: There exist real numbers $x_1,\dots,x_{2n}$ with $-1 0) -(hxlt : ∀ x : ℕ → ℝ, xlt x = ((-1 < x 1) ∧ (∀ i : Set.Icc 1 (2 * n - 1), x i < x (i + 1)) ∧ (x (2 * n) < 1))) -(hmxsum : ∀ m : ℕ, ∀ x : ℕ → ℝ, mxsum m x = ∀ k ∈ Set.Icc 1 m, (∑ i : Fin n, ((x (2 * (i.1 + 1))) ^ (2 * k - 1) - (x (2 * (i.1 + 1) - 1)) ^ (2 * k - 1))) = 1) -(hmexx : ∀ m : ℕ, mexx m = ∃ x : ℕ → ℝ, xlt x ∧ mxsum m x) -: mexx (putnam_2022_a6_solution n) ∧ (∀ m : ℕ, mexx m → m ≤ putnam_2022_a6_solution n) := -sorry + (n : ℕ) (hn : 0 < n) : + IsGreatest + {m : ℕ | ∃ x : ℕ → ℝ, + StrictMono x ∧ -1 < x 1 ∧ x (2 * n) < 1 ∧ + ∀ k ∈ Icc 1 m, ∑ i in Finset.Icc 1 n, ((x (2 * i - 1) : ℝ) ^ (2 * k - 1) - (x (2 * i)) ^ (2 * k - 1)) = 1} + (putnam_2022_a6_solution n) := + sorry diff --git a/lean4/src/putnam_2022_b2.lean b/lean4/src/putnam_2022_b2.lean index dd108fd..e20e5db 100644 --- a/lean4/src/putnam_2022_b2.lean +++ b/lean4/src/putnam_2022_b2.lean @@ -8,8 +8,9 @@ abbrev putnam_2022_b2_solution : Set ℕ := sorry Let $\times$ represent the cross product in $\mathbb{R}^3$. For what positive integers $n$ does there exist a set $S \subset \mathbb{R}^3$ with exactly $n$ elements such that $S=\{v \times w:v,w \in S\}$? -/ theorem putnam_2022_b2 -(n : ℕ) -(Scross : Finset (Fin 3 → ℝ) → Prop) -(hScross : ∀ S : Finset (Fin 3 → ℝ), Scross S = (S = {u : Fin 3 → ℝ | ∃ v w : S, u = crossProduct v w})) -: (n > 0 ∧ ∃ S : Finset (Fin 3 → ℝ), S.card = n ∧ Scross S) ↔ n ∈ putnam_2022_b2_solution := + (n : ℕ) + (hn : 0 < n) + (P : Finset (Fin 3 → ℝ) → Prop) + (P_def : ∀ S : Finset (Fin 3 → ℝ), P S ↔ (S = {u : Fin 3 → ℝ | ∃ v w : S, u = crossProduct v w})) : + (∃ S : Finset (Fin 3 → ℝ), S.card = n ∧ P S) ↔ n ∈ putnam_2022_b2_solution := sorry diff --git a/lean4/src/putnam_2022_b3.lean b/lean4/src/putnam_2022_b3.lean index 6932f61..c445168 100644 --- a/lean4/src/putnam_2022_b3.lean +++ b/lean4/src/putnam_2022_b3.lean @@ -9,7 +9,7 @@ abbrev putnam_2022_b3_solution : Prop := sorry Assign to each positive real number a color, either red or blue. Let $D$ be the set of all distances $d>0$ such that there are two points of the same color at distance $d$ apart. Recolor the positive reals so that the numbers in $D$ are red and the numbers not in $D$ are blue. If we iterate this recoloring process, will we always end up with all the numbers red after a finite number of steps? -/ theorem putnam_2022_b3 -(recolor : (ℝ → Prop) → (ℝ → Prop)) -(hrecolor : ∀ isred : ℝ → Prop, recolor isred = (fun d : ℝ => ∃ p q : ℝ, p > 0 ∧ q > 0 ∧ p < q ∧ isred p = isred q ∧ q - p = d)) -: (∀ isred : ℝ → Prop, (∃ k : ℕ, ∀ p > 0, (recolor^[k] isred) p)) ↔ putnam_2022_b3_solution := + (recolor : (ℝ → Prop) → (ℝ → Prop)) + (hrecolor : ∀ isred d, recolor isred d ↔ ∃ p q : ℝ, p > 0 ∧ q > 0 ∧ p < q ∧ isred p = isred q ∧ q - p = d) : + (∀ isred, (∃ k, ∀ p > 0, (recolor^[k] isred) p)) ↔ putnam_2022_b3_solution := sorry diff --git a/lean4/src/putnam_2022_b4.lean b/lean4/src/putnam_2022_b4.lean index fd53d10..46ba242 100644 --- a/lean4/src/putnam_2022_b4.lean +++ b/lean4/src/putnam_2022_b4.lean @@ -9,11 +9,12 @@ abbrev putnam_2022_b4_solution : Set ℕ := sorry Find all integers $n$ with $n \geq 4$ for which there exists a sequence of distinct real numbers $x_1,\dots,x_n$ such that each of the sets $\{x_1,x_2,x_3\},\{x_2,x_3,x_4\},\dots,\{x_{n-2},x_{n-1},x_n\},\{x_{n-1},x_n,x_1\}$, and $\{x_n,x_1,x_2\}$ forms a $3$-term arithmetic progression when arranged in increasing order. -/ theorem putnam_2022_b4 -(n : ℕ) -(hn : n ≥ 4) -(ap3 : ℝ → ℝ → ℝ → Prop) -(xprog : (ℕ → ℝ) → Prop) -(hap3 : ∀ x0 x1 x2 : ℝ, ap3 x0 x1 x2 = ∀ o0 o1 o2 : ℝ, (o0 < o1 ∧ o1 < o2 ∧ ({o0, o1, o2} : Set ℝ) = {x0, x1, x2}) → (o1 - o0 = o2 - o1)) -(hxprog : ∀ x : ℕ → ℝ, xprog x = ((∀ i j : Fin n, i.1 ≠ j.1 → x i.1 ≠ x j.1) ∧ (∀ i : Fin n, ap3 (x i.1) (x ((i.1 + 1) % n)) (x ((i.1 + 2) % n))))) -: (∃ x : ℕ → ℝ, xprog x) ↔ n ∈ putnam_2022_b4_solution := + (n : ℕ) + (hn : n ≥ 4) + (ap3 : ℝ → ℝ → ℝ → Prop) + (xprog : (ℕ → ℝ) → Prop) + (hap3 : ∀ x0 x1 x2, ap3 x0 x1 x2 ↔ ∀ o0 o1 o2 : ℝ, (o0 < o1 ∧ o1 < o2 ∧ ({o0, o1, o2} : Set ℝ) = {x0, x1, x2}) → (o1 - o0 = o2 - o1)) + (hxprog : ∀ x, xprog x ↔ + ((∀ i j : Fin n, i.1 ≠ j.1 → x i.1 ≠ x j.1) ∧ (∀ i : Fin n, ap3 (x i.1) (x ((i.1 + 1) % n)) (x ((i.1 + 2) % n))))) : + (∃ x, xprog x) ↔ n ∈ putnam_2022_b4_solution := sorry diff --git a/lean4/src/putnam_2022_b6.lean b/lean4/src/putnam_2022_b6.lean index 0c14235..1723b2c 100644 --- a/lean4/src/putnam_2022_b6.lean +++ b/lean4/src/putnam_2022_b6.lean @@ -9,8 +9,8 @@ abbrev putnam_2022_b6_solution : Set (Set.Ioi (0 : ℝ) → Set.Ioi (0 : ℝ)) : Find all continuous functions $f:\mathbb{R}^+ \to \mathbb{R}^+$ such that $f(xf(y))+f(yf(x))=1+f(x+y)$ for all $x,y>0$. -/ theorem putnam_2022_b6 -(f : Set.Ioi (0 : ℝ) → Set.Ioi (0 : ℝ)) -(eq : Prop) -(heq : eq = ∃ fr : ℝ → ℝ, (∀ x : Set.Ioi (0 : ℝ), fr x = f x) ∧ (∀ x y : Set.Ioi (0 : ℝ), fr (x * fr y) + fr (y * fr x) = 1 + fr (x + y))) -: (Continuous f ∧ eq) ↔ f ∈ putnam_2022_b6_solution := + (f : Set.Ioi (0 : ℝ) → Set.Ioi (0 : ℝ)) + (hf : Continuous f) : + (∃ fr : ℝ → ℝ, (∀ x : Set.Ioi (0 : ℝ), fr x = f x) ∧ (∀ x y : Set.Ioi (0 : ℝ), fr (x * fr y) + fr (y * fr x) = 1 + fr (x + y))) + ↔ f ∈ putnam_2022_b6_solution := sorry diff --git a/lean4/src/putnam_2023_a1.lean b/lean4/src/putnam_2023_a1.lean index 724bceb..22dc715 100644 --- a/lean4/src/putnam_2023_a1.lean +++ b/lean4/src/putnam_2023_a1.lean @@ -8,8 +8,7 @@ abbrev putnam_2023_a1_solution : ℕ := sorry For a positive integer $n$, let $f_n(x) = \cos(x) \cos(2x) \cos(3x) \cdots \cos(nx)$. Find the smallest $n$ such that $|f_n''(0)| > 2023$. -/ theorem putnam_2023_a1 -(f : ℕ → ℝ → ℝ) -(hf : ∀ n > 0, f n = fun x : ℝ => ∏ i in Finset.Icc 1 n, Real.cos (i * x)) -: putnam_2023_a1_solution > 0 ∧ |iteratedDeriv 2 (f putnam_2023_a1_solution) 0| > 2023 ∧ -∀ n > 0, n < putnam_2023_a1_solution → (|iteratedDeriv 2 (f n) 0| ≤ 2023) := + (f : ℕ → ℝ → ℝ) + (hf : ∀ n > 0, f n = fun x : ℝ => ∏ i in Finset.Icc 1 n, Real.cos (i * x)) : + IsLeast {n | 0 < n ∧ |iteratedDeriv 2 (f n) 0| > 2023} putnam_2023_a1_solution := sorry diff --git a/lean4/src/putnam_2023_a3.lean b/lean4/src/putnam_2023_a3.lean index 266ad2f..0db6b1c 100644 --- a/lean4/src/putnam_2023_a3.lean +++ b/lean4/src/putnam_2023_a3.lean @@ -14,6 +14,8 @@ Determine the smallest positive real number $r$ such that there exist differenti \item[(e)] $f(r) = 0$. \end{enumerate} -/ theorem putnam_2023_a3 -: sInf {r > 0 | ∃ f g : ℝ → ℝ, Differentiable ℝ f ∧ Differentiable ℝ g ∧ -f 0 > 0 ∧ g 0 = 0 ∧ (∀ x : ℝ, |deriv f x| ≤ |g x| ∧ |deriv g x| ≤ |f x|) ∧ f r = 0} = putnam_2023_a3_solution := +: sInf {r > 0 | ∃ f g : ℝ → ℝ, + Differentiable ℝ f ∧ Differentiable ℝ g ∧ + (∀ x : ℝ, |deriv f x| ≤ |g x| ∧ |deriv g x| ≤ |f x|) ∧ + f 0 > 0 ∧ g 0 = 0 ∧ f r = 0} = putnam_2023_a3_solution := sorry diff --git a/lean4/src/putnam_2023_b1.lean b/lean4/src/putnam_2023_b1.lean index f2dc2cb..b3add41 100644 --- a/lean4/src/putnam_2023_b1.lean +++ b/lean4/src/putnam_2023_b1.lean @@ -12,12 +12,16 @@ theorem putnam_2023_b1 (m n : ℕ) (initcoins : ℕ → ℕ → Bool) (hinitcoins : initcoins = (fun i j : ℕ => (i ≤ m - 2 ∧ j ≤ n - 2 : Bool))) -(legalmove : (ℕ → ℕ → Bool) → (ℕ → ℕ → Bool) → Prop) -(hlegalmove : legalmove = (fun (coins1 : ℕ → ℕ → Bool) (coins2 : ℕ → ℕ → Bool) => ∃ i j : ℕ, i < m - 1 ∧ j < n - 1 ∧ -coins1 i j ∧ !coins1 (i + 1) j ∧ !coins1 i (j + 1) ∧ !coins1 (i + 1) (j + 1) ∧ !coins2 i j ∧ !coins2 (i + 1) j ∧ !coins2 i (j + 1) ∧ coins2 (i + 1) (j + 1) ∧ -(∀ i' j' : ℕ, ((i', j') ≠ (i, j) ∧ (i', j') ≠ (i + 1, j) ∧ (i', j') ≠ (i, j + 1) ∧ (i', j') ≠ (i + 1, j + 1)) → coins1 i' j' = coins2 i' j'))) -(legalseq : List (ℕ → ℕ → Bool) → Prop) -(hlegalseq : legalseq = (fun seq : List (ℕ → ℕ → Bool) => seq.length ≥ 1 ∧ seq[0]! = initcoins ∧ (∀ i < seq.length - 1, legalmove seq[i]! seq[i + 1]!))) +(IsLegalMove : (ℕ → ℕ → Bool) → (ℕ → ℕ → Bool) → Prop) +(IsLegalMove_def : ∀ coins1 coins2, + IsLegalMove coins1 coins2 ↔ + ∃ i j, i < m - 1 ∧ j < n - 1 ∧ + coins1 i j ∧ !coins1 (i + 1) j ∧ !coins1 i (j + 1) ∧ !coins1 (i + 1) (j + 1) ∧ + !coins2 i j ∧ !coins2 (i + 1) j ∧ !coins2 i (j + 1) ∧ coins2 (i + 1) (j + 1) ∧ + (∀ i' j', ((i', j') ≠ (i, j) ∧ (i', j') ≠ (i + 1, j) ∧ (i', j') ≠ (i, j + 1) ∧ (i', j') ≠ (i + 1, j + 1)) + → coins1 i' j' = coins2 i' j')) +(IsLegalSeq : List (ℕ → ℕ → Bool) → Prop) +(IsLegalSeq_def : ∀ seq, IsLegalSeq seq ↔ seq.length ≥ 1 ∧ seq[0]! = initcoins ∧ (∀ i < seq.length - 1, IsLegalMove seq[i]! seq[i + 1]!)) (mnpos : m ≥ 1 ∧ n ≥ 1) -: {config : ℕ → ℕ → Bool | ∃ seq : List (ℕ → ℕ → Bool), legalseq seq ∧ config = seq.getLast!}.encard = putnam_2023_b1_solution m n := +: {config : ℕ → ℕ → Bool | ∃ seq : List (ℕ → ℕ → Bool), IsLegalSeq seq ∧ config = seq.getLast!}.encard = putnam_2023_b1_solution m n := sorry diff --git a/lean4/src/putnam_2023_b4.lean b/lean4/src/putnam_2023_b4.lean index 5ac2a1a..8239cd6 100644 --- a/lean4/src/putnam_2023_b4.lean +++ b/lean4/src/putnam_2023_b4.lean @@ -16,17 +16,22 @@ For a nonnegative integer $n$ and a strictly increasing sequence of real numbers Considering all choices of $n$ and $t_0,t_1,\dots,t_n$ such that $t_k \geq t_{k-1}+1$ for $1 \leq k \leq n$, what is the least possible value of $T$ for which $f(t_0+T)=2023$? -/ theorem putnam_2023_b4 -(tne : ℕ → (ℕ → ℝ) → Set ℝ) -(fdiff flim fderiv1 fderiv2 fall : ℕ → (ℕ → ℝ) → (ℝ → ℝ) → Prop) -(tinc : ℕ → (ℕ → ℝ) → Prop) -(Tall : ℝ → Prop) -(htne : ∀ n : ℕ, ∀ ts : ℕ → ℝ, tne n ts = {t : ℝ | t > ts 0 ∧ ∀ i : Fin n, t ≠ ts (i.1 + 1)}) -(hfdiff : ∀ n : ℕ, ∀ ts : ℕ → ℝ, ∀ f : ℝ → ℝ, fdiff n ts f = (ContinuousOn f (Set.Ici (ts 0)) ∧ ContDiffOn ℝ 1 f (tne n ts) ∧ DifferentiableOn ℝ (derivWithin f (tne n ts)) (tne n ts))) -(hflim : ∀ n : ℕ, ∀ ts : ℕ → ℝ, ∀ f : ℝ → ℝ, flim n ts f = ∀ k : Fin (n + 1), Tendsto (derivWithin f (tne n ts)) (𝓝[>] (ts k.1)) (𝓝 0)) -(hfderiv1 : ∀ n : ℕ, ∀ ts : ℕ → ℝ, ∀ f : ℝ → ℝ, fderiv1 n ts f = ∀ k : Fin n, ∀ t ∈ Set.Ioo (ts k.1) (ts (k.1 + 1)), iteratedDerivWithin 2 f (tne n ts) t = k.1 + 1) -(hfderiv2 : ∀ n : ℕ, ∀ ts : ℕ → ℝ, ∀ f : ℝ → ℝ, fderiv2 n ts f = ∀ t > ts n, iteratedDerivWithin 2 f (tne n ts) t = n + 1) -(hfall : ∀ n : ℕ, ∀ ts : ℕ → ℝ, ∀ f : ℝ → ℝ, fall n ts f = (fdiff n ts f ∧ f (ts 0) = 0.5 ∧ flim n ts f ∧ fderiv1 n ts f ∧ fderiv2 n ts f)) -(htinc : ∀ n : ℕ, ∀ ts : ℕ → ℝ, tinc n ts = ∀ k : Fin n, ts (k.1 + 1) ≥ ts k.1 + 1) -(hTall : ∀ T : ℝ, Tall T = ((T ≥ 0) ∧ ∃ n : ℕ, ∃ ts : ℕ → ℝ, ∃ f : ℝ → ℝ, tinc n ts ∧ fall n ts f ∧ f (ts 0 + T) = 2023)) -: Tall putnam_2023_b4_solution ∧ ∀ T : ℝ, Tall T → T ≥ putnam_2023_b4_solution := -sorry + (tne : ℕ → (ℕ → ℝ) → Set ℝ) + (htne : ∀ n ts, tne n ts = {t | t > ts 0 ∧ ∀ i : Fin n, t ≠ ts (i.1 + 1)}) : + IsLeast + {(T : ℝ) | 0 ≤ T ∧ ∃ (n : ℕ) (ts : ℕ → ℝ) (f : ℝ → ℝ), + ∀ k : Fin n, ts (k.1 + 1) ≥ ts k.1 + 1 ∧ + ContinuousOn f (Set.Ici (ts 0)) ∧ + ContDiffOn ℝ 1 f (tne n ts) ∧ + DifferentiableOn ℝ (derivWithin f (tne n ts)) (tne n ts) ∧ + f (ts 0) = 0.5 ∧ + (∀ k : Fin (n + 1), + Tendsto (derivWithin f (tne n ts)) (𝓝[>] (ts k.1)) (𝓝 0)) ∧ + (∀ k : Fin n, + ∀ t ∈ Set.Ioo (ts k.1) (ts (k.1 + 1)), + iteratedDerivWithin 2 f (tne n ts) t = k.1 + 1) ∧ + (∀ t > ts n, + iteratedDerivWithin 2 f (tne n ts) t = n + 1) ∧ + f (ts 0 + T) = 2023} + putnam_2023_b4_solution := + sorry diff --git a/lean4/src/putnam_2023_b5.lean b/lean4/src/putnam_2023_b5.lean index 49b5198..3fbac45 100644 --- a/lean4/src/putnam_2023_b5.lean +++ b/lean4/src/putnam_2023_b5.lean @@ -8,8 +8,9 @@ abbrev putnam_2023_b5_solution : Set ℕ := sorry Determine which positive integers $n$ have the following property: For all integers $m$ that are relatively prime to $n$, there exists a permutation $\pi:\{1,2,\dots,n\} \to \{1,2,\dots,n\}$ such that $\pi(\pi(k)) \equiv mk \pmod{n}$ for all $k \in \{1,2,\dots,n\}$. -/ theorem putnam_2023_b5 -(n : ℕ) -(perm : Prop) -(hperm : perm = ∀ m : ℤ, IsRelPrime m n → ∃ p : Equiv.Perm (Fin n), ∀ k : Fin n, (p (p k)).1 + 1 ≡ m * (k.1 + 1) [ZMOD n]) -: (n > 0 ∧ perm) ↔ n ∈ putnam_2023_b5_solution := + (n : ℕ) + (hn : 0 < n) : + n ∈ putnam_2023_b5_solution ↔ + (∀ m : ℤ, IsRelPrime m n → ∃ p : Equiv.Perm (Fin n), + ∀ k : Fin n, (p (p k)).1 + 1 ≡ m * (k.1 + 1) [ZMOD n]) := sorry