diff --git a/coq/src/commented_problems.v b/coq/src/commented_problems.v index 291ce5fa..dae049dc 100644 --- a/coq/src/commented_problems.v +++ b/coq/src/commented_problems.v @@ -76,4 +76,29 @@ Theorem putnam_2010_a5: forall (a b: R -> R -> R), cross_product a b = a. Proof. Admitted. -End putnam_2010_a5. *) \ No newline at end of file +End putnam_2010_a5. *) + +(* Require Import Reals +GeoCoq.Main.Tarski_dev.Ch16_coordinates_with_functions +GeoCoq.Axioms.Definitions +GeoCoq.Main.Highschool.triangles. +Context `{T2D:Tarski_2D} `{TE:@Tarski_euclidean Tn TnEQD}. +Open Scope R. +Definition putnam_2003_b5_solution (pt_to_R : Tpoint -> (R * R)) (dist : Tpoint -> Tpoint -> R) (P Op : Tpoint) := sqrt 3 * (1 - (dist P Op) ^ 2 - 1). +Theorem putnam_2003_b5 + (pt_to_R : Tpoint -> (R * R)) + (F_to_R : F -> R) + (dist : Tpoint -> Tpoint -> R := fun A B => let (a, b) := pt_to_R A in let (c, d) := pt_to_R B in dist_euc a b c d) + (Triangle : Tpoint -> Tpoint -> Tpoint -> Prop := fun x y z => ~ Col x y z) (* copied from GeoCoq.Axioms.euclidean_axioms *) + (A B C Op Op' P: Tpoint) + (fixpoint : dist Op Op' = R1) + (hABC : OnCircle A Op Op' /\ OnCircle B Op Op' /\ OnCircle C Op Op') + (hABC' : Main.Highschool.triangles.equilateral A B C) + (hp : InCircle P Op Op') + (a : R := dist P A) + (b : R := dist P B) + (c : R := dist P C) + : exists (A' B' C' : Tpoint) (D: Cs O E A' B' C'), + Triangle A' B' C' /\ dist A' B' = a /\ dist B' C' = b /\ dist C' A' = c /\ + F_to_R (signed_area A' B' C' D A' B' C') = (putnam_2003_b5_solution pt_to_R dist P Op). +Proof. Admitted. *) diff --git a/coq/src/putnam_2002_a1.v b/coq/src/putnam_2002_a1.v index 7868eb5f..61112d51 100644 --- a/coq/src/putnam_2002_a1.v +++ b/coq/src/putnam_2002_a1.v @@ -1,8 +1,10 @@ Require Import Reals Factorial Coquelicot.Coquelicot. Definition putnam_2002_a1_solution (k n: nat) := Rpower (-1 * INR k) (INR n) * INR (fact n). Theorem putnam_2002_a1 + (k : nat) (p : (nat -> R) -> R -> nat -> R := fun a x n => sum_n (fun i => a i * x ^ i) n) - : forall (N k: nat), gt k 0 -> exists (a: nat -> R) (n: nat), forall (x: R), - (Derive_n (fun x => 1 / (x ^ k - 1)) N) x = (p a x n) / (x ^ k - 1) ^ (n + 1) -> - p a x 1%nat = putnam_2002_a1_solution k n. + (kpos : gt k 0) + : forall (N: nat), forall (a: nat -> R) (n: nat), + (forall (x: R), (Derive_n (fun x => 1 / (x ^ k - 1)) N) x = (p a x n) / (x ^ k - 1) ^ (n + 1)) -> + p a 1 n = putnam_2002_a1_solution k n. Proof. Admitted. diff --git a/coq/src/putnam_2002_b3.v b/coq/src/putnam_2002_b3.v index cd842399..9fb0ceec 100644 --- a/coq/src/putnam_2002_b3.v +++ b/coq/src/putnam_2002_b3.v @@ -1,5 +1,5 @@ Require Import Reals Coquelicot.Coquelicot. Open Scope R. Theorem putnam_2002_b3 - : forall (n: nat), ge n 1 -> let n := INR n in 1 / (2 * n * exp 1) < 1 / (exp 1) - Rpower (1 - 1 / n) n < 1 / (n * (exp 1)). + : forall (n: nat), gt n 1 -> let n := INR n in 1 / (2 * n * exp 1) < 1 / (exp 1) - Rpower (1 - 1 / n) n < 1 / (n * (exp 1)). Proof. Admitted. diff --git a/coq/src/putnam_2003_a1.v b/coq/src/putnam_2003_a1.v index 9356672b..bb6da201 100644 --- a/coq/src/putnam_2003_a1.v +++ b/coq/src/putnam_2003_a1.v @@ -1,8 +1,8 @@ Require Import Nat List Ensembles Finite_sets Coquelicot.Coquelicot. -Definition putnam_2003_a1_solution (n: nat) := n. +Definition putnam_2003_a1_solution : nat -> nat := (fun n : nat => n). Theorem putnam_2003_a1 (n: nat) (hn : n > 0) - (E: Ensemble (list nat) := fun l => forall (i j: nat), i < length l /\ j < length l /\ i < j -> nth i l 0 <= nth j l 0 /\ fold_left add l 0 = n) + (E: Ensemble (list nat) := fun l => fold_left add l 0 = n /\ (forall i : nat, i < length l -> nth i l 0 > 0) /\ (forall (i j: nat), i < length l /\ j < length l /\ i < j -> nth i l 0 <= nth j l 0) /\ last l 0 <= hd 0 l + 1) : cardinal (list nat) E (putnam_2003_a1_solution n). Proof. Admitted. diff --git a/coq/src/putnam_2003_a2.v b/coq/src/putnam_2003_a2.v index 000b6df8..bfd920dc 100644 --- a/coq/src/putnam_2003_a2.v +++ b/coq/src/putnam_2003_a2.v @@ -6,7 +6,11 @@ Theorem putnam_2003_a2 | _, nil => nil | h1 :: t1, h2 :: t2 => (h1 + h2) :: suml t1 t2 end) - : forall (n: nat) (a b: list R), length a = n /\ length b = n -> - (fold_left Rmult a 1) ^ (1 / n) + (fold_left Rmult b 1) ^ (1 / n) <= + (n : nat) + (a b : list R) + (npos : ge n 1) + (ablen : length a = n /\ length b = n) + (abnneg : forall i : nat, lt i n -> nth i a 0 >= 0 /\ nth i b 0 >= 0) + : (fold_left Rmult a 1) ^ (1 / n) + (fold_left Rmult b 1) ^ (1 / n) <= (fold_left Rmult (Suml a b) 1) ^ (1 / n). Proof. Admitted. diff --git a/coq/src/putnam_2003_a3.v b/coq/src/putnam_2003_a3.v index fe21831b..6420ef70 100644 --- a/coq/src/putnam_2003_a3.v +++ b/coq/src/putnam_2003_a3.v @@ -2,5 +2,5 @@ Require Import Reals Coquelicot.Coquelicot. Definition putnam_2003_a3_solution := 2 * sqrt 2 - 1. Theorem putnam_2003_a3 (f : R -> R := fun x => Rabs (sin x + cos x + tan x + 1 / tan x + 1 / cos x + 1 / sin x)) - : exists (minx: R), forall (x: R), f minx <= f x -> f minx = putnam_2003_a3_solution. + : (exists x : R, f x = putnam_2003_a3_solution) /\ (forall x : R, f x >= putnam_2003_a3_solution). Proof. Admitted. diff --git a/coq/src/putnam_2003_a4.v b/coq/src/putnam_2003_a4.v index 2d3b5082..77306316 100644 --- a/coq/src/putnam_2003_a4.v +++ b/coq/src/putnam_2003_a4.v @@ -2,7 +2,7 @@ Require Import Reals Coquelicot.Coquelicot. Theorem putnam_2003_a4 (a b c A B C: R) (ha : a <> 0) - (hA :A <> 0) + (hA : A <> 0) (hp : forall (x: R), Rabs (a * x ^ 2 + b * x + c) <= Rabs (A * x ^ 2 + B * x + C)) : Rabs (b ^ 2 - 4 * a * c) <= Rabs (B ^ 2 - 4 * A * C). Proof. Admitted. diff --git a/coq/src/putnam_2003_b1.v b/coq/src/putnam_2003_b1.v index f9edb7a4..39c7ad94 100644 --- a/coq/src/putnam_2003_b1.v +++ b/coq/src/putnam_2003_b1.v @@ -1,7 +1,8 @@ Require Import Reals Coquelicot.Coquelicot. -Definition putnam_2003_b1_solution := True. +Definition putnam_2003_b1_solution := False. Theorem putnam_2003_b1 (p : (nat -> R) -> R -> nat -> R := fun coeff x n => sum_n (fun i => coeff i * x ^ i) n) - : exists (coeffa coeffb coeffc coeffd: nat -> R) (na nb nc nd: nat), forall (x y: R), - 1 + x * y * (x * y) ^ 2 = (p coeffa x na) * (p coeffc y nc) + (p coeffb x nb) * (p coeffd y nd). + : (exists (coeffa coeffb coeffc coeffd: nat -> R) (na nb nc nd: nat), forall (x y: R), + 1 + x * y + x ^ 2 * y ^ 2 = (p coeffa x na) * (p coeffc y nc) + (p coeffb x nb) * (p coeffd y nd)) + <-> putnam_2003_b1_solution. Proof. Admitted. diff --git a/coq/src/putnam_2003_b3.v b/coq/src/putnam_2003_b3.v index 28d0784f..4c262935 100644 --- a/coq/src/putnam_2003_b3.v +++ b/coq/src/putnam_2003_b3.v @@ -2,13 +2,15 @@ Require Import Nat List Reals Coquelicot.Coquelicot. Theorem putnam_2003_b3 (lcmn := fix lcm_n (args : list nat) : nat := match args with - | nil => 0%nat + | nil => 1%nat | h :: args' => div (h * (lcm_n args')) (gcd h (lcm_n args')) end) (prodn := fix prod_n (m: nat -> R) (n : nat) : R := match n with | O => m 0%nat - | S n' => m n' * prod_n m n' + | S n' => m n * prod_n m n' end) - : forall (n: nat), gt n 0 -> INR (fact n) = prodn (fun i => INR (lcmn (seq 0 (div n (i + 1))))) n. + (n : nat) + (npos : gt n 0) + : INR (fact n) = prodn (fun i => INR (lcmn (seq 1 (div n (i + 1))))) (sub n 1). Proof. Admitted. diff --git a/coq/src/putnam_2003_b4.v b/coq/src/putnam_2003_b4.v index 14a0d5aa..f0e8ac54 100644 --- a/coq/src/putnam_2003_b4.v +++ b/coq/src/putnam_2003_b4.v @@ -1,13 +1,13 @@ Require Import Reals ZArith Coquelicot.Coquelicot. Theorem putnam_2003_b4 (a b c d e: Z) + (r1 r2 r3 r4: R) (ha : ~ Z.eq a 0) : let a := IZR a in let b := IZR b in let c := IZR c in let d := IZR d in let e := IZR e in - exists (r1 r2 r3 r4: R), forall (z: R), - a * z ^ 4 + b * z ^ 3 + c * z ^ 2 + d * z + e = a * (z - r1) * (z - r2) * (z - r3) * (z - r4) -> - (exists (p q: Z), r1 + r2 = IZR p / IZR q) /\ r1 + r2 <> r3 + r4 -> exists (p q: Z), r1 * r2 = IZR p / IZR q. + (forall (z: R), a * z ^ 4 + b * z ^ 3 + c * z ^ 2 + d * z + e = a * (z - r1) * (z - r2) * (z - r3) * (z - r4)) -> + ((exists (p q: Z), r1 + r2 = IZR p / IZR q) /\ r1 + r2 <> r3 + r4) -> (exists (p q: Z), r1 * r2 = IZR p / IZR q). Proof. Admitted. diff --git a/coq/src/putnam_2003_b5.v b/coq/src/putnam_2003_b5.v deleted file mode 100644 index 5ae16f2a..00000000 --- a/coq/src/putnam_2003_b5.v +++ /dev/null @@ -1,24 +0,0 @@ -Require Import Reals -GeoCoq.Main.Tarski_dev.Ch16_coordinates_with_functions -GeoCoq.Axioms.Definitions -GeoCoq.Main.Highschool.triangles. -Context `{T2D:Tarski_2D} `{TE:@Tarski_euclidean Tn TnEQD}. -Open Scope R. -Definition putnam_2003_b5_solution (pt_to_R : Tpoint -> (R * R)) (dist : Tpoint -> Tpoint -> R) (P Op : Tpoint) := sqrt 3 * (1 - (dist P Op) ^ 2 - 1). -Theorem putnam_2003_b5 - (pt_to_R : Tpoint -> (R * R)) - (F_to_R : F -> R) - (dist : Tpoint -> Tpoint -> R := fun A B => let (a, b) := pt_to_R A in let (c, d) := pt_to_R B in dist_euc a b c d) - (Triangle : Tpoint -> Tpoint -> Tpoint -> Prop := fun x y z => ~ Col x y z) (* copied from GeoCoq.Axioms.euclidean_axioms *) - (A B C Op Op' P: Tpoint) - (fixpoint : dist Op Op' = R1) - (hABC : OnCircle A Op Op' /\ OnCircle B Op Op' /\ OnCircle C Op Op') - (hABC' : Main.Highschool.triangles.equilateral A B C) - (hp : InCircle P Op Op') - (a : R := dist P A) - (b : R := dist P B) - (c : R := dist P C) - : exists (A' B' C' : Tpoint) (D: Cs O E A' B' C'), - Triangle A' B' C' /\ dist A' B' = a /\ dist B' C' = b /\ dist C' A' = c /\ - F_to_R (signed_area A' B' C' D A' B' C') = (putnam_2003_b5_solution pt_to_R dist P Op). -Proof. Admitted. \ No newline at end of file diff --git a/coq/src/putnam_2003_b6.v b/coq/src/putnam_2003_b6.v index 783e7f46..da90c81d 100644 --- a/coq/src/putnam_2003_b6.v +++ b/coq/src/putnam_2003_b6.v @@ -1,5 +1,6 @@ Require Import Reals Coquelicot.Coquelicot. Theorem putnam_2003_b6 - : forall (f: R -> R) (x: R), 0 <= x <= 1 -> continuity_pt f x -> - RInt (fun x => RInt (fun y => Rabs (f x + f y)) 0 1) 0 1 >= RInt (fun x => Rabs (f x)) 0 1. + (f : R -> R) + (hf : continuity f) + : RInt (fun x => RInt (fun y => Rabs (f x + f y)) 0 1) 0 1 >= RInt (fun x => Rabs (f x)) 0 1. Proof. Admitted. diff --git a/isabelle/putnam_2002_a2.thy b/isabelle/putnam_2002_a2.thy index 87e84caa..c7444b4c 100644 --- a/isabelle/putnam_2002_a2.thy +++ b/isabelle/putnam_2002_a2.thy @@ -1,6 +1,7 @@ theory putnam_2002_a2 imports Complex_Main "HOL-Analysis.Finite_Cartesian_Product" "HOL-Analysis.Linear_Algebra" +"HOL-Analysis.Elementary_Metric_Spaces" begin theorem putnam_2002_a2: diff --git a/isabelle/putnam_2002_a3.thy b/isabelle/putnam_2002_a3.thy index db1f3160..a6d04ce6 100644 --- a/isabelle/putnam_2002_a3.thy +++ b/isabelle/putnam_2002_a3.thy @@ -3,7 +3,7 @@ begin theorem putnam_2002_a3: fixes n Tn :: "int" - defines "Tn \ card {S :: int set. S \ {1..16} \ S \ {} \ (\ k :: int. k = (real 1)/(card S) * (\ s \ S. s))}" + defines "Tn \ card {S :: int set. S \ {1..n} \ S \ {} \ (\ k :: int. k = (real 1)/(card S) * (\ s \ S. s))}" assumes hn : "n \ 2" shows "even (Tn - n)" sorry diff --git a/isabelle/putnam_2002_b5.thy b/isabelle/putnam_2002_b5.thy index 1bee3d35..4d07693d 100644 --- a/isabelle/putnam_2002_b5.thy +++ b/isabelle/putnam_2002_b5.thy @@ -5,8 +5,7 @@ begin fun digits :: "nat \ nat \ nat list" where "digits b n = (if b < 2 then (replicate n 1) else (if n < b then [n] else [n mod b] @ digits b (n div b)))" theorem putnam_2002_b5: - shows "\ n :: nat. card {b :: nat. length (digits b n) = 3 \ digits b n = rev (digits b n)} \ 2002" + shows "\ n :: nat. card {b :: nat. b \ 1 \ length (digits b n) = 3 \ digits b n = rev (digits b n)} \ 2002" sorry end - diff --git a/isabelle/putnam_2003_a2.thy b/isabelle/putnam_2003_a2.thy index 497e1ab6..f9ed8ac4 100644 --- a/isabelle/putnam_2003_a2.thy +++ b/isabelle/putnam_2003_a2.thy @@ -4,7 +4,8 @@ begin (* Note: Boosted domain to infinite set *) theorem putnam_2003_a2: fixes n::nat and a b::"nat\real" - assumes abnneg : "\i \ {0.. 0 \ b i \ 0" + assumes npos: "n \ 1" + and abnneg: "\i \ {0.. 0 \ b i \ 0" shows "((\i=0..i=0.. ((\i=0.. real" assumes hf : "continuous_on UNIV f" shows "set_lebesgue_integral lebesgue {(x, y). 0 \ x \ x \ 1 \ 0 \ y \ y \ 1} (\ t :: real \ real. \f (fst t) + f (snd t)\) - \ interval_lebesgue_integral lebesgue 0 1 f" + \ interval_lebesgue_integral lebesgue 0 1 (\x::real. \f x\)" sorry end \ No newline at end of file diff --git a/lean4/src/putnam_2002_b5.lean b/lean4/src/putnam_2002_b5.lean index 66febe98..30ca1bd9 100644 --- a/lean4/src/putnam_2002_b5.lean +++ b/lean4/src/putnam_2002_b5.lean @@ -4,5 +4,5 @@ open BigOperators open Nat Set Topology Filter theorem putnam_2002_b5 -: ∃ n : ℕ, {b : ℕ | (Nat.digits b n).length = 3 ∧ List.Palindrome (Nat.digits b n)}.ncard ≥ 2002 := +: ∃ n : ℕ, {b : ℕ | b ≥ 1 ∧ (Nat.digits b n).length = 3 ∧ List.Palindrome (Nat.digits b n)}.ncard ≥ 2002 := sorry