X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdama%2Fdama%2Fmodels%2Fq_support.ma;h=9d73f7ab9b209fdfa887c3aed98764d3cd737265;hb=200bb81b91b7c4ebf479906d09c290353c763289;hp=82832c4bc64ee90cea02b82e867f4a321cc83328;hpb=052140f5d87dabd49f798b5cf42e35e1df411db3;p=helm.git diff --git a/helm/software/matita/contribs/dama/dama/models/q_support.ma b/helm/software/matita/contribs/dama/dama/models/q_support.ma index 82832c4bc..9d73f7ab9 100644 --- a/helm/software/matita/contribs/dama/dama/models/q_support.ma +++ b/helm/software/matita/contribs/dama/dama/models/q_support.ma @@ -33,6 +33,7 @@ axiom q_minus_r: ∀x,y. y + Qpos x = y - Qneg x. axiom q_plus_assoc: ∀x,y,z.x + (y + z) = x + y + z. axiom q_elim_minus: ∀x,y.x - y = x + Qopp y. axiom q_elim_opp: ∀x,y.x - Qopp y = x + y. +axiom q_minus_distrib:∀x,y,z:Q.x - (y + z) = x - y - z. (* order over Q *) axiom qlt : ℚ → ℚ → CProp. @@ -68,6 +69,8 @@ axiom q_pos_lt_OQ: ∀x.OQ < Qpos x. axiom q_le_plus_trans: ∀x,y:Q. OQ ≤ x → OQ ≤ y → OQ ≤ x + y. axiom q_le_S: ∀x,y,z.OQ ≤ x → x + y ≤ z → y ≤ z. axiom q_eq_to_le: ∀x,y. x = y → x ≤ y. +axiom q_le_OQ_Qpos: ∀x.OQ ≤ Qpos x. + inductive q_le_elimination (a,b:ℚ) : CProp ≝ | q_le_from_eq : a = b → q_le_elimination a b @@ -90,3 +93,33 @@ axiom q_d_sym: ∀x,y. ⅆ[x,y] = ⅆ[y,x]. (* integral part *) axiom nat_of_q: ℚ → nat. +(* derived *) +lemma q_lt_canc_plus_r: + ∀x,y,z:Q.x + z < y + z → x < y. +intros; rewrite < (q_plus_OQ y); rewrite < (q_plus_minus z); +rewrite > q_elim_minus; rewrite > q_plus_assoc; +apply q_lt_plus; rewrite > q_elim_opp; assumption; +qed. + +lemma q_lt_inj_plus_r: + ∀x,y,z:Q.x < y → x + z < y + z. +intros; apply (q_lt_canc_plus_r ?? (Qopp z)); +do 2 (rewrite < q_plus_assoc;rewrite < q_elim_minus); +rewrite > q_plus_minus; +do 2 rewrite > q_plus_OQ; assumption; +qed. + +lemma q_le_inj_plus_r: + ∀x,y,z:Q.x ≤ y → x + z ≤ y + z. +intros;cases (q_le_cases ?? H); +[1: rewrite > H1; apply q_eq_to_le; reflexivity; +|2: apply q_lt_to_le; apply q_lt_inj_plus_r; assumption;] +qed. + +lemma q_le_canc_plus_r: + ∀x,y,z:Q.x + z ≤ y + z → x ≤ y. +intros; lapply (q_le_inj_plus_r ?? (Qopp z) H) as H1; +do 2 rewrite < q_plus_assoc in H1; +rewrite < q_elim_minus in H1; rewrite > q_plus_minus in H1; +do 2 rewrite > q_plus_OQ in H1; assumption; +qed.