X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdama%2Fdama%2Fnat_ordered_set.ma;h=26e2f0d2953a75ac9e3706babd7ec22a3b4f695d;hb=e2a6d130d5274760f4591197bbbc66c3191e8a6a;hp=23865418c48fb898f8eeb896ed0894d89ac5e080;hpb=7e33e23e18dc5d008b3b3dc0052aa4d7b236415e;p=helm.git diff --git a/helm/software/matita/contribs/dama/dama/nat_ordered_set.ma b/helm/software/matita/contribs/dama/dama/nat_ordered_set.ma index 23865418c..26e2f0d29 100644 --- a/helm/software/matita/contribs/dama/dama/nat_ordered_set.ma +++ b/helm/software/matita/contribs/dama/dama/nat_ordered_set.ma @@ -44,18 +44,18 @@ cases (nat_discriminable x z); [2: left; assumption] cases H1; clear H1; qed. lemma nat_ordered_set : ordered_set. -apply (mk_ordered_set ? nat_excess); -[1: intro x; intro; apply (not_le_Sn_n ? H); -|2: apply nat_excess_cotransitive] +letin hos ≝ (mk_half_ordered_set nat (λT,R:Type.λf:T→T→R.f) ? nat_excess ? nat_excess_cotransitive); +[ intros; left; intros; reflexivity; +| intro x; intro H; apply (not_le_Sn_n ? H);] +constructor 1; apply hos; qed. -notation "\naturals" non associative with precedence 90 for @{'nat}. -interpretation "ordered set N" 'nat = nat_ordered_set. +interpretation "ordered set N" 'N = nat_ordered_set. alias id "le" = "cic:/matita/nat/orders/le.ind#xpointer(1/1)". lemma os_le_to_nat_le: ∀a,b:nat_ordered_set.a ≤ b → le a b. -intros; normalize in H; apply (not_lt_to_le ?? H); +intros; normalize in H; apply (not_lt_to_le b a H); qed. lemma nat_le_to_os_le: @@ -63,15 +63,3 @@ lemma nat_le_to_os_le: intros 3; apply (le_to_not_lt a b);assumption; qed. -lemma nat_lt_to_os_lt: - ∀a,b:nat_ordered_set.a < b → lt nat_ordered_set a b. -intros 3; split; -[1: apply nat_le_to_os_le; apply lt_to_le;assumption; -|2: right; apply H;] -qed. - -lemma os_lt_to_nat_lt: - ∀a,b:nat_ordered_set. lt nat_ordered_set a b → a < b. -intros; cases H; clear H; cases H2; -[2: apply H;| cases (H1 H)] -qed. \ No newline at end of file