X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2FRELATIONAL%2FNPlus%2Fmonoid.ma;h=7a9bb8da047c4ee1443eef83ae70691cc6f5c504;hb=d9824956d9132109ed5f23380a0a1f9c5181d18a;hp=a521c9a1dfef29b49f0176f484f0bf4e9eac3383;hpb=eefb7b4c9f5c4c531199c95e4bb72d8b8c88bc2e;p=helm.git diff --git a/helm/software/matita/contribs/RELATIONAL/NPlus/monoid.ma b/helm/software/matita/contribs/RELATIONAL/NPlus/monoid.ma index a521c9a1d..7a9bb8da0 100644 --- a/helm/software/matita/contribs/RELATIONAL/NPlus/monoid.ma +++ b/helm/software/matita/contribs/RELATIONAL/NPlus/monoid.ma @@ -12,38 +12,34 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/RELATIONAL/NPlus/monoid". + include "NPlus/fun.ma". (* Monoidal properties ******************************************************) -theorem nplus_zero_1: \forall q. zero + q == q. - intros. elim q; clear q; autobatch. +theorem nplus_zero_1: ∀q. zero ⊕ q ≍ q. + intros; elim q; clear q; autobatch. qed. -theorem nplus_succ_1: \forall p,q,r. (p + q == r) \to - (succ p) + q == (succ r). - intros. elim H; clear H q r; autobatch. +theorem nplus_succ_1: ∀p,q,r. p ⊕ q ≍ r → succ p ⊕ q ≍ succ r. + intros; elim H; clear H q r; autobatch. qed. -theorem nplus_comm: \forall p, q, x. (p + q == x) \to - \forall y. (q + p == y) \to x = y. +theorem nplus_comm: ∀p, q, x. p ⊕ q ≍ x → ∀y. q ⊕ p ≍ y → x = y. intros 4; elim H; clear H q x; [ lapply linear nplus_inv_zero_1 to H1 | lapply linear nplus_inv_succ_1 to H3. decompose ]; destruct; autobatch. qed. -theorem nplus_comm_rew: \forall p,q,r. (p + q == r) \to q + p == r. - intros. elim H; clear H q r; autobatch. +theorem nplus_comm_rew: ∀p,q,r. p ⊕ q ≍ r → q ⊕ p ≍ r. + intros; elim H; clear H q r; autobatch. qed. -theorem nplus_ass: \forall p1, p2, r1. (p1 + p2 == r1) \to - \forall p3, s1. (r1 + p3 == s1) \to - \forall r3. (p2 + p3 == r3) \to - \forall s3. (p1 + r3 == s3) \to s1 = s3. - intros 4. elim H; clear H p2 r1; +theorem nplus_ass: ∀p1, p2, r1. p1 ⊕ p2 ≍ r1 → ∀p3, s1. r1 ⊕ p3 ≍ s1 → + ∀r3. p2 ⊕ p3 ≍ r3 → ∀s3. p1 ⊕ r3 ≍ s3 → s1 = s3. + intros 4; elim H; clear H p2 r1; [ lapply linear nplus_inv_zero_1 to H2. destruct. lapply nplus_mono to H1, H3. destruct. autobatch | lapply linear nplus_inv_succ_1 to H3. decompose. destruct. @@ -54,19 +50,15 @@ qed. (* Corollaries of functional properties **************************************) -theorem nplus_inj_2: \forall p, q1, r. (p + q1 == r) \to - \forall q2. (p + q2 == r) \to q1 = q2. +theorem nplus_inj_2: ∀p, q1, r. p ⊕ q1 ≍ r → ∀q2. p ⊕ q2 ≍ r → q1 = q2. intros. autobatch. qed. (* Corollaries of nonoidal properties ***************************************) -theorem nplus_comm_1: \forall p1, q, r1. (p1 + q == r1) \to - \forall p2, r2. (p2 + q == r2) \to - \forall x. (p2 + r1 == x) \to - \forall y. (p1 + r2 == y) \to - x = y. - intros 4. elim H; clear H q r1; +theorem nplus_comm_1: ∀p1, q, r1. p1 ⊕ q ≍ r1 → ∀p2, r2. p2 ⊕ q ≍ r2 → + ∀x. p2 ⊕ r1 ≍ x → ∀y. p1 ⊕ r2 ≍ y → x = y. + intros 4; elim H; clear H q r1; [ lapply linear nplus_inv_zero_2 to H1 | lapply linear nplus_inv_succ_2 to H3. lapply linear nplus_inv_succ_2 to H4. decompose. destruct. @@ -74,10 +66,9 @@ theorem nplus_comm_1: \forall p1, q, r1. (p1 + q == r1) \to ]; destruct; autobatch. qed. -theorem nplus_comm_1_rew: \forall p1,q,r1. (p1 + q == r1) \to - \forall p2,r2. (p2 + q == r2) \to - \forall s. (p1 + r2 == s) \to (p2 + r1 == s). - intros 4. elim H; clear H q r1; +theorem nplus_comm_1_rew: ∀p1,q,r1. p1 ⊕ q ≍ r1 → ∀p2,r2. p2 ⊕ q ≍ r2 → + ∀s. p1 ⊕ r2 ≍ s → p2 ⊕ r1 ≍ s. + intros 4; elim H; clear H q r1; [ lapply linear nplus_inv_zero_2 to H1. destruct | lapply linear nplus_inv_succ_2 to H3. decompose. destruct. lapply linear nplus_inv_succ_2 to H4. decompose. destruct @@ -86,22 +77,22 @@ qed. (* theorem nplus_shift_succ_sx: \forall p,q,r. - (p + (succ q) == r) \to (succ p) + q == r. + (p \oplus (succ q) \asymp r) \to (succ p) \oplus q \asymp r. intros. lapply linear nplus_inv_succ_2 to H as H0. decompose. destruct. auto new timeout=100. qed. theorem nplus_shift_succ_dx: \forall p,q,r. - ((succ p) + q == r) \to p + (succ q) == r. + ((succ p) \oplus q \asymp r) \to p \oplus (succ q) \asymp r. intros. lapply linear nplus_inv_succ_1 to H as H0. decompose. destruct. auto new timeout=100. qed. -theorem nplus_trans_1: \forall p,q1,r1. (p + q1 == r1) \to - \forall q2,r2. (r1 + q2 == r2) \to - \exists q. (q1 + q2 == q) \land p + q == r2. +theorem nplus_trans_1: \forall p,q1,r1. (p \oplus q1 \asymp r1) \to + \forall q2,r2. (r1 \oplus q2 \asymp r2) \to + \exists q. (q1 \oplus q2 \asymp q) \land p \oplus q \asymp r2. intros 2; elim q1; clear q1; intros; [ lapply linear nplus_inv_zero_2 to H as H0. destruct. @@ -114,9 +105,8 @@ theorem nplus_trans_1: \forall p,q1,r1. (p + q1 == r1) \to ]; apply ex_intro; [| auto new timeout=100 || auto new timeout=100 ]. (**) qed. -theorem nplus_trans_2: \forall p1,q,r1. (p1 + q == r1) \to - \forall p2,r2. (p2 + r1 == r2) \to - \exists p. (p1 + p2 == p) \land p + q == r2. +theorem nplus_trans_2: ∀p1,q,r1. p1 ⊕ q ≍ r1 → ∀p2,r2. p2 ⊕ r1 ≍ r2 → + ∃p. p1 ⊕ p2 ≍ p ∧ p ⊕ q ≍ r2. intros 2; elim q; clear q; intros; [ lapply linear nplus_inv_zero_2 to H as H0. destruct @@ -126,6 +116,6 @@ theorem nplus_trans_2: \forall p1,q,r1. (p1 + q == r1) \to decompose. destruct. lapply linear H to H4, H3 as H0. decompose. - ]; apply ex_intro; [| auto new timeout=100 || auto new timeout=100 ]. (**) + ]; autobatch. apply ex_intro; [| auto new timeout=100 || auto new timeout=100 ]. (**) qed. *)