X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fdama%2Fordered_sets.ma;h=3424edbb57ee42a840a565488c821d667892ae5f;hb=59d7f64e2e1a22ded8f9017942ca640fe62d886a;hp=dcae29e18f3d8b7c26de5255525ce349c600a60c;hpb=fdcfbe23f5e11b1856ca6adbc78e5374b493d199;p=helm.git diff --git a/helm/software/matita/dama/ordered_sets.ma b/helm/software/matita/dama/ordered_sets.ma index dcae29e18..3424edbb5 100644 --- a/helm/software/matita/dama/ordered_sets.ma +++ b/helm/software/matita/dama/ordered_sets.ma @@ -14,99 +14,91 @@ set "baseuri" "cic:/matita/ordered_sets/". -include "higher_order_defs/relations.ma". -include "nat/plus.ma". -include "constructive_connectives.ma". - -definition cotransitive ≝ - λC:Type.λle:C→C→Prop.∀x,y,z:C. le x y → le x z ∨ le z y. - -definition antisimmetric ≝ - λC:Type.λle:C→C→Prop.∀x,y:C. le x y → le y x → x=y. - -record is_order_relation (C:Type) (le:C→C→Prop) : Type ≝ - { or_reflexive: reflexive ? le; - or_transitive: transitive ? le; - or_antisimmetric: antisimmetric ? le - }. - -record ordered_set: Type ≝ - { os_carrier:> Type; - os_le: os_carrier → os_carrier → Prop; - os_order_relation_properties:> is_order_relation ? os_le - }. - -interpretation "Ordered Sets le" 'leq a b = - (cic:/matita/ordered_sets/os_le.con _ a b). - +include "excedence.ma". + +record is_porder_relation (C:Type) (le:C→C→Prop) (eq:C→C→Prop) : Type ≝ { + por_reflexive: reflexive ? le; + por_transitive: transitive ? le; + por_antisimmetric: antisymmetric ? le eq +}. + +record pordered_set: Type ≝ { + pos_carr:> excedence; + pos_order_relation_properties:> is_porder_relation ? (le pos_carr) (eq pos_carr) +}. + +lemma pordered_set_of_excedence: excedence → pordered_set. +intros (E); apply (mk_pordered_set E); apply (mk_is_porder_relation); +[apply le_reflexive|apply le_transitive|apply le_antisymmetric] +qed. + +definition total_order : ∀E:excedence. Type ≝ + λE:excedence. ∀a,b:E. a ≰ b → a < b. + +alias id "transitive" = "cic:/matita/higher_order_defs/relations/transitive.con". +alias id "cotransitive" = "cic:/matita/higher_order_defs/relations/cotransitive.con". +alias id "antisymmetric" = "cic:/matita/higher_order_defs/relations/antisymmetric.con". theorem antisimmetric_to_cotransitive_to_transitive: - ∀C.∀le:relation C. antisimmetric ? le → cotransitive ? le → - transitive ? le. - intros; - unfold transitive; - intros; - elim (c ? ? z H1); - [ assumption - | rewrite < (H ? ? H2 t); - assumption - ]. + ∀C:Type.∀le:C→C→Prop. antisymmetric ? le → cotransitive ? le → transitive ? le. +intros (T f Af cT); unfold transitive; intros (x y z fxy fyz); +lapply (cT ? ? fxy z) as H; cases H; [assumption] cases (Af ? ? fyz H1); qed. -definition is_increasing ≝ λO:ordered_set.λa:nat→O.∀n:nat.a n ≤ a (S n). -definition is_decreasing ≝ λO:ordered_set.λa:nat→O.∀n:nat.a (S n) ≤ a n. +definition is_increasing ≝ λO:pordered_set.λa:nat→O.∀n:nat.a n ≤ a (S n). +definition is_decreasing ≝ λO:pordered_set.λa:nat→O.∀n:nat.a (S n) ≤ a n. -definition is_upper_bound ≝ λO:ordered_set.λa:nat→O.λu:O.∀n:nat.a n ≤ u. -definition is_lower_bound ≝ λO:ordered_set.λa:nat→O.λu:O.∀n:nat.u ≤ a n. +definition is_upper_bound ≝ λO:pordered_set.λa:nat→O.λu:O.∀n:nat.a n ≤ u. +definition is_lower_bound ≝ λO:pordered_set.λa:nat→O.λu:O.∀n:nat.u ≤ a n. -record is_sup (O:ordered_set) (a:nat→O) (u:O) : Prop ≝ +record is_sup (O:pordered_set) (a:nat→O) (u:O) : Prop ≝ { sup_upper_bound: is_upper_bound O a u; sup_least_upper_bound: ∀v:O. is_upper_bound O a v → u≤v }. -record is_inf (O:ordered_set) (a:nat→O) (u:O) : Prop ≝ +record is_inf (O:pordered_set) (a:nat→O) (u:O) : Prop ≝ { inf_lower_bound: is_lower_bound O a u; inf_greatest_lower_bound: ∀v:O. is_lower_bound O a v → v≤u }. -record is_bounded_below (O:ordered_set) (a:nat→O) : Type ≝ +record is_bounded_below (O:pordered_set) (a:nat→O) : Type ≝ { ib_lower_bound: O; ib_lower_bound_is_lower_bound: is_lower_bound ? a ib_lower_bound }. -record is_bounded_above (O:ordered_set) (a:nat→O) : Type ≝ +record is_bounded_above (O:pordered_set) (a:nat→O) : Type ≝ { ib_upper_bound: O; ib_upper_bound_is_upper_bound: is_upper_bound ? a ib_upper_bound }. -record is_bounded (O:ordered_set) (a:nat→O) : Type ≝ +record is_bounded (O:pordered_set) (a:nat→O) : Type ≝ { ib_bounded_below:> is_bounded_below ? a; ib_bounded_above:> is_bounded_above ? a }. -record bounded_below_sequence (O:ordered_set) : Type ≝ +record bounded_below_sequence (O:pordered_set) : Type ≝ { bbs_seq:1> nat→O; bbs_is_bounded_below:> is_bounded_below ? bbs_seq }. -record bounded_above_sequence (O:ordered_set) : Type ≝ +record bounded_above_sequence (O:pordered_set) : Type ≝ { bas_seq:1> nat→O; bas_is_bounded_above:> is_bounded_above ? bas_seq }. -record bounded_sequence (O:ordered_set) : Type ≝ +record bounded_sequence (O:pordered_set) : Type ≝ { bs_seq:1> nat → O; bs_is_bounded_below: is_bounded_below ? bs_seq; bs_is_bounded_above: is_bounded_above ? bs_seq }. definition bounded_below_sequence_of_bounded_sequence ≝ - λO:ordered_set.λb:bounded_sequence O. + λO:pordered_set.λb:bounded_sequence O. mk_bounded_below_sequence ? b (bs_is_bounded_below ? b). coercion cic:/matita/ordered_sets/bounded_below_sequence_of_bounded_sequence.con. definition bounded_above_sequence_of_bounded_sequence ≝ - λO:ordered_set.λb:bounded_sequence O. + λO:pordered_set.λb:bounded_sequence O. mk_bounded_above_sequence ? b (bs_is_bounded_above ? b). coercion cic:/matita/ordered_sets/bounded_above_sequence_of_bounded_sequence.con.