Skip to content
Snippets Groups Projects
Commit 57d111b4 authored by Sergey Bozhko's avatar Sergey Bozhko :eyes:
Browse files

small clean up in util/step_function.v

parent 3706139b
No related branches found
No related tags found
No related merge requests found
...@@ -5,8 +5,8 @@ Section StepFunction. ...@@ -5,8 +5,8 @@ Section StepFunction.
Section Defs. Section Defs.
(* We say that a function f... *) (* We say that a function [f]... *)
Variable f: nat -> nat. Variable f : nat -> nat.
(* ...is a step function iff the following holds. *) (* ...is a step function iff the following holds. *)
Definition is_step_function := Definition is_step_function :=
...@@ -16,25 +16,25 @@ Section StepFunction. ...@@ -16,25 +16,25 @@ Section StepFunction.
Section Lemmas. Section Lemmas.
(* Let f be any step function over natural numbers. *) (* Let [f] be any step function over natural numbers. *)
Variable f: nat -> nat. Variable f : nat -> nat.
Hypothesis H_step_function: is_step_function f. Hypothesis H_step_function : is_step_function f.
(* In this section, we prove a result similar to the intermediate (* In this section, we prove a result similar to the intermediate
value theorem for continuous functions. *) value theorem for continuous functions. *)
Section ExistsIntermediateValue. Section ExistsIntermediateValue.
(* Consider any interval [x1, x2]. *) (* Consider any interval [x1, x2]. *)
Variable x1 x2: nat. Variable x1 x2 : nat.
Hypothesis H_is_interval: x1 <= x2. Hypothesis H_is_interval : x1 <= x2.
(* Let t be any value such that f x1 < y < f x2. *) (* Let [t] be any value such that [f x1 <= y < f x2]. *)
Variable y: nat. Variable y : nat.
Hypothesis H_between: f x1 <= y < f x2. Hypothesis H_between : f x1 <= y < f x2.
(* Then, we prove that there exists an intermediate point x_mid such that (* Then, we prove that there exists an intermediate point [x_mid] such that
f x_mid = y. *) [f x_mid = y]. *)
Lemma exists_intermediate_point: Lemma exists_intermediate_point :
exists x_mid, x1 <= x_mid < x2 /\ f x_mid = y. exists x_mid, x1 <= x_mid < x2 /\ f x_mid = y.
Proof. Proof.
rename H_is_interval into INT, H_step_function into STEP, H_between into BETWEEN. rename H_is_interval into INT, H_step_function into STEP, H_between into BETWEEN.
...@@ -46,11 +46,11 @@ Section StepFunction. ...@@ -46,11 +46,11 @@ Section StepFunction.
{ move => x2 LE /andP [GEy LTy]. { move => x2 LE /andP [GEy LTy].
exploit (DELTA (x2 - x1)); exploit (DELTA (x2 - x1));
first by apply/andP; split; last by rewrite addnBA // addKn. first by apply/andP; split; last by rewrite addnBA // addKn.
by rewrite addnBA // addKn. by rewrite addnBA // addKn.
} }
induction delta. induction delta.
{ rewrite addn0; move => /andP [GE0 LT0]. { rewrite addn0; move => /andP [GE0 LT0].
by apply (leq_ltn_trans GE0) in LT0; rewrite ltnn in LT0. by apply (leq_ltn_trans GE0) in LT0; rewrite ltnn in LT0.
} }
{ move => /andP [GT LT]. { move => /andP [GT LT].
specialize (STEP (x1 + delta)); rewrite leq_eqVlt in STEP. specialize (STEP (x1 + delta)); rewrite leq_eqVlt in STEP.
...@@ -59,18 +59,18 @@ Section StepFunction. ...@@ -59,18 +59,18 @@ Section StepFunction.
first by rewrite !addn1 in EQ; rewrite addnS EQ ltnS in LT. first by rewrite !addn1 in EQ; rewrite addnS EQ ltnS in LT.
rewrite [X in _ < X]addn1 ltnS in STEP. rewrite [X in _ < X]addn1 ltnS in STEP.
apply: (leq_trans _ STEP). apply: (leq_trans _ STEP).
by rewrite addn1 -addnS ltnW. by rewrite addn1 -addnS ltnW.
} clear STEP LT. } clear STEP LT.
rewrite leq_eqVlt in LE. rewrite leq_eqVlt in LE.
move: LE => /orP [/eqP EQy | LT]. move: LE => /orP [/eqP EQy | LT].
{ exists (x1 + delta); split; last by rewrite EQy. { exists (x1 + delta); split; last by rewrite EQy.
by apply/andP; split; [by apply leq_addr | by rewrite addnS]. by apply/andP; split; [apply leq_addr | rewrite addnS].
} }
{ feed (IHdelta); first by apply/andP; split. { feed (IHdelta); first by apply/andP; split.
move: IHdelta => [x_mid [/andP [GE0 LT0] EQ0]]. move: IHdelta => [x_mid [/andP [GE0 LT0] EQ0]].
exists x_mid; split; last by done. exists x_mid; split; last by done.
apply/andP; split; first by done. apply/andP; split; first by done.
by apply: (leq_trans LT0); rewrite addnS. by apply: (leq_trans LT0); rewrite addnS.
} }
} }
Qed. Qed.
...@@ -83,22 +83,22 @@ Section StepFunction. ...@@ -83,22 +83,22 @@ Section StepFunction.
value theorem, but for predicates of natural numbers. *) value theorem, but for predicates of natural numbers. *)
Section ExistsIntermediateValuePredicates. Section ExistsIntermediateValuePredicates.
(* Let P be any predicate on natural numbers. *) (* Let [P] be any predicate on natural numbers. *)
Variable P : nat -> bool. Variable P : nat -> bool.
(* Consider a time interval [t1,t2] such that ... *) (* Consider a time interval [t1,t2] such that ... *)
Variables t1 t2 : nat. Variables t1 t2 : nat.
Hypothesis H_t1_le_t2 : t1 <= t2. Hypothesis H_t1_le_t2 : t1 <= t2.
(* ... P doesn't hold for t1 ... *) (* ... [P] doesn't hold for [t1] ... *)
Hypothesis H_not_P_at_t1 : ~~ P t1. Hypothesis H_not_P_at_t1 : ~~ P t1.
(* ... but holds for t2. *) (* ... but holds for [t2]. *)
Hypothesis H_P_at_t2 : P t2. Hypothesis H_P_at_t2 : P t2.
(* Then we prove that within time interval [t1,t2] there exists time (* Then we prove that within time interval [t1,t2] there exists time
instant t such that t is the first time instant when P holds. *) instant [t] such that [t] is the first time instant when [P] holds. *)
Lemma exists_first_intermediate_point: Lemma exists_first_intermediate_point :
exists t, (t1 < t <= t2) /\ (forall x, t1 <= x < t -> ~~ P x) /\ P t. exists t, (t1 < t <= t2) /\ (forall x, t1 <= x < t -> ~~ P x) /\ P t.
Proof. Proof.
have EX: exists x, P x && (t1 < x <= t2). have EX: exists x, P x && (t1 < x <= t2).
...@@ -106,7 +106,7 @@ Section StepFunction. ...@@ -106,7 +106,7 @@ Section StepFunction.
apply/andP; split; first by done. apply/andP; split; first by done.
apply/andP; split; last by done. apply/andP; split; last by done.
move: H_t1_le_t2; rewrite leq_eqVlt; move => /orP [/eqP EQ | NEQ1]; last by done. move: H_t1_le_t2; rewrite leq_eqVlt; move => /orP [/eqP EQ | NEQ1]; last by done.
by exfalso; subst t2; move: H_not_P_at_t1 => /negP NPt1. by exfalso; subst t2; move: H_not_P_at_t1 => /negP NPt1.
} }
have MIN := ex_minnP EX. have MIN := ex_minnP EX.
move: MIN => [x /andP [Px /andP [LT1 LT2]] MIN]; clear EX. move: MIN => [x /andP [Px /andP [LT1 LT2]] MIN]; clear EX.
...@@ -116,12 +116,12 @@ Section StepFunction. ...@@ -116,12 +116,12 @@ Section StepFunction.
{ apply/andP; split; first by done. { apply/andP; split; first by done.
apply/andP; split. apply/andP; split.
- move: NEQ1. rewrite leq_eqVlt; move => /orP [/eqP EQ | NEQ1]; last by done. - move: NEQ1. rewrite leq_eqVlt; move => /orP [/eqP EQ | NEQ1]; last by done.
by exfalso; subst y; move: H_not_P_at_t1 => /negP NPt1. by exfalso; subst y; move: H_not_P_at_t1 => /negP NPt1.
- by apply ltnW, leq_trans with x. - by apply ltnW, leq_trans with x.
} }
by move: NEQ2; rewrite ltnNge; move => /negP NEQ2. by move: NEQ2; rewrite ltnNge; move => /negP NEQ2.
Qed. Qed.
End ExistsIntermediateValuePredicates. End ExistsIntermediateValuePredicates.
End StepFunction. End StepFunction.
\ No newline at end of file
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment