X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fdama%2Fordered_sets.ma;h=946d89573242cbea8907e867816645e256671e22;hb=f2b44fbe63263f5ed70d251304f0df670302182f;hp=a7f772f34a5b90ec4a59bd1158d2d17660dfbb1a;hpb=3f70fa72abe1d8453ea3565f4a33a05832dbc2e0;p=helm.git diff --git a/matita/dama/ordered_sets.ma b/matita/dama/ordered_sets.ma index a7f772f34..946d89573 100644 --- a/matita/dama/ordered_sets.ma +++ b/matita/dama/ordered_sets.ma @@ -134,9 +134,9 @@ qed. definition reverse_pordered_set: pordered_set → pordered_set. intros (p); apply (mk_pordered_set (reverse_excedence p)); -generalize in match (reverse_excedence p); intros (E); cases E (T f cRf cTf); -simplify; apply mk_is_porder_relation; unfold; intros; -[apply le_reflexive|apply (le_transitive ???? H H1);|apply (le_antisymmetric ??? H H1)] +generalize in match (reverse_excedence p); intros (E); +apply mk_is_porder_relation; +[apply le_reflexive|apply le_transitive|apply le_antisymmetric] qed. lemma is_lower_bound_reverse_is_upper_bound: @@ -205,15 +205,8 @@ intros (O a l H); apply mk_is_sup; apply is_upper_bound_reverse_is_lower_bound; assumption;] qed. -(* -record cotransitively_ordered_set: Type := - { cos_ordered_set :> ordered_set; - cos_cotransitive: cotransitive ? (os_le cos_ordered_set) - }. -*) - definition total_order_property : ∀E:excedence. Type ≝ - λE:excedence. ∀a,b:E. a ≰ b → a < b. + λE:excedence. ∀a,b:E. a ≰ b → b < a. record tordered_set : Type ≝ { tos_poset:> pordered_set;