(************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) (* √a == b := sqrt_unique. Lemma sqrt_square : forall a, √(a*a) == a. Proof. wrap sqrt_square. Qed. Definition sqrt_le_mono : forall a b, a<=b -> √a <= √b := sqrt_le_mono. Definition sqrt_lt_cancel : forall a b, √a < √b -> a < b := sqrt_lt_cancel. Lemma sqrt_le_square : forall a b, b*b<=a <-> b <= √a. Proof. wrap sqrt_le_square. Qed. Lemma sqrt_lt_square : forall a b, a √a < b. Proof. wrap sqrt_lt_square. Qed. Definition sqrt_0 := sqrt_0. Definition sqrt_1 := sqrt_1. Definition sqrt_2 := sqrt_2. Definition sqrt_lt_lin : forall a, 1 √a