diff --git a/theories/algebra/frac.v b/theories/algebra/frac.v index 96f1a3d6dec1b2edc7191f701444afbaada3c65d..7dfd5b92e87e9ca624fbce5e9b7254c71ff2f4c4 100644 --- a/theories/algebra/frac.v +++ b/theories/algebra/frac.v @@ -11,14 +11,19 @@ Instance frac_valid : Valid frac := λ x, (x ≤ 1)%Qc. Instance frac_pcore : PCore frac := λ _, None. Instance frac_op : Op frac := λ x y, (x + y)%Qp. -Lemma frac_included (x y : frac) : x ≼ y ↔ (x < y)%Qc. +(* TODO: Find better place for this lemma. *) +Lemma Qp_le_sum (x y : Qp) : (x < y)%Qc ↔ (∃ z, y = x + z)%Qp. Proof. split. - - intros [z ->%leibniz_equiv]; simpl. - rewrite -{1}(Qcplus_0_r x). apply Qcplus_lt_mono_l, Qp_prf. - intros Hlt%Qclt_minus_iff. exists (mk_Qp (y - x) Hlt). apply Qp_eq; simpl. by rewrite (Qcplus_comm y) Qcplus_assoc Qcplus_opp_r Qcplus_0_l. + - intros [z ->%leibniz_equiv]; simpl. + rewrite -{1}(Qcplus_0_r x). apply Qcplus_lt_mono_l, Qp_prf. Qed. + +Lemma frac_included (x y : frac) : x ≼ y ↔ (x < y)%Qc. +Proof. symmetry. exact: Qp_le_sum. Qed. + Corollary frac_included_weak (x y : frac) : x ≼ y → (x ≤ y)%Qc. Proof. intros ?%frac_included. auto using Qclt_le_weak. Qed.