Update for Coq 8.9

This commit is contained in:
Adam Chlipala 2019-03-04 11:23:01 -05:00
parent e92a697e33
commit e032ab4240
7 changed files with 23 additions and 23 deletions

2
.gitignore vendored
View file

@ -12,7 +12,7 @@
Makefile.coq Makefile.coq
Makefile.coq.conf Makefile.coq.conf
*.glob *.glob
*.v.d *.d
*.vo *.vo
frap.tgz frap.tgz
.coq-native .coq-native

View file

@ -886,7 +886,7 @@ Module DeeperWithFail.
| Stepped (h : heap) (c : cmd result) | Stepped (h : heap) (c : cmd result)
| Failed. | Failed.
Implicit Arguments Failed [result]. Arguments Failed {result}.
Fixpoint step {result} (c : cmd result) (h : heap) : stepResult result := Fixpoint step {result} (c : cmd result) (h : heap) : stepResult result :=
match c with match c with

View file

@ -731,7 +731,7 @@ Module DeeperWithFail.
| Stepped (h : heap) (c : cmd result) | Stepped (h : heap) (c : cmd result)
| Failed. | Failed.
Implicit Arguments Failed [result]. Arguments Failed {result}.
Fixpoint step {result} (c : cmd result) (h : heap) : stepResult result := Fixpoint step {result} (c : cmd result) (h : heap) : stepResult result :=
match c with match c with

View file

@ -1217,7 +1217,7 @@ Section split.
Defined. Defined.
End split. End split.
Implicit Arguments split [P1 P2]. Arguments split {P1 P2}.
(* And now, a few more boring lemmas. Rejoin at "BOREDOM VANQUISHED", if you (* And now, a few more boring lemmas. Rejoin at "BOREDOM VANQUISHED", if you
* like. *) * like. *)

View file

@ -599,7 +599,7 @@ Section split.
Defined. Defined.
End split. End split.
Implicit Arguments split [P1 P2]. Arguments split {P1 P2}.
(* And now, a few more boring lemmas. Rejoin at "BOREDOM VANQUISHED", if you (* And now, a few more boring lemmas. Rejoin at "BOREDOM VANQUISHED", if you
* like. *) * like. *)

View file

@ -1,5 +1,5 @@
Require Import Eqdep String Arith Omega Program Sets Relations Map Var Invariant Bool ModelCheck. Require Import Eqdep String Arith Omega Program Sets Relations Map Var Invariant Bool ModelCheck.
Export String Arith Sets Relations Map Var Invariant Bool ModelCheck. Export Ascii String Arith Sets Relations Map Var Invariant Bool ModelCheck.
Require Import List. Require Import List.
Export List ListNotations. Export List ListNotations.
Open Scope string_scope. Open Scope string_scope.

View file

@ -776,25 +776,25 @@ Inductive boundRunningTime : cmd -> nat -> Prop :=
(* Perhaps surprisingly, there exist commands that have no finite time bounds! (* Perhaps surprisingly, there exist commands that have no finite time bounds!
* Mixed-embedding languages often have these counterintuitive properties. *) * Mixed-embedding languages often have these counterintuitive properties. *)
Fixpoint scribbly (n : nat) : cmd :=
match n with
| O => Return 0
| S n' => _ <- Write n' 0; scribbly n'
end.
Lemma scribbly_time : forall n m,
boundRunningTime (scribbly n) m
-> m >= n.
Proof.
induct n; invert 1; auto.
invert H2.
specialize (H4 n0).
apply IHn in H4.
linear_arithmetic.
Qed.
Theorem boundRunningTime_not_total : exists c, forall n, ~boundRunningTime c n. Theorem boundRunningTime_not_total : exists c, forall n, ~boundRunningTime c n.
Proof. Proof.
Fixpoint scribbly (n : nat) : cmd :=
match n with
| O => Return 0
| S n' => _ <- Write n' 0; scribbly n'
end.
Lemma scribbly_time : forall n m,
boundRunningTime (scribbly n) m
-> m >= n.
Proof.
induct n; invert 1; auto.
invert H2.
specialize (H4 n0).
apply IHn in H4.
linear_arithmetic.
Qed.
exists (n <- Read 0; scribbly n); propositional. exists (n <- Read 0; scribbly n); propositional.
invert H. invert H.
specialize (H4 (S n2)). specialize (H4 (S n2)).