(* *)
(**************************************************************************)
+include "nat/minus.ma".
include "list/list.ma".
interpretation "list nth" 'nth = (nth _).
intros; elim n; [reflexivity] simplify; rewrite > H; reflexivity;
qed.
-record rel : Type ≝ {
- rel_T :> Type;
+record rel (rel_T : Type) : Type ≝ {
rel_op :2> rel_T → rel_T → Prop
}.
record trans_rel : Type ≝ {
- o_rel :> rel;
- o_tra : ∀x,y,z: o_rel.o_rel x y → o_rel y z → o_rel x z
+ o_T :> Type;
+ o_rel :> rel o_T;
+ o_tra : ∀x,y,z: o_T.o_rel x y → o_rel y z → o_rel x z
}.
lemma trans: ∀r:trans_rel.∀x,y,z:r.r x y → r y z → r x z.
apply o_tra;
qed.
-inductive sorted (lt : trans_rel): list (rel_T lt) → Prop ≝
+inductive sorted (lt : trans_rel): list (o_T lt) → Prop ≝
| sorted_nil : sorted lt []
| sorted_one : ∀x. sorted lt [x]
| sorted_cons : ∀x,y,tl. lt x y → sorted lt (y::tl) → sorted lt (x::y::tl).
intros (r x y l H1); inversion H1; intros; [1,2: destruct H]
destruct H4; inversion H2; intros; [destruct H4]
[1: destruct H4; constructor 2;
-|2: destruct H7; constructor 3; [ apply (o_tra ? ??? H H4);]
- apply (sorted_tail ??? H2);]
+|2: destruct H7; constructor 3;
+ [ apply (o_tra ? ??? H H4); | apply (sorted_tail ??? H2);]]
qed.
lemma sorted_tail_bigger : ∀r,x,l,d.sorted r (x::l) → ∀i. i < \len l → r x (\nth l d i).
destruct H6; simplify; assumption;]
qed.
+(*
lemma all_bases_positive : ∀f:q_f.∀i. OQ < nth_base (bars f) (S i).
intro f; generalize in match (bars_begin_OQ f); generalize in match (bars_sorted f);
cases (bars_not_nil f); intros;
cases n in H3; intros; [cases (not_le_Sn_O ? H3)] apply (H2 n1);
apply (le_S_S_to_le ?? H3);]
qed.
+*)
+(* move in nat/ *)
lemma lt_n_plus_n_Sm : ∀n,m:nat.n < n + S m.
-intros; rewrite > sym_plus; apply (le_S_S n (m+n)); apply (le_plus_n m n); qed.
+intros; rewrite > sym_plus; apply (le_S_S n (m+n)); alias id "le_plus_n" = "cic:/matita/nat/le_arith/le_plus_n.con".
+apply (le_plus_n m n); qed.
-lemma nth_concat_lt_len:
- ∀T:Type.∀l1,l2:list T.∀def.∀i.i < len l1 → nth (l1@l2) def i = nth l1 def i.
+lemma nth_append_lt_len:
+ ∀T:Type.∀l1,l2:list T.∀def.∀i.i < \len l1 → \nth (l1@l2) def i = \nth l1 def i.
intros 4; elim l1; [cases (not_le_Sn_O ? H)] cases i in H H1; simplify; intros;
[reflexivity| rewrite < H;[reflexivity] apply le_S_S_to_le; apply H1]
qed.
-lemma nth_concat_ge_len:
+lemma nth_append_ge_len:
∀T:Type.∀l1,l2:list T.∀def.∀i.
- len l1 ≤ i → nth (l1@l2) def i = nth l2 def (i - len l1).
+ \len l1 ≤ i → \nth (l1@l2) def i = \nth l2 def (i - \len l1).
intros 4; elim l1; [ rewrite < minus_n_O; reflexivity]
cases i in H1; simplify; intros; [cases (not_le_Sn_O ? H1)]
apply H; apply le_S_S_to_le; apply H1;
lemma nth_len:
∀T:Type.∀l1,l2:list T.∀def,x.
- nth (l1@x::l2) def (len l1) = x.
+ \nth (l1@x::l2) def (\len l1) = x.
intros 2; elim l1;[reflexivity] simplify; apply H; qed.
-lemma all_bigger_can_concat_bigger:
- ∀l1,l2,start,b,x,n.
- (∀i.i< len l1 → nth_base l1 i < \fst b) →
- (∀i.i< len l2 → \fst b ≤ nth_base l2 i) →
- (∀i.i< len l1 → start ≤ i → x ≤ nth_base l1 i) →
- start ≤ n → n < len (l1@b::l2) → x ≤ \fst b → x ≤ nth_base (l1@b::l2) n.
-intros; cases (cmp_nat n (len l1));
-[1: unfold nth_base; rewrite > (nth_concat_lt_len ????? H6);
- apply (H2 n); assumption;
-|2: rewrite > H6; unfold nth_base; rewrite > nth_len; assumption;
-|3: unfold nth_base; rewrite > nth_concat_ge_len; [2: apply lt_to_le; assumption]
- rewrite > len_concat in H4; simplify in H4; rewrite < plus_n_Sm in H4;
- lapply linear le_S_S_to_le to H4 as K; rewrite > sym_plus in K;
- lapply linear le_plus_to_minus to K as X;
- generalize in match X; generalize in match (n - len l1); intro W; cases W; clear W X;
- [intros; assumption] intros;
- apply (q_le_trans ??? H5); apply (H1 n1); assumption;]
-qed.
-
lemma sorted_head_smaller:
- ∀l,p. sorted (p::l) → ∀i.i < len l → \fst p < nth_base l i.
-intro l; elim l; intros; [cases (not_le_Sn_O ? H1)] cases i in H2; simplify; intros;
+ ∀r,l,p,d. sorted r (p::l) → ∀i.i < \len l → r p (\nth l d i).
+intros 2 (r l); elim l; [cases (not_le_Sn_O ? H1)] cases i in H2; simplify; intros;
[1: inversion H1; [1,2: simplify; intros; destruct H3] intros; destruct H6; assumption;
-|2: apply (H p ? n ?); [apply (sorted_skip ??? H1)] apply le_S_S_to_le; apply H2]
-qed.
+|2: apply (H p ?? n ?); [apply (sorted_skip ???? H1)] apply le_S_S_to_le; apply H2]
+qed.
+
+alias symbol "lt" = "natural 'less than'".
+lemma sorted_pivot:
+ ∀r,l1,l2,p,d. sorted r (l1@p::l2) →
+ (∀i. i < \len l1 → r (\nth l1 d i) p) ∧
+ (∀i. i < \len l2 → r p (\nth l2 d i)).
+intros 2 (r l); elim l;
+[1: split; [intros; cases (not_le_Sn_O ? H1);] intros;
+ apply sorted_head_smaller; assumption;
+|2: simplify in H1; cases (H ?? d (sorted_tail ??? H1));
+ lapply depth = 0 (sorted_head_smaller ??? d H1) as Hs;
+ split; simplify; intros;
+ [1: cases i in H4; simplify; intros;
+ [1: lapply depth = 0 (Hs (\len l1)) as HS;
+ rewrite > nth_len in HS; apply HS;
+ rewrite > len_append; simplify; apply lt_n_plus_n_Sm;
+ |2: apply (H2 n); apply le_S_S_to_le; apply H4]
+ |2: apply H3; assumption]]
+qed.
+
+inductive cases_bool (p:bool) : bool → CProp ≝
+| case_true : p = true → cases_bool p true
+| cases_false : p = false → cases_bool p false.
+
+lemma case_b : ∀A:Type.∀f:A → bool. ∀x.cases_bool (f x) (f x).
+intros; cases (f x);[left;|right] reflexivity;
+qed.
+
+include "cprop_connectives.ma".
+
+definition eject_N ≝
+ λP.λp:∃x:nat.P x.match p with [ex_introT p _ ⇒ p].
+coercion eject_N.
+definition inject_N ≝ λP.λp:nat.λh:P p. ex_introT ? P p h.
+coercion inject_N with 0 1 nocomposites.
+
+inductive find_spec (T:Type) (P:T→bool) (l:list T) (d:T) (res : nat) : nat → CProp ≝
+| found:
+ ∀i. i < \len l → P (\nth l d i) = true → res = i →
+ (∀j. j < i → P (\nth l d j) = false) → find_spec T P l d res i
+| not_found: ∀i. i = \len l → res = i →
+ (∀j.j < \len l → P (\nth l d j) = false) → find_spec T P l d res i.
+
+lemma find_lemma : ∀T.∀P:T→bool.∀l:list T.∀d.∃i.find_spec ? P l d i i.
+intros;
+letin find ≝ (
+ let rec aux (acc: nat) (l : list T) on l : nat ≝
+ match l with
+ [ nil ⇒ acc
+ | cons x tl ⇒
+ match P x with
+ [ false ⇒ aux (S acc) tl
+ | true ⇒ acc]]
+ in aux :
+ ∀acc,l1.∃p:nat.
+ ∀story. story @ l1 = l → acc = \len story →
+ find_spec ? P story d acc acc →
+ find_spec ? P (story @ l1) d p p);
+[4: clearbody find; cases (find 0 l);
+ lapply (H [] (refl_eq ??) (refl_eq ??)) as K;
+ [2: apply (not_found ?? [] d); intros; try reflexivity; cases (not_le_Sn_O ? H1);
+ |1: cases K; clear K;
+ [2: exists[apply (\len l)]
+ apply not_found; try reflexivity; apply H3;
+ |1: exists[apply i] apply found; try reflexivity; assumption;]]
+|1: intros; cases (aux (S n) l2); simplify; clear aux;
+ lapply depth = 0 (H5 (story@[t])) as K; clear H5;
+ change with (find_spec ? P (story @ ([t] @ l2)) d w w);
+ rewrite < associative_append; apply K; clear K;
+ [1: rewrite > associative_append; apply H2;
+ |2: rewrite > H3; rewrite > len_append; rewrite > sym_plus; reflexivity;
+ |3: cases H4; clear H4; destruct H7;
+ [2: rewrite > H5; rewrite > (?:S (\len story) = \len (story @ [t])); [2:
+ rewrite > len_append; rewrite > sym_plus; reflexivity;]
+ apply not_found; try reflexivity; intros; cases (cmp_nat (S j) (\len story));
+ [1: rewrite > (nth_append_lt_len ????? H8); apply H7; apply H8;
+ |2: rewrite > (nth_append_ge_len ????? (le_S_S_to_le ?? H8));
+ rewrite > (?: j - \len story = 0);[assumption]
+ rewrite > (?:j = \len story);[rewrite > minus_n_n; reflexivity]
+ apply le_to_le_to_eq; [2: apply le_S_S_to_le; assumption;]
+ rewrite > len_append in H4;rewrite > sym_plus in H4;
+ apply le_S_S_to_le; apply H4;]
+ |1: rewrite < H3 in H5; cases (not_le_Sn_n ? H5);]]
+|2: intros; cases H4; clear H4;
+ [1: destruct H7; rewrite > H3 in H5; cases (not_le_Sn_n ? H5);
+ |2: apply found; try reflexivity;
+ [1: rewrite > len_append; simplify; rewrite > H5; apply lt_n_plus_n_Sm;
+ |2: rewrite > H5; rewrite > nth_append_ge_len; [2: apply le_n]
+ rewrite < minus_n_n; assumption;
+ |3: intros; rewrite > H5 in H4; rewrite > nth_append_lt_len; [2:assumption]
+ apply H7; assumption]]
+|3: intros; rewrite > append_nil; assumption;]
+qed.
+
+lemma find : ∀T:Type.∀P:T→bool.∀l:list T.∀d:T.nat.
+intros; cases (find_lemma ? f l t); apply w; qed.
+
+lemma cases_find: ∀T,P,l,d. find_spec T P l d (find T P l d) (find T P l d).
+intros; unfold find; cases (find_lemma T P l d); simplify; assumption; qed.
include "nat_ordered_set.ma".
include "models/q_support.ma".
-include "models/list_support.ma".
+include "models/list_support.ma".
include "cprop_connectives.ma".
definition bar ≝ ℚ × ℚ.
notation < "\ldots\rect\square\EmptySmallSquare\ldots" with precedence 90 for @{'lq2}.
interpretation "lq2" 'lq2 = (list bar).
-inductive sorted : list bar → Prop ≝
-| sorted_nil : sorted []
-| sorted_one : ∀x. sorted [x]
-| sorted_cons : ∀x,y,tl. \fst x < \fst y → sorted (y::tl) → sorted (x::y::tl).
+definition q2_lt := mk_rel bar (λx,y:bar.\fst x < \fst y).
-definition nth_base ≝ λf,n. \fst (nth f ▭ n).
-definition nth_height ≝ λf,n. \snd (nth f ▭ n).
+interpretation "bar lt" 'lt x y = (rel_op _ q2_lt x y).
-record q_f : Type ≝ {
- bars: list bar;
- bars_sorted : sorted bars;
- bars_begin_OQ : nth_base bars O = OQ;
- bars_tail_OQ : nth_height bars (pred (len bars)) = OQ
-}.
+lemma q2_trans : ∀a,b,c:bar. a < b → b < c → a < c.
+intros 3; cases a; cases b; cases c; unfold q2_lt; simplify; intros;
+apply (q_lt_trans ??? H H1);
+qed.
-lemma nth_nil: ∀T,i.∀def:T. nth [] def i = def.
-intros; elim i; simplify; [reflexivity;] assumption; qed.
+definition q2_trel := mk_trans_rel bar q2_lt q2_trans.
-lemma len_concat: ∀T:Type.∀l1,l2:list T. len (l1@l2) = len l1 + len l2.
-intros; elim l1; [reflexivity] simplify; rewrite < H; reflexivity;
-qed.
+interpretation "bar lt" 'lt x y = (FunClass_2_OF_trans_rel q2_trel x y).
-inductive non_empty_list (A:Type) : list A → Type :=
-| show_head: ∀x,l. non_empty_list A (x::l).
+definition canonical_q_lt : rel bar → trans_rel ≝ λx:rel bar.q2_trel.
-lemma bars_not_nil: ∀f:q_f.non_empty_list ? (bars f).
-intro f; generalize in match (bars_begin_OQ f); cases (bars f);
-[1: intro X; normalize in X; destruct X;
-|2: intros; constructor 1;]
-qed.
+coercion canonical_q_lt with nocomposites.
-lemma sorted_tail: ∀x,l.sorted (x::l) → sorted l.
-intros; inversion H; intros; [destruct H1;|destruct H1;constructor 1;]
-destruct H4; assumption;
-qed.
+interpretation "bar lt" 'lt x y = (FunClass_2_OF_trans_rel (canonical_q_lt _) x y).
-lemma sorted_skip: ∀x,y,l. sorted (x::y::l) → sorted (x::l).
-intros; inversion H; intros; [1,2: destruct H1]
-destruct H4; inversion H2; intros; [destruct H4]
-[1: destruct H4; constructor 2;
-|2: destruct H7; constructor 3; [apply (q_lt_trans ??? H1 H4);]
- apply (sorted_tail ?? H2);]
-qed.
+definition nth_base ≝ λf,n. \fst (\nth f ▭ n).
+definition nth_height ≝ λf,n. \snd (\nth f ▭ n).
+
+record q_f : Type ≝ {
+ bars: list bar;
+ bars_sorted : sorted q2_lt bars;
+ bars_begin_OQ : nth_base bars O = OQ;
+ bars_end_OQ : nth_height bars (pred (\len bars)) = OQ
+}.
-lemma sorted_tail_bigger : ∀x,l.sorted (x::l) → ∀i. i < len l → \fst x < nth_base l i.
-intros 2; elim l; [ cases (not_le_Sn_O i H1);]
-cases i in H2;
-[2: intros; apply (H ? n);[apply (sorted_skip ??? H1)|apply le_S_S_to_le; apply H2]
-|1: intros; inversion H1; intros; [1,2: destruct H3]
- destruct H6; simplify; assumption;]
+lemma len_bases_gt_O: ∀f.O < \len (bars f).
+intros; generalize in match (bars_begin_OQ f); cases (bars f); intros;
+[2: simplify; apply le_S_S; apply le_O_n;
+|1: normalize in H; destruct H;]
qed.
lemma all_bases_positive : ∀f:q_f.∀i. OQ < nth_base (bars f) (S i).
intro f; generalize in match (bars_begin_OQ f); generalize in match (bars_sorted f);
-cases (bars_not_nil f); intros;
-cases (cmp_nat i (len l));
-[1: lapply (sorted_tail_bigger ?? H ? H2) as K; simplify in H1;
- rewrite > H1 in K; apply K;
-|2: rewrite > H2; simplify; elim l; simplify; [apply (q_pos_OQ one)]
- assumption;
-|3: simplify; elim l in i H2;[simplify; rewrite > nth_nil; apply (q_pos_OQ one)]
- cases n in H3; intros; [cases (not_le_Sn_O ? H3)] apply (H2 n1);
- apply (le_S_S_to_le ?? H3);]
+cases (len_gt_non_empty ?? (len_bases_gt_O f)); intros;
+cases (cmp_nat (\len l) i);
+[2: lapply (sorted_tail_bigger q2_lt ?? ▭ H ? H2) as K;
+ simplify in H1; rewrite < H1; apply K;
+|1: simplify; elim l in i H2;[simplify; rewrite > nth_nil; apply (q_pos_OQ one)]
+ cases n in H3; intros; [simplify in H3; cases (not_le_Sn_O ? H3)]
+ apply (H2 n1); simplify in H3; apply (le_S_S_to_le ?? H3);]
qed.
+(*
lemma lt_n_plus_n_Sm : ∀n,m:nat.n < n + S m.
intros; rewrite > sym_plus; apply (le_S_S n (m+n)); apply (le_plus_n m n); qed.
+*)
-lemma nth_concat_lt_len:
- ∀T:Type.∀l1,l2:list T.∀def.∀i.i < len l1 → nth (l1@l2) def i = nth l1 def i.
-intros 4; elim l1; [cases (not_le_Sn_O ? H)] cases i in H H1; simplify; intros;
-[reflexivity| rewrite < H;[reflexivity] apply le_S_S_to_le; apply H1]
-qed.
-
-lemma nth_concat_ge_len:
- ∀T:Type.∀l1,l2:list T.∀def.∀i.
- len l1 ≤ i → nth (l1@l2) def i = nth l2 def (i - len l1).
-intros 4; elim l1; [ rewrite < minus_n_O; reflexivity]
-cases i in H1; simplify; intros; [cases (not_le_Sn_O ? H1)]
-apply H; apply le_S_S_to_le; apply H1;
-qed.
-
-lemma nth_len:
- ∀T:Type.∀l1,l2:list T.∀def,x.
- nth (l1@x::l2) def (len l1) = x.
-intros 2; elim l1;[reflexivity] simplify; apply H; qed.
-
+(*
lemma all_bigger_can_concat_bigger:
∀l1,l2,start,b,x,n.
(∀i.i< len l1 → nth_base l1 i < \fst b) →
[intros; assumption] intros;
apply (q_le_trans ??? H5); apply (H1 n1); assumption;]
qed.
+*)
+
+
+inductive value_spec (f : q_f) (i : ℚ) : ℚ → nat → CProp ≝
+ value_of : ∀q,j.
+ nth_height (bars f) j = q →
+ nth_base (bars f) j < i →
+ (∀n.j < n → n < \len (bars f) → i ≤ nth_base (bars f) n) → value_spec f i q j.
+
+
+definition value : ∀f:q_f.∀i:ratio.∃p:ℚ.∃j.value_spec f (Qpos i) p j.
+intros; letin P ≝ (λx:bar.match q_cmp (Qpos i) (\fst x) with
+ [ q_leq _ ⇒ true
+ | q_gt _ ⇒ false]);
+exists [apply (nth_height (bars f) (pred (find ? P (bars f) ▭)));]
+exists [apply (pred (find ? P (bars f) ▭))] apply value_of;
+[1: reflexivity
+|2: cases (cases_find bar P (bars f) ▭);
+ [1: cases i1 in H H1 H2 H3; simplify; intros;
+ [1: generalize in match (bars_begin_OQ f);
+ cases (len_gt_non_empty ?? (len_bases_gt_O f)); simplify; intros;
+ rewrite > H4; apply q_pos_OQ;
+ |2: cases (len_gt_non_empty ?? (len_bases_gt_O f)) in H3;
+ intros; lapply (H3 n (le_n ?)) as K; unfold P in K;
+ cases (q_cmp (Qpos i) (\fst (\nth (x::l) ▭ n))) in K;
+ simplify; intros; [destruct H5] assumption]
+ |2: destruct H; cases (len_gt_non_empty ?? (len_bases_gt_O f)) in H2;
+ simplify; intros; lapply (H (\len l) (le_n ?)) as K; clear H;
+ unfold P in K; cases (q_cmp (Qpos i) (\fst (\nth (x::l) ▭ (\len l)))) in K;
+ simplify; intros; [destruct H2] assumption;]
+|3: intro; cases (cases_find bar P (bars f) ▭); intros;
+ [1:
+
+generalize in match (bars_begin_OQ f); generalize in match (bars_sorted f);
+generalize in match (bars_end_OQ f);
+cases (len_gt_non_empty ?? (len_bases_gt_O f)); simplify;
+intros;
+[1:
-lemma sorted_head_smaller:
- ∀l,p. sorted (p::l) → ∀i.i < len l → \fst p < nth_base l i.
-intro l; elim l; intros; [cases (not_le_Sn_O ? H1)] cases i in H2; simplify; intros;
-[1: inversion H1; [1,2: simplify; intros; destruct H3] intros; destruct H6; assumption;
-|2: apply (H p ? n ?); [apply (sorted_skip ??? H1)] apply le_S_S_to_le; apply H2]
-qed.
-
-
-alias symbol "pi1" = "pair pi1".
-alias symbol "lt" (instance 6) = "Q less than".
-alias symbol "lt" (instance 2) = "Q less than".
-alias symbol "and" = "logical and".
-lemma sorted_pivot:
- ∀l1,l2,p. sorted (l1@p::l2) →
- (∀i. i < len l1 → nth_base l1 i < \fst p) ∧
- (∀i. i < len l2 → \fst p < nth_base l2 i).
-intro l; elim l;
-[1: split; [intros; cases (not_le_Sn_O ? H1);] intros;
- apply sorted_head_smaller; assumption;
-|2: cases (H ?? (sorted_tail a (l1@p::l2) H1));
- lapply depth = 0 (sorted_head_smaller (l1@p::l2) a H1) as Hs;
- split; simplify; intros;
- [1: cases i in H4; simplify; intros;
- [1: lapply depth = 0 (Hs (len l1)) as HS;
- unfold nth_base in HS; rewrite > nth_len in HS; apply HS;
- rewrite > len_concat; simplify; apply lt_n_plus_n_Sm;
- |2: apply (H2 n); apply le_S_S_to_le; apply H4]
- |2: apply H3; assumption]]
-qed.
-
-definition eject_NxQ ≝
- λP.λp:∃x:nat × ℚ.P x.match p with [ex_introT p _ ⇒ p].
-coercion eject_NxQ.
-definition inject_NxQ ≝ λP.λp:nat × ℚ.λh:P p. ex_introT ? P p h.
-coercion inject_NxQ with 0 1 nocomposites.
-
-definition value_spec : q_f → ℚ → nat × ℚ → Prop ≝
- λf,i,q. nth_height (bars f) (\fst q) = \snd q ∧
- (nth_base (bars f) (\fst q) < i ∧
- ∀n.\fst q < n → n < len (bars f) → i ≤ nth_base (bars f) n).
-definition value : ∀f:q_f.∀i:ratio.∃p:ℚ.∃j.value_spec f (Qpos i) 〈j,p〉.
-intros;
alias symbol "pi2" = "pair pi2".
alias symbol "pi1" = "pair pi1".
alias symbol "lt" (instance 7) = "Q less than".
qed.
inductive cmp_cases (n,m:nat) : CProp ≝
- | cmp_lt : n < m → cmp_cases n m
- | cmp_eq : n = m → cmp_cases n m
+ | cmp_le : n ≤ m → cmp_cases n m
| cmp_gt : m < n → cmp_cases n m.
lemma cmp_nat: ∀n,m.cmp_cases n m.
intros; generalize in match (nat_compare_to_Prop n m);
cases (nat_compare n m); intros;
-[constructor 1|constructor 2|constructor 3] assumption;
+[constructor 1;apply lt_to_le|constructor 1;rewrite > H|constructor 2]
+try assumption; apply le_n;
qed.