X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Flist%2Fsort.ma;h=c67b28227101efee63462b06b59020441638b945;hb=f48faa46342b557486e06c17fd574d1eeb386239;hp=d182ed6d68bd1148b663a3d494cd18baf996d306;hpb=c5608d2704215a7004d23f4801e19df3939f0285;p=helm.git diff --git a/helm/software/matita/library/list/sort.ma b/helm/software/matita/library/list/sort.ma index d182ed6d6..c67b28227 100644 --- a/helm/software/matita/library/list/sort.ma +++ b/helm/software/matita/library/list/sort.ma @@ -12,22 +12,29 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/list/sort/". - include "datatypes/bool.ma". include "datatypes/constructors.ma". -include "list/list.ma". +include "list/in.ma". + +inductive sorted (A:Type) (P:A \to A \to Prop): list A \to Prop \def +| sort_nil : sorted A P [] +| sort_cons : \forall x,l.sorted A P l \to (\forall y.in_list ? y l \to P x y) + \to sorted A P (x::l). + +lemma sorted_cons_to_sorted : \forall A,P,x,l.sorted A P (x::l) \to sorted A P l. +intros;inversion H;intros + [destruct H1 + |destruct H4;assumption] +qed. + +lemma sorted_to_minimum : \forall A,P,x,l.sorted A P (x::l) \to + \forall y.in_list ? y l \to P x y. +intros;inversion H;intros; + [destruct H2 + |destruct H5;apply H4;assumption] +qed. + -let rec mem (A:Type) (eq: A → A → bool) x (l: list A) on l ≝ - match l with - [ nil ⇒ false - | (cons a l') ⇒ - match eq x a with - [ true ⇒ true - | false ⇒ mem A eq x l' - ] - ]. - let rec ordered (A:Type) (le: A → A → bool) (l: list A) on l ≝ match l with [ nil ⇒ true @@ -39,6 +46,45 @@ let rec ordered (A:Type) (le: A → A → bool) (l: list A) on l ≝ ] ]. +lemma sorted_to_eq_sorted_b_true : + \forall A,ord,ordb. + (\forall x,y.ord x y \to ordb x y = true) \to + \forall l.sorted A ord l \to ordered A ordb l = true. +intros 5;elim l + [reflexivity + |simplify;rewrite > H1;generalize in match H2;cases l1;intros + [reflexivity + |simplify;rewrite > H + [reflexivity + |apply (sorted_to_minimum ? ? ? ? H3);apply in_list_head] + |apply sort_nil + |apply (sorted_cons_to_sorted ? ? ? ? H3)]] +qed. + +(* + TODO + per far funzionare questa dimostrazione bisogna appoggiarsi a un + eqb (e utilizzare quindi in_list <-> mem), oppure modificare la definizione + di sorted in modo che non si appoggi più a in_list: + sorted [] + sorted [x] per ogni x + sorted x::y::l se ord x y e sorted y::l + +lemma eq_sorted_b_true_to_sorted : + \forall A,ord,ordb. + (\forall x,y.ordb x y = true \to ord x y) \to + \forall l.ordered A ordb l = true \to sorted A ord l. +intros 5;elim l + [apply sort_nil + |apply sort_cons + [apply H1;simplify in H2;generalize in match H2;cases l1;intros + [reflexivity + |simplify in H3;simplify;apply (andb_true_true_r ? ? H3)] + |intros;apply H;generalize in match H2; + generalize in match (in_list_to_mem_true ? ? ? ? H + [ +*) + let rec insert (A:Type) (le: A → A → bool) x (l: list A) on l ≝ match l with [ nil ⇒ [x]