This commit is contained in:
Michael Zhang 2020-06-03 21:46:06 -05:00
commit bc27534fa2
Signed by: michael
GPG key ID: BDA47A31A3C8EE6B
46 changed files with 22279 additions and 0 deletions

665
Auto.v Normal file
View file

@ -0,0 +1,665 @@
(** * Auto: More Automation *)
Set Warnings "-notation-overridden,-parsing".
From Coq Require Import omega.Omega.
From LF Require Import Maps.
From LF Require Import Imp.
(** Up to now, we've used the more manual part of Coq's tactic
facilities. In this chapter, we'll learn more about some of Coq's
powerful automation features: proof search via the [auto] tactic,
automated forward reasoning via the [Ltac] hypothesis matching
machinery, and deferred instantiation of existential variables
using [eapply] and [eauto]. Using these features together with
Ltac's scripting facilities will enable us to make our proofs
startlingly short! Used properly, they can also make proofs more
maintainable and robust to changes in underlying definitions. A
deeper treatment of [auto] and [eauto] can be found in the
[UseAuto] chapter in _Programming Language Foundations_.
There's another major category of automation we haven't discussed
much yet, namely built-in decision procedures for specific kinds
of problems: [omega] is one example, but there are others. This
topic will be deferred for a while longer.
Our motivating example will be this proof, repeated with just a
few small changes from the [Imp] chapter. We will simplify
this proof in several stages. *)
(** First, define a little Ltac macro to compress a common
pattern into a single command. *)
Ltac inv H := inversion H; subst; clear H.
Theorem ceval_deterministic: forall c st st1 st2,
st =[ c ]=> st1 ->
st =[ c ]=> st2 ->
st1 = st2.
Proof.
intros c st st1 st2 E1 E2;
generalize dependent st2;
induction E1; intros st2 E2; inv E2.
- (* E_Skip *) reflexivity.
- (* E_Ass *) reflexivity.
- (* E_Seq *)
assert (st' = st'0) as EQ1.
{ (* Proof of assertion *) apply IHE1_1; apply H1. }
subst st'0.
apply IHE1_2. assumption.
(* E_IfTrue *)
- (* b evaluates to true *)
apply IHE1. assumption.
- (* b evaluates to false (contradiction) *)
rewrite H in H5. inversion H5.
(* E_IfFalse *)
- (* b evaluates to true (contradiction) *)
rewrite H in H5. inversion H5.
- (* b evaluates to false *)
apply IHE1. assumption.
(* E_WhileFalse *)
- (* b evaluates to false *)
reflexivity.
- (* b evaluates to true (contradiction) *)
rewrite H in H2. inversion H2.
(* E_WhileTrue *)
- (* b evaluates to false (contradiction) *)
rewrite H in H4. inversion H4.
- (* b evaluates to true *)
assert (st' = st'0) as EQ1.
{ (* Proof of assertion *) apply IHE1_1; assumption. }
subst st'0.
apply IHE1_2. assumption. Qed.
(* ################################################################# *)
(** * The [auto] Tactic *)
(** Thus far, our proof scripts mostly apply relevant hypotheses or
lemmas by name, and one at a time. *)
Example auto_example_1 : forall (P Q R: Prop),
(P -> Q) -> (Q -> R) -> P -> R.
Proof.
intros P Q R H1 H2 H3.
apply H2. apply H1. assumption.
Qed.
(** The [auto] tactic frees us from this drudgery by _searching_ for a
sequence of applications that will prove the goal: *)
Example auto_example_1' : forall (P Q R: Prop),
(P -> Q) -> (Q -> R) -> P -> R.
Proof.
auto.
Qed.
(** The [auto] tactic solves goals that are solvable by any combination of
- [intros] and
- [apply] (of hypotheses from the local context, by default). *)
(** Using [auto] is always "safe" in the sense that it will never fail
and will never change the proof state: either it completely solves
the current goal, or it does nothing. *)
(** Here is a more interesting example showing [auto]'s power: *)
Example auto_example_2 : forall P Q R S T U : Prop,
(P -> Q) ->
(P -> R) ->
(T -> R) ->
(S -> T -> U) ->
((P->Q) -> (P->S)) ->
T ->
P ->
U.
Proof. auto. Qed.
(** Proof search could, in principle, take an arbitrarily long time,
so there are limits to how far [auto] will search by default. *)
Example auto_example_3 : forall (P Q R S T U: Prop),
(P -> Q) ->
(Q -> R) ->
(R -> S) ->
(S -> T) ->
(T -> U) ->
P ->
U.
Proof.
(* When it cannot solve the goal, [auto] does nothing *)
auto.
(* Optional argument says how deep to search (default is 5) *)
auto 6.
Qed.
(** When searching for potential proofs of the current goal,
[auto] considers the hypotheses in the current context together
with a _hint database_ of other lemmas and constructors. Some
common lemmas about equality and logical operators are installed
in this hint database by default. *)
Example auto_example_4 : forall P Q R : Prop,
Q ->
(Q -> R) ->
P \/ (Q /\ R).
Proof. auto. Qed.
(** We can extend the hint database just for the purposes of one
application of [auto] by writing "[auto using ...]". *)
Lemma le_antisym : forall n m: nat, (n <= m /\ m <= n) -> n = m.
Proof. intros. omega. Qed.
Example auto_example_6 : forall n m p : nat,
(n <= p -> (n <= m /\ m <= n)) ->
n <= p ->
n = m.
Proof.
intros.
auto using le_antisym.
Qed.
(** Of course, in any given development there will probably be
some specific constructors and lemmas that are used very often in
proofs. We can add these to the global hint database by writing
Hint Resolve T.
at the top level, where [T] is a top-level theorem or a
constructor of an inductively defined proposition (i.e., anything
whose type is an implication). As a shorthand, we can write
Hint Constructors c.
to tell Coq to do a [Hint Resolve] for _all_ of the constructors
from the inductive definition of [c].
It is also sometimes necessary to add
Hint Unfold d.
where [d] is a defined symbol, so that [auto] knows to expand uses
of [d], thus enabling further possibilities for applying lemmas that
it knows about. *)
(** It is also possible to define specialized hint databases that can
be activated only when needed. See the Coq reference manual for
more. *)
Hint Resolve le_antisym.
Example auto_example_6' : forall n m p : nat,
(n<= p -> (n <= m /\ m <= n)) ->
n <= p ->
n = m.
Proof.
intros.
auto. (* picks up hint from database *)
Qed.
Definition is_fortytwo x := (x = 42).
Example auto_example_7: forall x,
(x <= 42 /\ 42 <= x) -> is_fortytwo x.
Proof.
auto. (* does nothing *)
Abort.
Hint Unfold is_fortytwo.
Example auto_example_7' : forall x,
(x <= 42 /\ 42 <= x) -> is_fortytwo x.
Proof. auto. Qed.
(** Let's take a first pass over [ceval_deterministic] to simplify the
proof script. *)
Theorem ceval_deterministic': forall c st st1 st2,
st =[ c ]=> st1 ->
st =[ c ]=> st2 ->
st1 = st2.
Proof.
intros c st st1 st2 E1 E2.
generalize dependent st2;
induction E1; intros st2 E2; inv E2; auto.
- (* E_Seq *)
assert (st' = st'0) as EQ1 by auto.
subst st'0.
auto.
- (* E_IfTrue *)
+ (* b evaluates to false (contradiction) *)
rewrite H in H5. inversion H5.
- (* E_IfFalse *)
+ (* b evaluates to true (contradiction) *)
rewrite H in H5. inversion H5.
- (* E_WhileFalse *)
+ (* b evaluates to true (contradiction) *)
rewrite H in H2. inversion H2.
(* E_WhileTrue *)
- (* b evaluates to false (contradiction) *)
rewrite H in H4. inversion H4.
- (* b evaluates to true *)
assert (st' = st'0) as EQ1 by auto.
subst st'0.
auto.
Qed.
(** When we are using a particular tactic many times in a proof, we
can use a variant of the [Proof] command to make that tactic into
a default within the proof. Saying [Proof with t] (where [t] is
an arbitrary tactic) allows us to use [t1...] as a shorthand for
[t1;t] within the proof. As an illustration, here is an alternate
version of the previous proof, using [Proof with auto]. *)
Theorem ceval_deterministic'_alt: forall c st st1 st2,
st =[ c ]=> st1 ->
st =[ c ]=> st2 ->
st1 = st2.
Proof with auto.
intros c st st1 st2 E1 E2;
generalize dependent st2;
induction E1;
intros st2 E2; inv E2...
- (* E_Seq *)
assert (st' = st'0) as EQ1...
subst st'0...
- (* E_IfTrue *)
+ (* b evaluates to false (contradiction) *)
rewrite H in H5. inversion H5.
- (* E_IfFalse *)
+ (* b evaluates to true (contradiction) *)
rewrite H in H5. inversion H5.
- (* E_WhileFalse *)
+ (* b evaluates to true (contradiction) *)
rewrite H in H2. inversion H2.
(* E_WhileTrue *)
- (* b evaluates to false (contradiction) *)
rewrite H in H4. inversion H4.
- (* b evaluates to true *)
assert (st' = st'0) as EQ1...
subst st'0...
Qed.
(* ################################################################# *)
(** * Searching For Hypotheses *)
(** The proof has become simpler, but there is still an annoying
amount of repetition. Let's start by tackling the contradiction
cases. Each of them occurs in a situation where we have both
H1: beval st b = false
and
H2: beval st b = true
as hypotheses. The contradiction is evident, but demonstrating it
is a little complicated: we have to locate the two hypotheses [H1]
and [H2] and do a [rewrite] following by an [inversion]. We'd
like to automate this process.
(In fact, Coq has a built-in tactic [congruence] that will do the
job in this case. But we'll ignore the existence of this tactic
for now, in order to demonstrate how to build forward search
tactics by hand.)
As a first step, we can abstract out the piece of script in
question by writing a little function in Ltac. *)
Ltac rwinv H1 H2 := rewrite H1 in H2; inv H2.
Theorem ceval_deterministic'': forall c st st1 st2,
st =[ c ]=> st1 ->
st =[ c ]=> st2 ->
st1 = st2.
Proof.
intros c st st1 st2 E1 E2.
generalize dependent st2;
induction E1; intros st2 E2; inv E2; auto.
- (* E_Seq *)
assert (st' = st'0) as EQ1 by auto.
subst st'0.
auto.
- (* E_IfTrue *)
+ (* b evaluates to false (contradiction) *)
rwinv H H5.
- (* E_IfFalse *)
+ (* b evaluates to true (contradiction) *)
rwinv H H5.
- (* E_WhileFalse *)
+ (* b evaluates to true (contradiction) *)
rwinv H H2.
(* E_WhileTrue *)
- (* b evaluates to false (contradiction) *)
rwinv H H4.
- (* b evaluates to true *)
assert (st' = st'0) as EQ1 by auto.
subst st'0.
auto. Qed.
(** That was a bit better, but we really want Coq to discover the
relevant hypotheses for us. We can do this by using the [match
goal] facility of Ltac. *)
Ltac find_rwinv :=
match goal with
H1: ?E = true,
H2: ?E = false
|- _ => rwinv H1 H2
end.
(** This [match goal] looks for two distinct hypotheses that
have the form of equalities, with the same arbitrary expression
[E] on the left and with conflicting boolean values on the right.
If such hypotheses are found, it binds [H1] and [H2] to their
names and applies the [rwinv] tactic to [H1] and [H2].
Adding this tactic to the ones that we invoke in each case of the
induction handles all of the contradictory cases. *)
Theorem ceval_deterministic''': forall c st st1 st2,
st =[ c ]=> st1 ->
st =[ c ]=> st2 ->
st1 = st2.
Proof.
intros c st st1 st2 E1 E2.
generalize dependent st2;
induction E1; intros st2 E2; inv E2; try find_rwinv; auto.
- (* E_Seq *)
assert (st' = st'0) as EQ1 by auto.
subst st'0.
auto.
- (* E_WhileTrue *)
+ (* b evaluates to true *)
assert (st' = st'0) as EQ1 by auto.
subst st'0.
auto. Qed.
(** Let's see about the remaining cases. Each of them involves
applying a conditional hypothesis to extract an equality.
Currently we have phrased these as assertions, so that we have to
predict what the resulting equality will be (although we can then
use [auto] to prove it). An alternative is to pick the relevant
hypotheses to use and then [rewrite] with them, as follows: *)
Theorem ceval_deterministic'''': forall c st st1 st2,
st =[ c ]=> st1 ->
st =[ c ]=> st2 ->
st1 = st2.
Proof.
intros c st st1 st2 E1 E2.
generalize dependent st2;
induction E1; intros st2 E2; inv E2; try find_rwinv; auto.
- (* E_Seq *)
rewrite (IHE1_1 st'0 H1) in *. auto.
- (* E_WhileTrue *)
+ (* b evaluates to true *)
rewrite (IHE1_1 st'0 H3) in *. auto. Qed.
(** Now we can automate the task of finding the relevant hypotheses to
rewrite with. *)
Ltac find_eqn :=
match goal with
H1: forall x, ?P x -> ?L = ?R,
H2: ?P ?X
|- _ => rewrite (H1 X H2) in *
end.
(** The pattern [forall x, ?P x -> ?L = ?R] matches any hypothesis of
the form "for all [x], _some property of [x]_ implies _some
equality_." The property of [x] is bound to the pattern variable
[P], and the left- and right-hand sides of the equality are bound
to [L] and [R]. The name of this hypothesis is bound to [H1].
Then the pattern [?P ?X] matches any hypothesis that provides
evidence that [P] holds for some concrete [X]. If both patterns
succeed, we apply the [rewrite] tactic (instantiating the
quantified [x] with [X] and providing [H2] as the required
evidence for [P X]) in all hypotheses and the goal.
One problem remains: in general, there may be several pairs of
hypotheses that have the right general form, and it seems tricky
to pick out the ones we actually need. A key trick is to realize
that we can _try them all_! Here's how this works:
- each execution of [match goal] will keep trying to find a valid
pair of hypotheses until the tactic on the RHS of the match
succeeds; if there are no such pairs, it fails;
- [rewrite] will fail given a trivial equation of the form [X = X];
- we can wrap the whole thing in a [repeat], which will keep doing
useful rewrites until only trivial ones are left. *)
Theorem ceval_deterministic''''': forall c st st1 st2,
st =[ c ]=> st1 ->
st =[ c ]=> st2 ->
st1 = st2.
Proof.
intros c st st1 st2 E1 E2.
generalize dependent st2;
induction E1; intros st2 E2; inv E2; try find_rwinv;
repeat find_eqn; auto.
Qed.
(** The big payoff in this approach is that our proof script should be
more robust in the face of modest changes to our language. To
test this, let's try adding a [REPEAT] command to the language. *)
Module Repeat.
Inductive com : Type :=
| CSkip
| CAsgn (x : string) (a : aexp)
| CSeq (c1 c2 : com)
| CIf (b : bexp) (c1 c2 : com)
| CWhile (b : bexp) (c : com)
| CRepeat (c : com) (b : bexp).
(** [REPEAT] behaves like [WHILE], except that the loop guard is
checked _after_ each execution of the body, with the loop
repeating as long as the guard stays _false_. Because of this,
the body will always execute at least once. *)
Notation "'SKIP'" :=
CSkip.
Notation "c1 ; c2" :=
(CSeq c1 c2) (at level 80, right associativity).
Notation "X '::=' a" :=
(CAsgn X a) (at level 60).
Notation "'WHILE' b 'DO' c 'END'" :=
(CWhile b c) (at level 80, right associativity).
Notation "'TEST' e1 'THEN' e2 'ELSE' e3 'FI'" :=
(CIf e1 e2 e3) (at level 80, right associativity).
Notation "'REPEAT' e1 'UNTIL' b2 'END'" :=
(CRepeat e1 b2) (at level 80, right associativity).
Inductive ceval : state -> com -> state -> Prop :=
| E_Skip : forall st,
ceval st SKIP st
| E_Ass : forall st a1 n X,
aeval st a1 = n ->
ceval st (X ::= a1) (t_update st X n)
| E_Seq : forall c1 c2 st st' st'',
ceval st c1 st' ->
ceval st' c2 st'' ->
ceval st (c1 ; c2) st''
| E_IfTrue : forall st st' b1 c1 c2,
beval st b1 = true ->
ceval st c1 st' ->
ceval st (TEST b1 THEN c1 ELSE c2 FI) st'
| E_IfFalse : forall st st' b1 c1 c2,
beval st b1 = false ->
ceval st c2 st' ->
ceval st (TEST b1 THEN c1 ELSE c2 FI) st'
| E_WhileFalse : forall b1 st c1,
beval st b1 = false ->
ceval st (WHILE b1 DO c1 END) st
| E_WhileTrue : forall st st' st'' b1 c1,
beval st b1 = true ->
ceval st c1 st' ->
ceval st' (WHILE b1 DO c1 END) st'' ->
ceval st (WHILE b1 DO c1 END) st''
| E_RepeatEnd : forall st st' b1 c1,
ceval st c1 st' ->
beval st' b1 = true ->
ceval st (CRepeat c1 b1) st'
| E_RepeatLoop : forall st st' st'' b1 c1,
ceval st c1 st' ->
beval st' b1 = false ->
ceval st' (CRepeat c1 b1) st'' ->
ceval st (CRepeat c1 b1) st''.
Notation "st '=[' c ']=>' st'" := (ceval st c st')
(at level 40).
(** Our first attempt at the determinacy proof does not quite succeed:
the [E_RepeatEnd] and [E_RepeatLoop] cases are not handled by our
previous automation. *)
Theorem ceval_deterministic: forall c st st1 st2,
st =[ c ]=> st1 ->
st =[ c ]=> st2 ->
st1 = st2.
Proof.
intros c st st1 st2 E1 E2.
generalize dependent st2;
induction E1;
intros st2 E2; inv E2; try find_rwinv; repeat find_eqn; auto.
- (* E_RepeatEnd *)
+ (* b evaluates to false (contradiction) *)
find_rwinv.
(* oops: why didn't [find_rwinv] solve this for us already?
answer: we did things in the wrong order. *)
- (* E_RepeatLoop *)
+ (* b evaluates to true (contradiction) *)
find_rwinv.
Qed.
(** Fortunately, to fix this, we just have to swap the invocations of
[find_eqn] and [find_rwinv]. *)
Theorem ceval_deterministic': forall c st st1 st2,
st =[ c ]=> st1 ->
st =[ c ]=> st2 ->
st1 = st2.
Proof.
intros c st st1 st2 E1 E2.
generalize dependent st2;
induction E1;
intros st2 E2; inv E2; repeat find_eqn; try find_rwinv; auto.
Qed.
End Repeat.
(** These examples just give a flavor of what "hyper-automation"
can achieve in Coq. The details of [match goal] are a bit
tricky (and debugging scripts using it is, frankly, not very
pleasant). But it is well worth adding at least simple uses to
your proofs, both to avoid tedium and to "future proof" them. *)
(* ================================================================= *)
(** ** The [eapply] and [eauto] variants *)
(** To close the chapter, we'll introduce one more convenient feature
of Coq: its ability to delay instantiation of quantifiers. To
motivate this feature, recall this example from the [Imp]
chapter: *)
Example ceval_example1:
empty_st =[
X ::= 2;;
TEST X <= 1
THEN Y ::= 3
ELSE Z ::= 4
FI
]=> (Z !-> 4 ; X !-> 2).
Proof.
(* We supply the intermediate state [st']... *)
apply E_Seq with (X !-> 2).
- apply E_Ass. reflexivity.
- apply E_IfFalse. reflexivity. apply E_Ass. reflexivity.
Qed.
(** In the first step of the proof, we had to explicitly provide a
longish expression to help Coq instantiate a "hidden" argument to
the [E_Seq] constructor. This was needed because the definition
of [E_Seq]...
E_Seq : forall c1 c2 st st' st'',
st =[ c1 ]=> st' ->
st' =[ c2 ]=> st'' ->
st =[ c1 ;; c2 ]=> st''
is quantified over a variable, [st'], that does not appear in its
conclusion, so unifying its conclusion with the goal state doesn't
help Coq find a suitable value for this variable. If we leave
out the [with], this step fails ("Error: Unable to find an
instance for the variable [st']").
What's silly about this error is that the appropriate value for [st']
will actually become obvious in the very next step, where we apply
[E_Ass]. If Coq could just wait until we get to this step, there
would be no need to give the value explicitly. This is exactly what
the [eapply] tactic gives us: *)
Example ceval'_example1:
empty_st =[
X ::= 2;;
TEST X <= 1
THEN Y ::= 3
ELSE Z ::= 4
FI
]=> (Z !-> 4 ; X !-> 2).
Proof.
eapply E_Seq. (* 1 *)
- apply E_Ass. (* 2 *)
reflexivity. (* 3 *)
- (* 4 *) apply E_IfFalse. reflexivity. apply E_Ass. reflexivity.
Qed.
(** The [eapply H] tactic behaves just like [apply H] except
that, after it finishes unifying the goal state with the
conclusion of [H], it does not bother to check whether all the
variables that were introduced in the process have been given
concrete values during unification.
If you step through the proof above, you'll see that the goal
state at position [1] mentions the _existential variable_ [?st']
in both of the generated subgoals. The next step (which gets us
to position [2]) replaces [?st'] with a concrete value. This new
value contains a new existential variable [?n], which is
instantiated in its turn by the following [reflexivity] step,
position [3]. When we start working on the second
subgoal (position [4]), we observe that the occurrence of [?st']
in this subgoal has been replaced by the value that it was given
during the first subgoal. *)
(** Several of the tactics that we've seen so far, including [exists],
[constructor], and [auto], have similar variants. For example,
here's a proof using [eauto]: *)
Hint Constructors ceval.
Hint Transparent state.
Hint Transparent total_map.
Definition st12 := (Y !-> 2 ; X !-> 1).
Definition st21 := (Y !-> 1 ; X !-> 2).
Example eauto_example : exists s',
st21 =[
TEST X <= Y
THEN Z ::= Y - X
ELSE Y ::= X + Z
FI
]=> s'.
Proof. eauto. Qed.
(** The [eauto] tactic works just like [auto], except that it uses
[eapply] instead of [apply].
Pro tip: One might think that, since [eapply] and [eauto] are more
powerful than [apply] and [auto], it would be a good idea to use
them all the time. Unfortunately, they are also significantly
slower -- especially [eauto]. Coq experts tend to use [apply] and
[auto] most of the time, only switching to the [e] variants when
the ordinary variants don't do the job. *)
(* Wed Jan 9 12:02:47 EST 2019 *)

47
AutoTest.v Normal file
View file

@ -0,0 +1,47 @@
Set Warnings "-notation-overridden,-parsing".
From Coq Require Export String.
From LF Require Import Auto.
Parameter MISSING: Type.
Module Check.
Ltac check_type A B :=
match type of A with
| context[MISSING] => idtac "Missing:" A
| ?T => first [unify T B; idtac "Type: ok" | idtac "Type: wrong - should be (" B ")"]
end.
Ltac print_manual_grade A :=
match eval compute in A with
| Some (_ ?S ?C) =>
idtac "Score:" S;
match eval compute in C with
| ""%string => idtac "Comment: None"
| _ => idtac "Comment:" C
end
| None =>
idtac "Score: Ungraded";
idtac "Comment: None"
end.
End Check.
From LF Require Import Auto.
Import Check.
Goal True.
idtac " ".
idtac "Max points - standard: 0".
idtac "Max points - advanced: 0".
idtac "".
idtac "********** Summary **********".
idtac "".
idtac "********** Standard **********".
idtac "".
idtac "********** Advanced **********".
Abort.
(* Wed Jan 9 12:02:31 EST 2019 *)

1422
Basics.v Normal file

File diff suppressed because it is too large Load diff

195
BasicsTest.v Normal file
View file

@ -0,0 +1,195 @@
Set Warnings "-notation-overridden,-parsing".
From Coq Require Export String.
From LF Require Import Basics.
Parameter MISSING: Type.
Module Check.
Ltac check_type A B :=
match type of A with
| context[MISSING] => idtac "Missing:" A
| ?T => first [unify T B; idtac "Type: ok" | idtac "Type: wrong - should be (" B ")"]
end.
Ltac print_manual_grade A :=
match eval compute in A with
| Some (_ ?S ?C) =>
idtac "Score:" S;
match eval compute in C with
| ""%string => idtac "Comment: None"
| _ => idtac "Comment:" C
end
| None =>
idtac "Score: Ungraded";
idtac "Comment: None"
end.
End Check.
From LF Require Import Basics.
Import Check.
Goal True.
idtac "------------------- nandb --------------------".
idtac " ".
idtac "#> test_nandb4".
idtac "Possible points: 1".
check_type @test_nandb4 ((nandb true true = false)).
idtac "Assumptions:".
Abort.
Print Assumptions test_nandb4.
Goal True.
idtac " ".
idtac "------------------- andb3 --------------------".
idtac " ".
idtac "#> test_andb34".
idtac "Possible points: 1".
check_type @test_andb34 ((andb3 true true false = false)).
idtac "Assumptions:".
Abort.
Print Assumptions test_andb34.
Goal True.
idtac " ".
idtac "------------------- factorial --------------------".
idtac " ".
idtac "#> test_factorial2".
idtac "Possible points: 1".
check_type @test_factorial2 ((factorial 5 = 10 * 12)).
idtac "Assumptions:".
Abort.
Print Assumptions test_factorial2.
Goal True.
idtac " ".
idtac "------------------- ltb --------------------".
idtac " ".
idtac "#> test_ltb3".
idtac "Possible points: 1".
check_type @test_ltb3 (((4 <? 2) = false)).
idtac "Assumptions:".
Abort.
Print Assumptions test_ltb3.
Goal True.
idtac " ".
idtac "------------------- plus_id_exercise --------------------".
idtac " ".
idtac "#> plus_id_exercise".
idtac "Possible points: 1".
check_type @plus_id_exercise ((forall n m o : nat, n = m -> m = o -> n + m = m + o)).
idtac "Assumptions:".
Abort.
Print Assumptions plus_id_exercise.
Goal True.
idtac " ".
idtac "------------------- mult_S_1 --------------------".
idtac " ".
idtac "#> mult_S_1".
idtac "Possible points: 2".
check_type @mult_S_1 ((forall n m : nat, m = S n -> m * (1 + n) = m * m)).
idtac "Assumptions:".
Abort.
Print Assumptions mult_S_1.
Goal True.
idtac " ".
idtac "------------------- andb_true_elim2 --------------------".
idtac " ".
idtac "#> andb_true_elim2".
idtac "Possible points: 2".
check_type @andb_true_elim2 ((forall b c : bool, b && c = true -> c = true)).
idtac "Assumptions:".
Abort.
Print Assumptions andb_true_elim2.
Goal True.
idtac " ".
idtac "------------------- zero_nbeq_plus_1 --------------------".
idtac " ".
idtac "#> zero_nbeq_plus_1".
idtac "Possible points: 1".
check_type @zero_nbeq_plus_1 ((forall n : nat, (0 =? n + 1) = false)).
idtac "Assumptions:".
Abort.
Print Assumptions zero_nbeq_plus_1.
Goal True.
idtac " ".
idtac "------------------- indentity_fn_applied_twice --------------------".
idtac " ".
idtac "#> identity_fn_applied_twice".
idtac "Possible points: 1".
check_type @identity_fn_applied_twice (
(forall f : bool -> bool,
(forall x : bool, f x = x) -> forall b : bool, f (f b) = b)).
idtac "Assumptions:".
Abort.
Print Assumptions identity_fn_applied_twice.
Goal True.
idtac " ".
idtac "------------------- negation_fn_applied_twice --------------------".
idtac " ".
idtac "#> Manually graded: negation_fn_applied_twice".
idtac "Possible points: 1".
print_manual_grade manual_grade_for_negation_fn_applied_twice.
idtac " ".
idtac "------------------- binary --------------------".
idtac " ".
idtac "#> Manually graded: binary".
idtac "Possible points: 3".
print_manual_grade manual_grade_for_binary.
idtac " ".
idtac " ".
idtac "Max points - standard: 15".
idtac "Max points - advanced: 15".
idtac "".
idtac "********** Summary **********".
idtac "".
idtac "********** Standard **********".
idtac "---------- test_nandb4 ---------".
Print Assumptions test_nandb4.
idtac "---------- test_andb34 ---------".
Print Assumptions test_andb34.
idtac "---------- test_factorial2 ---------".
Print Assumptions test_factorial2.
idtac "---------- test_ltb3 ---------".
Print Assumptions test_ltb3.
idtac "---------- plus_id_exercise ---------".
Print Assumptions plus_id_exercise.
idtac "---------- mult_S_1 ---------".
Print Assumptions mult_S_1.
idtac "---------- andb_true_elim2 ---------".
Print Assumptions andb_true_elim2.
idtac "---------- zero_nbeq_plus_1 ---------".
Print Assumptions zero_nbeq_plus_1.
idtac "---------- identity_fn_applied_twice ---------".
Print Assumptions identity_fn_applied_twice.
idtac "---------- negation_fn_applied_twice ---------".
idtac "MANUAL".
idtac "---------- binary ---------".
idtac "MANUAL".
idtac "".
idtac "********** Advanced **********".
Abort.
(* Wed Jan 9 12:02:05 EST 2019 *)

35
Bib.v Normal file
View file

@ -0,0 +1,35 @@
(** * Bib: Bibliography *)
(* ################################################################# *)
(** * Resources cited in this volume *)
(**
[Bertot 2004] Interactive Theorem Proving and Program Development:
Coq'Art: The Calculus of Inductive Constructions, by Yves Bertot and
Pierre Casteran. Springer-Verlag, 2004.
http://tinyurl.com/z3o7nqu
[Chlipala 2013] Certified Programming with Dependent Types, by
Adam Chlipala. MIT Press. 2013. http://tinyurl.com/zqdnyg2
[Lipovaca 2011] Learn You a Haskell for Great Good! A Beginner's
Guide, by Miran Lipovaca, No Starch Press, April 2011.
http://learnyouahaskell.com
[O'Sullivan 2008] Bryan O'Sullivan, John Goerzen, and Don Stewart:
Real world Haskell - code you can believe in. O'Reilly
2008. http://book.realworldhaskell.org
[Pugh 1991] Pugh, William. "The Omega test: a fast and practical
integer programming algorithm for dependence analysis." Proceedings
of the 1991 ACM/IEEE conference on Supercomputing. ACM, 1991.
http://dl.acm.org/citation.cfm?id=125848
[Wadler 2015] Philip Wadler. "Propositions as types."
Communications of the ACM 58, no. 12 (2015): 75-84.
http://dl.acm.org/citation.cfm?id=2699407
*)
(* Wed Jan 9 12:02:47 EST 2019 *)

47
BibTest.v Normal file
View file

@ -0,0 +1,47 @@
Set Warnings "-notation-overridden,-parsing".
From Coq Require Export String.
From LF Require Import Bib.
Parameter MISSING: Type.
Module Check.
Ltac check_type A B :=
match type of A with
| context[MISSING] => idtac "Missing:" A
| ?T => first [unify T B; idtac "Type: ok" | idtac "Type: wrong - should be (" B ")"]
end.
Ltac print_manual_grade A :=
match eval compute in A with
| Some (_ ?S ?C) =>
idtac "Score:" S;
match eval compute in C with
| ""%string => idtac "Comment: None"
| _ => idtac "Comment:" C
end
| None =>
idtac "Score: Ungraded";
idtac "Comment: None"
end.
End Check.
From LF Require Import Bib.
Import Check.
Goal True.
idtac " ".
idtac "Max points - standard: 0".
idtac "Max points - advanced: 0".
idtac "".
idtac "********** Summary **********".
idtac "".
idtac "********** Standard **********".
idtac "".
idtac "********** Advanced **********".
Abort.
(* Wed Jan 9 12:02:32 EST 2019 *)

132
Extraction.v Normal file
View file

@ -0,0 +1,132 @@
(** * Extraction: Extracting ML from Coq *)
(* ################################################################# *)
(** * Basic Extraction *)
(** In its simplest form, extracting an efficient program from one
written in Coq is completely straightforward.
First we say what language we want to extract into. Options are
OCaml (the most mature), Haskell (mostly works), and Scheme (a bit
out of date). *)
Require Coq.extraction.Extraction.
Extraction Language OCaml.
(** Now we load up the Coq environment with some definitions, either
directly or by importing them from other modules. *)
From Coq Require Import Arith.Arith.
From Coq Require Import Init.Nat.
From Coq Require Import Arith.EqNat.
From LF Require Import ImpCEvalFun.
(** Finally, we tell Coq the name of a definition to extract and the
name of a file to put the extracted code into. *)
Extraction "imp1.ml" ceval_step.
(** When Coq processes this command, it generates a file [imp1.ml]
containing an extracted version of [ceval_step], together with
everything that it recursively depends on. Compile the present
[.v] file and have a look at [imp1.ml] now. *)
(* ################################################################# *)
(** * Controlling Extraction of Specific Types *)
(** We can tell Coq to extract certain [Inductive] definitions to
specific OCaml types. For each one, we must say
- how the Coq type itself should be represented in OCaml, and
- how each constructor should be translated. *)
Extract Inductive bool => "bool" [ "true" "false" ].
(** Also, for non-enumeration types (where the constructors take
arguments), we give an OCaml expression that can be used as a
"recursor" over elements of the type. (Think Church numerals.) *)
Extract Inductive nat => "int"
[ "0" "(fun x -> x + 1)" ]
"(fun zero succ n ->
if n=0 then zero () else succ (n-1))".
(** We can also extract defined constants to specific OCaml terms or
operators. *)
Extract Constant plus => "( + )".
Extract Constant mult => "( * )".
Extract Constant eqb => "( = )".
(** Important: It is entirely _your responsibility_ to make sure that
the translations you're proving make sense. For example, it might
be tempting to include this one
Extract Constant minus => "( - )".
but doing so could lead to serious confusion! (Why?)
*)
Extraction "imp2.ml" ceval_step.
(** Have a look at the file [imp2.ml]. Notice how the fundamental
definitions have changed from [imp1.ml]. *)
(* ################################################################# *)
(** * A Complete Example *)
(** To use our extracted evaluator to run Imp programs, all we need to
add is a tiny driver program that calls the evaluator and prints
out the result.
For simplicity, we'll print results by dumping out the first four
memory locations in the final state.
Also, to make it easier to type in examples, let's extract a
parser from the [ImpParser] Coq module. To do this, we first need
to set up the right correspondence between Coq strings and lists
of OCaml characters. *)
Require Import ExtrOcamlBasic.
Require Import ExtrOcamlString.
(** We also need one more variant of booleans. *)
Extract Inductive sumbool => "bool" ["true" "false"].
(** The extraction is the same as always. *)
From LF Require Import Imp.
From LF Require Import ImpParser.
From LF Require Import Maps.
Extraction "imp.ml" empty_st ceval_step parse.
(** Now let's run our generated Imp evaluator. First, have a look at
[impdriver.ml]. (This was written by hand, not extracted.)
Next, compile the driver together with the extracted code and
execute it, as follows.
ocamlc -w -20 -w -26 -o impdriver imp.mli imp.ml impdriver.ml
./impdriver
(The [-w] flags to [ocamlc] are just there to suppress a few
spurious warnings.) *)
(* ################################################################# *)
(** * Discussion *)
(** Since we've proved that the [ceval_step] function behaves the same
as the [ceval] relation in an appropriate sense, the extracted
program can be viewed as a _certified_ Imp interpreter. Of
course, the parser we're using is not certified, since we didn't
prove anything about it! *)
(* ################################################################# *)
(** * Going Further *)
(** Further details about extraction can be found in the Extract
chapter in _Verified Functional Algorithms_ (_Software
Foundations_ volume 3). *)
(* Wed Jan 9 12:02:46 EST 2019 *)

47
ExtractionTest.v Normal file
View file

@ -0,0 +1,47 @@
Set Warnings "-notation-overridden,-parsing".
From Coq Require Export String.
From LF Require Import Extraction.
Parameter MISSING: Type.
Module Check.
Ltac check_type A B :=
match type of A with
| context[MISSING] => idtac "Missing:" A
| ?T => first [unify T B; idtac "Type: ok" | idtac "Type: wrong - should be (" B ")"]
end.
Ltac print_manual_grade A :=
match eval compute in A with
| Some (_ ?S ?C) =>
idtac "Score:" S;
match eval compute in C with
| ""%string => idtac "Comment: None"
| _ => idtac "Comment:" C
end
| None =>
idtac "Score: Ungraded";
idtac "Comment: None"
end.
End Check.
From LF Require Import Extraction.
Import Check.
Goal True.
idtac " ".
idtac "Max points - standard: 0".
idtac "Max points - advanced: 0".
idtac "".
idtac "********** Summary **********".
idtac "".
idtac "********** Standard **********".
idtac "".
idtac "********** Advanced **********".
Abort.
(* Wed Jan 9 12:02:29 EST 2019 *)

1956
Imp.v Normal file

File diff suppressed because it is too large Load diff

399
ImpCEvalFun.v Normal file
View file

@ -0,0 +1,399 @@
(** * ImpCEvalFun: An Evaluation Function for Imp *)
(** We saw in the [Imp] chapter how a naive approach to defining a
function representing evaluation for Imp runs into difficulties.
There, we adopted the solution of changing from a functional to a
relational definition of evaluation. In this optional chapter, we
consider strategies for getting the functional approach to
work. *)
(* ################################################################# *)
(** * A Broken Evaluator *)
From Coq Require Import omega.Omega.
From Coq Require Import Arith.Arith.
From LF Require Import Imp Maps.
(** Here was our first try at an evaluation function for commands,
omitting [WHILE]. *)
Open Scope imp_scope.
Fixpoint ceval_step1 (st : state) (c : com) : state :=
match c with
| SKIP =>
st
| l ::= a1 =>
(l !-> aeval st a1 ; st)
| c1 ;; c2 =>
let st' := ceval_step1 st c1 in
ceval_step1 st' c2
| TEST b THEN c1 ELSE c2 FI =>
if (beval st b)
then ceval_step1 st c1
else ceval_step1 st c2
| WHILE b1 DO c1 END =>
st (* bogus *)
end.
Close Scope imp_scope.
(** As we remarked in chapter [Imp], in a traditional functional
programming language like ML or Haskell we could write the WHILE
case as follows:
| WHILE b1 DO c1 END =>
if (beval st b1) then
ceval_step1 st (c1;; WHILE b1 DO c1 END)
else st
Coq doesn't accept such a definition ([Error: Cannot guess
decreasing argument of fix]) because the function we want to
define is not guaranteed to terminate. Indeed, the changed
[ceval_step1] function applied to the [loop] program from [Imp.v]
would never terminate. Since Coq is not just a functional
programming language, but also a consistent logic, any potentially
non-terminating function needs to be rejected. Here is an
invalid(!) Coq program showing what would go wrong if Coq allowed
non-terminating recursive functions:
Fixpoint loop_false (n : nat) : False := loop_false n.
That is, propositions like [False] would become
provable (e.g., [loop_false 0] would be a proof of [False]), which
would be a disaster for Coq's logical consistency.
Thus, because it doesn't terminate on all inputs, the full version
of [ceval_step1] cannot be written in Coq -- at least not without
one additional trick... *)
(* ################################################################# *)
(** * A Step-Indexed Evaluator *)
(** The trick we need is to pass an _additional_ parameter to the
evaluation function that tells it how long to run. Informally, we
start the evaluator with a certain amount of "gas" in its tank,
and we allow it to run until either it terminates in the usual way
_or_ it runs out of gas, at which point we simply stop evaluating
and say that the final result is the empty memory. (We could also
say that the result is the current state at the point where the
evaluator runs out of gas -- it doesn't really matter because the
result is going to be wrong in either case!) *)
Open Scope imp_scope.
Fixpoint ceval_step2 (st : state) (c : com) (i : nat) : state :=
match i with
| O => empty_st
| S i' =>
match c with
| SKIP =>
st
| l ::= a1 =>
(l !-> aeval st a1 ; st)
| c1 ;; c2 =>
let st' := ceval_step2 st c1 i' in
ceval_step2 st' c2 i'
| TEST b THEN c1 ELSE c2 FI =>
if (beval st b)
then ceval_step2 st c1 i'
else ceval_step2 st c2 i'
| WHILE b1 DO c1 END =>
if (beval st b1)
then let st' := ceval_step2 st c1 i' in
ceval_step2 st' c i'
else st
end
end.
Close Scope imp_scope.
(** _Note_: It is tempting to think that the index [i] here is
counting the "number of steps of evaluation." But if you look
closely you'll see that this is not the case: for example, in the
rule for sequencing, the same [i] is passed to both recursive
calls. Understanding the exact way that [i] is treated will be
important in the proof of [ceval__ceval_step], which is given as
an exercise below.
One thing that is not so nice about this evaluator is that we
can't tell, from its result, whether it stopped because the
program terminated normally or because it ran out of gas. Our
next version returns an [option state] instead of just a [state],
so that we can distinguish between normal and abnormal
termination. *)
Open Scope imp_scope.
Fixpoint ceval_step3 (st : state) (c : com) (i : nat)
: option state :=
match i with
| O => None
| S i' =>
match c with
| SKIP =>
Some st
| l ::= a1 =>
Some (l !-> aeval st a1 ; st)
| c1 ;; c2 =>
match (ceval_step3 st c1 i') with
| Some st' => ceval_step3 st' c2 i'
| None => None
end
| TEST b THEN c1 ELSE c2 FI =>
if (beval st b)
then ceval_step3 st c1 i'
else ceval_step3 st c2 i'
| WHILE b1 DO c1 END =>
if (beval st b1)
then match (ceval_step3 st c1 i') with
| Some st' => ceval_step3 st' c i'
| None => None
end
else Some st
end
end.
Close Scope imp_scope.
(** We can improve the readability of this version by introducing a
bit of auxiliary notation to hide the plumbing involved in
repeatedly matching against optional states. *)
Notation "'LETOPT' x <== e1 'IN' e2"
:= (match e1 with
| Some x => e2
| None => None
end)
(right associativity, at level 60).
Open Scope imp_scope.
Fixpoint ceval_step (st : state) (c : com) (i : nat)
: option state :=
match i with
| O => None
| S i' =>
match c with
| SKIP =>
Some st
| l ::= a1 =>
Some (l !-> aeval st a1 ; st)
| c1 ;; c2 =>
LETOPT st' <== ceval_step st c1 i' IN
ceval_step st' c2 i'
| TEST b THEN c1 ELSE c2 FI =>
if (beval st b)
then ceval_step st c1 i'
else ceval_step st c2 i'
| WHILE b1 DO c1 END =>
if (beval st b1)
then LETOPT st' <== ceval_step st c1 i' IN
ceval_step st' c i'
else Some st
end
end.
Close Scope imp_scope.
Definition test_ceval (st:state) (c:com) :=
match ceval_step st c 500 with
| None => None
| Some st => Some (st X, st Y, st Z)
end.
(* Compute
(test_ceval empty_st
(X ::= 2;;
TEST (X <= 1)
THEN Y ::= 3
ELSE Z ::= 4
FI)).
====>
Some (2, 0, 4) *)
(** **** Exercise: 2 stars, standard, recommended (pup_to_n)
Write an Imp program that sums the numbers from [1] to
[X] (inclusive: [1 + 2 + ... + X]) in the variable [Y]. Make sure
your solution satisfies the test that follows. *)
Definition pup_to_n : com
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
(*
Example pup_to_n_1 :
test_ceval (X !-> 5) pup_to_n
= Some (0, 15, 0).
Proof. reflexivity. Qed.
[] *)
(** **** Exercise: 2 stars, standard, optional (peven)
Write an [Imp] program that sets [Z] to [0] if [X] is even and
sets [Z] to [1] otherwise. Use [test_ceval] to test your
program. *)
(* FILL IN HERE
[] *)
(* ################################################################# *)
(** * Relational vs. Step-Indexed Evaluation *)
(** As for arithmetic and boolean expressions, we'd hope that
the two alternative definitions of evaluation would actually
amount to the same thing in the end. This section shows that this
is the case. *)
Theorem ceval_step__ceval: forall c st st',
(exists i, ceval_step st c i = Some st') ->
st =[ c ]=> st'.
Proof.
intros c st st' H.
inversion H as [i E].
clear H.
generalize dependent st'.
generalize dependent st.
generalize dependent c.
induction i as [| i' ].
- (* i = 0 -- contradictory *)
intros c st st' H. discriminate H.
- (* i = S i' *)
intros c st st' H.
destruct c;
simpl in H; inversion H; subst; clear H.
+ (* SKIP *) apply E_Skip.
+ (* ::= *) apply E_Ass. reflexivity.
+ (* ;; *)
destruct (ceval_step st c1 i') eqn:Heqr1.
* (* Evaluation of r1 terminates normally *)
apply E_Seq with s.
apply IHi'. rewrite Heqr1. reflexivity.
apply IHi'. simpl in H1. assumption.
* (* Otherwise -- contradiction *)
discriminate H1.
+ (* TEST *)
destruct (beval st b) eqn:Heqr.
* (* r = true *)
apply E_IfTrue. rewrite Heqr. reflexivity.
apply IHi'. assumption.
* (* r = false *)
apply E_IfFalse. rewrite Heqr. reflexivity.
apply IHi'. assumption.
+ (* WHILE *) destruct (beval st b) eqn :Heqr.
* (* r = true *)
destruct (ceval_step st c i') eqn:Heqr1.
{ (* r1 = Some s *)
apply E_WhileTrue with s. rewrite Heqr.
reflexivity.
apply IHi'. rewrite Heqr1. reflexivity.
apply IHi'. simpl in H1. assumption. }
{ (* r1 = None *) discriminate H1. }
* (* r = false *)
injection H1. intros H2. rewrite <- H2.
apply E_WhileFalse. apply Heqr. Qed.
(** **** Exercise: 4 stars, standard (ceval_step__ceval_inf)
Write an informal proof of [ceval_step__ceval], following the
usual template. (The template for case analysis on an inductively
defined value should look the same as for induction, except that
there is no induction hypothesis.) Make your proof communicate
the main ideas to a human reader; do not simply transcribe the
steps of the formal proof. *)
(* FILL IN HERE *)
(* Do not modify the following line: *)
Definition manual_grade_for_ceval_step__ceval_inf : option (nat*string) := None.
(** [] *)
Theorem ceval_step_more: forall i1 i2 st st' c,
i1 <= i2 ->
ceval_step st c i1 = Some st' ->
ceval_step st c i2 = Some st'.
Proof.
induction i1 as [|i1']; intros i2 st st' c Hle Hceval.
- (* i1 = 0 *)
simpl in Hceval. discriminate Hceval.
- (* i1 = S i1' *)
destruct i2 as [|i2']. inversion Hle.
assert (Hle': i1' <= i2') by omega.
destruct c.
+ (* SKIP *)
simpl in Hceval. inversion Hceval.
reflexivity.
+ (* ::= *)
simpl in Hceval. inversion Hceval.
reflexivity.
+ (* ;; *)
simpl in Hceval. simpl.
destruct (ceval_step st c1 i1') eqn:Heqst1'o.
* (* st1'o = Some *)
apply (IHi1' i2') in Heqst1'o; try assumption.
rewrite Heqst1'o. simpl. simpl in Hceval.
apply (IHi1' i2') in Hceval; try assumption.
* (* st1'o = None *)
discriminate Hceval.
+ (* TEST *)
simpl in Hceval. simpl.
destruct (beval st b); apply (IHi1' i2') in Hceval;
assumption.
+ (* WHILE *)
simpl in Hceval. simpl.
destruct (beval st b); try assumption.
destruct (ceval_step st c i1') eqn: Heqst1'o.
* (* st1'o = Some *)
apply (IHi1' i2') in Heqst1'o; try assumption.
rewrite -> Heqst1'o. simpl. simpl in Hceval.
apply (IHi1' i2') in Hceval; try assumption.
* (* i1'o = None *)
simpl in Hceval. discriminate Hceval. Qed.
(** **** Exercise: 3 stars, standard, recommended (ceval__ceval_step)
Finish the following proof. You'll need [ceval_step_more] in a
few places, as well as some basic facts about [<=] and [plus]. *)
Theorem ceval__ceval_step: forall c st st',
st =[ c ]=> st' ->
exists i, ceval_step st c i = Some st'.
Proof.
intros c st st' Hce.
induction Hce.
(* FILL IN HERE *) Admitted.
(** [] *)
Theorem ceval_and_ceval_step_coincide: forall c st st',
st =[ c ]=> st'
<-> exists i, ceval_step st c i = Some st'.
Proof.
intros c st st'.
split. apply ceval__ceval_step. apply ceval_step__ceval.
Qed.
(* ################################################################# *)
(** * Determinism of Evaluation Again *)
(** Using the fact that the relational and step-indexed definition of
evaluation are the same, we can give a slicker proof that the
evaluation _relation_ is deterministic. *)
Theorem ceval_deterministic' : forall c st st1 st2,
st =[ c ]=> st1 ->
st =[ c ]=> st2 ->
st1 = st2.
Proof.
intros c st st1 st2 He1 He2.
apply ceval__ceval_step in He1.
apply ceval__ceval_step in He2.
inversion He1 as [i1 E1].
inversion He2 as [i2 E2].
apply ceval_step_more with (i2 := i1 + i2) in E1.
apply ceval_step_more with (i2 := i1 + i2) in E2.
rewrite E1 in E2. inversion E2. reflexivity.
omega. omega. Qed.
(* Wed Jan 9 12:02:46 EST 2019 *)

88
ImpCEvalFunTest.v Normal file
View file

@ -0,0 +1,88 @@
Set Warnings "-notation-overridden,-parsing".
From Coq Require Export String.
From LF Require Import ImpCEvalFun.
Parameter MISSING: Type.
Module Check.
Ltac check_type A B :=
match type of A with
| context[MISSING] => idtac "Missing:" A
| ?T => first [unify T B; idtac "Type: ok" | idtac "Type: wrong - should be (" B ")"]
end.
Ltac print_manual_grade A :=
match eval compute in A with
| Some (_ ?S ?C) =>
idtac "Score:" S;
match eval compute in C with
| ""%string => idtac "Comment: None"
| _ => idtac "Comment:" C
end
| None =>
idtac "Score: Ungraded";
idtac "Comment: None"
end.
End Check.
From LF Require Import ImpCEvalFun.
Import Check.
Goal True.
idtac "------------------- pup_to_n --------------------".
idtac " ".
idtac "#> pup_to_n".
idtac "Possible points: 2".
check_type @pup_to_n (Imp.com).
idtac "Assumptions:".
Abort.
Print Assumptions pup_to_n.
Goal True.
idtac " ".
idtac "------------------- ceval_step__ceval_inf --------------------".
idtac " ".
idtac "#> Manually graded: ceval_step__ceval_inf".
idtac "Possible points: 4".
print_manual_grade manual_grade_for_ceval_step__ceval_inf.
idtac " ".
idtac "------------------- ceval__ceval_step --------------------".
idtac " ".
idtac "#> ceval__ceval_step".
idtac "Possible points: 3".
check_type @ceval__ceval_step (
(forall (c : Imp.com) (st st' : Imp.state),
Imp.ceval c st st' ->
exists i : nat, ceval_step st c i = @Some Imp.state st')).
idtac "Assumptions:".
Abort.
Print Assumptions ceval__ceval_step.
Goal True.
idtac " ".
idtac " ".
idtac "Max points - standard: 9".
idtac "Max points - advanced: 9".
idtac "".
idtac "********** Summary **********".
idtac "".
idtac "********** Standard **********".
idtac "---------- pup_to_n ---------".
Print Assumptions pup_to_n.
idtac "---------- ceval_step__ceval_inf ---------".
idtac "MANUAL".
idtac "---------- ceval__ceval_step ---------".
Print Assumptions ceval__ceval_step.
idtac "".
idtac "********** Advanced **********".
Abort.
(* Wed Jan 9 12:02:27 EST 2019 *)

464
ImpParser.v Normal file
View file

@ -0,0 +1,464 @@
(** * ImpParser: Lexing and Parsing in Coq *)
(** The development of the Imp language in [Imp.v] completely ignores
issues of concrete syntax -- how an ascii string that a programmer
might write gets translated into abstract syntax trees defined by
the datatypes [aexp], [bexp], and [com]. In this chapter, we
illustrate how the rest of the story can be filled in by building
a simple lexical analyzer and parser using Coq's functional
programming facilities. *)
(** It is not important to understand all the details here (and
accordingly, the explanations are fairly terse and there are no
exercises). The main point is simply to demonstrate that it can
be done. You are invited to look through the code -- most of it
is not very complicated, though the parser relies on some
"monadic" programming idioms that may require a little work to
make out -- but most readers will probably want to just skim down
to the Examples section at the very end to get the punchline. *)
Set Warnings "-notation-overridden,-parsing".
From Coq Require Import Strings.String.
From Coq Require Import Strings.Ascii.
From Coq Require Import Arith.Arith.
From Coq Require Import Init.Nat.
From Coq Require Import Arith.EqNat.
From Coq Require Import Lists.List.
Import ListNotations.
From LF Require Import Maps Imp.
(* ################################################################# *)
(** * Internals *)
(* ================================================================= *)
(** ** Lexical Analysis *)
Definition isWhite (c : ascii) : bool :=
let n := nat_of_ascii c in
orb (orb (n =? 32) (* space *)
(n =? 9)) (* tab *)
(orb (n =? 10) (* linefeed *)
(n =? 13)). (* Carriage return. *)
Notation "x '<=?' y" := (x <=? y)
(at level 70, no associativity) : nat_scope.
Definition isLowerAlpha (c : ascii) : bool :=
let n := nat_of_ascii c in
andb (97 <=? n) (n <=? 122).
Definition isAlpha (c : ascii) : bool :=
let n := nat_of_ascii c in
orb (andb (65 <=? n) (n <=? 90))
(andb (97 <=? n) (n <=? 122)).
Definition isDigit (c : ascii) : bool :=
let n := nat_of_ascii c in
andb (48 <=? n) (n <=? 57).
Inductive chartype := white | alpha | digit | other.
Definition classifyChar (c : ascii) : chartype :=
if isWhite c then
white
else if isAlpha c then
alpha
else if isDigit c then
digit
else
other.
Fixpoint list_of_string (s : string) : list ascii :=
match s with
| EmptyString => []
| String c s => c :: (list_of_string s)
end.
Fixpoint string_of_list (xs : list ascii) : string :=
fold_right String EmptyString xs.
Definition token := string.
Fixpoint tokenize_helper (cls : chartype) (acc xs : list ascii)
: list (list ascii) :=
let tk := match acc with [] => [] | _::_ => [rev acc] end in
match xs with
| [] => tk
| (x::xs') =>
match cls, classifyChar x, x with
| _, _, "(" =>
tk ++ ["("]::(tokenize_helper other [] xs')
| _, _, ")" =>
tk ++ [")"]::(tokenize_helper other [] xs')
| _, white, _ =>
tk ++ (tokenize_helper white [] xs')
| alpha,alpha,x =>
tokenize_helper alpha (x::acc) xs'
| digit,digit,x =>
tokenize_helper digit (x::acc) xs'
| other,other,x =>
tokenize_helper other (x::acc) xs'
| _,tp,x =>
tk ++ (tokenize_helper tp [x] xs')
end
end %char.
Definition tokenize (s : string) : list string :=
map string_of_list (tokenize_helper white [] (list_of_string s)).
Example tokenize_ex1 :
tokenize "abc12=3 223*(3+(a+c))" %string
= ["abc"; "12"; "="; "3"; "223";
"*"; "("; "3"; "+"; "(";
"a"; "+"; "c"; ")"; ")"]%string.
Proof. reflexivity. Qed.
(* ================================================================= *)
(** ** Parsing *)
(* ----------------------------------------------------------------- *)
(** *** Options With Errors *)
(** An [option] type with error messages: *)
Inductive optionE (X:Type) : Type :=
| SomeE (x : X)
| NoneE (s : string).
Arguments SomeE {X}.
Arguments NoneE {X}.
(** Some syntactic sugar to make writing nested match-expressions on
optionE more convenient. *)
Notation "' p <- e1 ;; e2"
:= (match e1 with
| SomeE p => e2
| NoneE err => NoneE err
end)
(right associativity, p pattern, at level 60, e1 at next level).
Notation "'TRY' ' p <- e1 ;; e2 'OR' e3"
:= (match e1 with
| SomeE p => e2
| NoneE _ => e3
end)
(right associativity, p pattern,
at level 60, e1 at next level, e2 at next level).
(* ----------------------------------------------------------------- *)
(** *** Generic Combinators for Building Parsers *)
Open Scope string_scope.
Definition parser (T : Type) :=
list token -> optionE (T * list token).
Fixpoint many_helper {T} (p : parser T) acc steps xs :=
match steps, p xs with
| 0, _ =>
NoneE "Too many recursive calls"
| _, NoneE _ =>
SomeE ((rev acc), xs)
| S steps', SomeE (t, xs') =>
many_helper p (t :: acc) steps' xs'
end.
(** A (step-indexed) parser that expects zero or more [p]s: *)
Fixpoint many {T} (p : parser T) (steps : nat) : parser (list T) :=
many_helper p [] steps.
(** A parser that expects a given token, followed by [p]: *)
Definition firstExpect {T} (t : token) (p : parser T)
: parser T :=
fun xs => match xs with
| x::xs' =>
if string_dec x t
then p xs'
else NoneE ("expected '" ++ t ++ "'.")
| [] =>
NoneE ("expected '" ++ t ++ "'.")
end.
(** A parser that expects a particular token: *)
Definition expect (t : token) : parser unit :=
firstExpect t (fun xs => SomeE (tt, xs)).
(* ----------------------------------------------------------------- *)
(** *** A Recursive-Descent Parser for Imp *)
(** Identifiers: *)
Definition parseIdentifier (xs : list token)
: optionE (string * list token) :=
match xs with
| [] => NoneE "Expected identifier"
| x::xs' =>
if forallb isLowerAlpha (list_of_string x) then
SomeE (x, xs')
else
NoneE ("Illegal identifier:'" ++ x ++ "'")
end.
(** Numbers: *)
Definition parseNumber (xs : list token)
: optionE (nat * list token) :=
match xs with
| [] => NoneE "Expected number"
| x::xs' =>
if forallb isDigit (list_of_string x) then
SomeE (fold_left
(fun n d =>
10 * n + (nat_of_ascii d -
nat_of_ascii "0"%char))
(list_of_string x)
0,
xs')
else
NoneE "Expected number"
end.
(** Parse arithmetic expressions *)
Fixpoint parsePrimaryExp (steps:nat)
(xs : list token)
: optionE (aexp * list token) :=
match steps with
| 0 => NoneE "Too many recursive calls"
| S steps' =>
TRY ' (i, rest) <- parseIdentifier xs ;;
SomeE (AId i, rest)
OR
TRY ' (n, rest) <- parseNumber xs ;;
SomeE (ANum n, rest)
OR
' (e, rest) <- firstExpect "(" (parseSumExp steps') xs ;;
' (u, rest') <- expect ")" rest ;;
SomeE (e,rest')
end
with parseProductExp (steps:nat)
(xs : list token) :=
match steps with
| 0 => NoneE "Too many recursive calls"
| S steps' =>
' (e, rest) <- parsePrimaryExp steps' xs ;;
' (es, rest') <- many (firstExpect "*" (parsePrimaryExp steps'))
steps' rest ;;
SomeE (fold_left AMult es e, rest')
end
with parseSumExp (steps:nat) (xs : list token) :=
match steps with
| 0 => NoneE "Too many recursive calls"
| S steps' =>
' (e, rest) <- parseProductExp steps' xs ;;
' (es, rest') <-
many (fun xs =>
TRY ' (e,rest') <-
firstExpect "+"
(parseProductExp steps') xs ;;
SomeE ( (true, e), rest')
OR
' (e, rest') <-
firstExpect "-"
(parseProductExp steps') xs ;;
SomeE ( (false, e), rest'))
steps' rest ;;
SomeE (fold_left (fun e0 term =>
match term with
| (true, e) => APlus e0 e
| (false, e) => AMinus e0 e
end)
es e,
rest')
end.
Definition parseAExp := parseSumExp.
(** Parsing boolean expressions: *)
Fixpoint parseAtomicExp (steps:nat)
(xs : list token) :=
match steps with
| 0 => NoneE "Too many recursive calls"
| S steps' =>
TRY ' (u,rest) <- expect "true" xs ;;
SomeE (BTrue,rest)
OR
TRY ' (u,rest) <- expect "false" xs ;;
SomeE (BFalse,rest)
OR
TRY ' (e,rest) <- firstExpect "~"
(parseAtomicExp steps')
xs ;;
SomeE (BNot e, rest)
OR
TRY ' (e,rest) <- firstExpect "("
(parseConjunctionExp steps')
xs ;;
' (u,rest') <- expect ")" rest ;;
SomeE (e, rest')
OR
' (e, rest) <- parseProductExp steps' xs ;;
TRY ' (e', rest') <- firstExpect "="
(parseAExp steps') rest ;;
SomeE (BEq e e', rest')
OR
TRY ' (e', rest') <- firstExpect "<="
(parseAExp steps') rest ;;
SomeE (BLe e e', rest')
OR
NoneE "Expected '=' or '<=' after arithmetic expression"
end
with parseConjunctionExp (steps:nat)
(xs : list token) :=
match steps with
| 0 => NoneE "Too many recursive calls"
| S steps' =>
' (e, rest) <- parseAtomicExp steps' xs ;;
' (es, rest') <- many (firstExpect "&&"
(parseAtomicExp steps'))
steps' rest ;;
SomeE (fold_left BAnd es e, rest')
end.
Definition parseBExp := parseConjunctionExp.
Check parseConjunctionExp.
Definition testParsing {X : Type}
(p : nat ->
list token ->
optionE (X * list token))
(s : string) :=
let t := tokenize s in
p 100 t.
(*
Eval compute in
testParsing parseProductExp "x.y.(x.x).x".
Eval compute in
testParsing parseConjunctionExp "~(x=x&&x*x<=(x*x)*x)&&x=x".
*)
(** Parsing commands: *)
Fixpoint parseSimpleCommand (steps:nat)
(xs : list token) :=
match steps with
| 0 => NoneE "Too many recursive calls"
| S steps' =>
TRY ' (u, rest) <- expect "SKIP" xs ;;
SomeE (SKIP%imp, rest)
OR
TRY ' (e,rest) <-
firstExpect "TEST"
(parseBExp steps') xs ;;
' (c,rest') <-
firstExpect "THEN"
(parseSequencedCommand steps') rest ;;
' (c',rest'') <-
firstExpect "ELSE"
(parseSequencedCommand steps') rest' ;;
' (tt,rest''') <-
expect "END" rest'' ;;
SomeE(TEST e THEN c ELSE c' FI%imp, rest''')
OR
TRY ' (e,rest) <-
firstExpect "WHILE"
(parseBExp steps') xs ;;
' (c,rest') <-
firstExpect "DO"
(parseSequencedCommand steps') rest ;;
' (u,rest'') <-
expect "END" rest' ;;
SomeE(WHILE e DO c END%imp, rest'')
OR
TRY ' (i, rest) <- parseIdentifier xs ;;
' (e, rest') <- firstExpect "::=" (parseAExp steps') rest ;;
SomeE ((i ::= e)%imp, rest')
OR
NoneE "Expecting a command"
end
with parseSequencedCommand (steps:nat)
(xs : list token) :=
match steps with
| 0 => NoneE "Too many recursive calls"
| S steps' =>
' (c, rest) <- parseSimpleCommand steps' xs ;;
TRY ' (c', rest') <-
firstExpect ";;"
(parseSequencedCommand steps') rest ;;
SomeE ((c ;; c')%imp, rest')
OR
SomeE (c, rest)
end.
Definition bignumber := 1000.
Definition parse (str : string) : optionE com :=
let tokens := tokenize str in
match parseSequencedCommand bignumber tokens with
| SomeE (c, []) => SomeE c
| SomeE (_, t::_) => NoneE ("Trailing tokens remaining: " ++ t)
| NoneE err => NoneE err
end.
(* ################################################################# *)
(** * Examples *)
Example eg1 : parse "
TEST x = y + 1 + 2 - y * 6 + 3 THEN
x ::= x * 1;;
y ::= 0
ELSE
SKIP
END "
=
SomeE (
TEST "x" = "y" + 1 + 2 - "y" * 6 + 3 THEN
"x" ::= "x" * 1;;
"y" ::= 0
ELSE
SKIP
FI)%imp.
Proof. cbv. reflexivity. Qed.
Example eg2 : parse "
SKIP;;
z::=x*y*(x*x);;
WHILE x=x DO
TEST (z <= z*z) && ~(x = 2) THEN
x ::= z;;
y ::= z
ELSE
SKIP
END;;
SKIP
END;;
x::=z "
=
SomeE (
SKIP;;
"z" ::= "x" * "y" * ("x" * "x");;
WHILE "x" = "x" DO
TEST ("z" <= "z" * "z") && ~("x" = 2) THEN
"x" ::= "z";;
"y" ::= "z"
ELSE
SKIP
FI;;
SKIP
END;;
"x" ::= "z")%imp.
Proof. cbv. reflexivity. Qed.
(* Wed Jan 9 12:02:46 EST 2019 *)

47
ImpParserTest.v Normal file
View file

@ -0,0 +1,47 @@
Set Warnings "-notation-overridden,-parsing".
From Coq Require Export String.
From LF Require Import ImpParser.
Parameter MISSING: Type.
Module Check.
Ltac check_type A B :=
match type of A with
| context[MISSING] => idtac "Missing:" A
| ?T => first [unify T B; idtac "Type: ok" | idtac "Type: wrong - should be (" B ")"]
end.
Ltac print_manual_grade A :=
match eval compute in A with
| Some (_ ?S ?C) =>
idtac "Score:" S;
match eval compute in C with
| ""%string => idtac "Comment: None"
| _ => idtac "Comment:" C
end
| None =>
idtac "Score: Ungraded";
idtac "Comment: None"
end.
End Check.
From LF Require Import ImpParser.
Import Check.
Goal True.
idtac " ".
idtac "Max points - standard: 0".
idtac "Max points - advanced: 0".
idtac "".
idtac "********** Summary **********".
idtac "".
idtac "********** Standard **********".
idtac "".
idtac "********** Advanced **********".
Abort.
(* Wed Jan 9 12:02:26 EST 2019 *)

249
ImpTest.v Normal file
View file

@ -0,0 +1,249 @@
Set Warnings "-notation-overridden,-parsing".
From Coq Require Export String.
From LF Require Import Imp.
Parameter MISSING: Type.
Module Check.
Ltac check_type A B :=
match type of A with
| context[MISSING] => idtac "Missing:" A
| ?T => first [unify T B; idtac "Type: ok" | idtac "Type: wrong - should be (" B ")"]
end.
Ltac print_manual_grade A :=
match eval compute in A with
| Some (_ ?S ?C) =>
idtac "Score:" S;
match eval compute in C with
| ""%string => idtac "Comment: None"
| _ => idtac "Comment:" C
end
| None =>
idtac "Score: Ungraded";
idtac "Comment: None"
end.
End Check.
From LF Require Import Imp.
Import Check.
Goal True.
idtac "------------------- optimize_0plus_b_sound --------------------".
idtac " ".
idtac "#> AExp.optimize_0plus_b_sound".
idtac "Possible points: 3".
check_type @AExp.optimize_0plus_b_sound (
(forall b : AExp.bexp, AExp.beval (AExp.optimize_0plus_b b) = AExp.beval b)).
idtac "Assumptions:".
Abort.
Print Assumptions AExp.optimize_0plus_b_sound.
Goal True.
idtac " ".
idtac "------------------- bevalR --------------------".
idtac " ".
idtac "#> AExp.beval_iff_bevalR".
idtac "Possible points: 3".
check_type @AExp.beval_iff_bevalR (
(forall (b : AExp.bexp) (bv : bool), AExp.bevalR b bv <-> AExp.beval b = bv)).
idtac "Assumptions:".
Abort.
Print Assumptions AExp.beval_iff_bevalR.
Goal True.
idtac " ".
idtac "------------------- ceval_example2 --------------------".
idtac " ".
idtac "#> ceval_example2".
idtac "Possible points: 2".
check_type @ceval_example2 (
(empty_st =[ X ::= 0;; Y ::= 1;; Z ::= 2
]=> @Maps.t_update nat (@Maps.t_update nat (X !-> 0) Y 1) Z 2)).
idtac "Assumptions:".
Abort.
Print Assumptions ceval_example2.
Goal True.
idtac " ".
idtac "------------------- XtimesYinZ_spec --------------------".
idtac " ".
idtac "#> Manually graded: XtimesYinZ_spec".
idtac "Possible points: 3".
print_manual_grade manual_grade_for_XtimesYinZ_spec.
idtac " ".
idtac "------------------- loop_never_stops --------------------".
idtac " ".
idtac "#> loop_never_stops".
idtac "Possible points: 3".
check_type @loop_never_stops ((forall st st' : state, ~ st =[ loop ]=> st')).
idtac "Assumptions:".
Abort.
Print Assumptions loop_never_stops.
Goal True.
idtac " ".
idtac "------------------- no_whiles_eqv --------------------".
idtac " ".
idtac "#> no_whiles_eqv".
idtac "Possible points: 3".
check_type @no_whiles_eqv ((forall c : com, no_whiles c = true <-> no_whilesR c)).
idtac "Assumptions:".
Abort.
Print Assumptions no_whiles_eqv.
Goal True.
idtac " ".
idtac "------------------- no_whiles_terminating --------------------".
idtac " ".
idtac "#> Manually graded: no_whiles_terminating".
idtac "Possible points: 4".
print_manual_grade manual_grade_for_no_whiles_terminating.
idtac " ".
idtac "------------------- stack_compiler --------------------".
idtac " ".
idtac "#> s_execute1".
idtac "Possible points: 0.5".
check_type @s_execute1 (
(s_execute empty_st (@nil nat)
(SPush 5 :: (SPush 3 :: SPush 1 :: SMinus :: @nil sinstr)%list) =
(2 :: 5 :: @nil nat)%list)).
idtac "Assumptions:".
Abort.
Print Assumptions s_execute1.
Goal True.
idtac " ".
idtac "#> s_execute2".
idtac "Possible points: 0.5".
check_type @s_execute2 (
(s_execute (X !-> 3) (3 :: (4 :: @nil nat)%list)
(SPush 4 :: (SLoad X :: SMult :: SPlus :: @nil sinstr)%list) =
(15 :: 4 :: @nil nat)%list)).
idtac "Assumptions:".
Abort.
Print Assumptions s_execute2.
Goal True.
idtac " ".
idtac "#> s_compile1".
idtac "Possible points: 2".
check_type @s_compile1 (
(s_compile (X - 2 * Y) =
(SLoad X :: SPush 2 :: SLoad Y :: SMult :: SMinus :: @nil sinstr)%list)).
idtac "Assumptions:".
Abort.
Print Assumptions s_compile1.
Goal True.
idtac " ".
idtac "------------------- stack_compiler_correct --------------------".
idtac " ".
idtac "#> s_compile_correct".
idtac "Advanced".
idtac "Possible points: 4".
check_type @s_compile_correct (
(forall (st : state) (e : aexp),
s_execute st (@nil nat) (s_compile e) = (aeval st e :: @nil nat)%list)).
idtac "Assumptions:".
Abort.
Print Assumptions s_compile_correct.
Goal True.
idtac " ".
idtac "------------------- break_imp --------------------".
idtac " ".
idtac "#> BreakImp.break_ignore".
idtac "Advanced".
idtac "Possible points: 1".
check_type @BreakImp.break_ignore (
(forall (c : BreakImp.com) (st st' : state) (s : BreakImp.result),
BreakImp.ceval (BreakImp.CSeq BreakImp.CBreak c) st s st' -> st = st')).
idtac "Assumptions:".
Abort.
Print Assumptions BreakImp.break_ignore.
Goal True.
idtac " ".
idtac "#> BreakImp.while_continue".
idtac "Advanced".
idtac "Possible points: 1".
check_type @BreakImp.while_continue (
(forall (b : bexp) (c : BreakImp.com) (st st' : state) (s : BreakImp.result),
BreakImp.ceval (BreakImp.CWhile b c) st s st' -> s = BreakImp.SContinue)).
idtac "Assumptions:".
Abort.
Print Assumptions BreakImp.while_continue.
Goal True.
idtac " ".
idtac "#> BreakImp.while_stops_on_break".
idtac "Advanced".
idtac "Possible points: 2".
check_type @BreakImp.while_stops_on_break (
(forall (b : bexp) (c : BreakImp.com) (st st' : state),
beval st b = true ->
BreakImp.ceval c st BreakImp.SBreak st' ->
BreakImp.ceval (BreakImp.CWhile b c) st BreakImp.SContinue st')).
idtac "Assumptions:".
Abort.
Print Assumptions BreakImp.while_stops_on_break.
Goal True.
idtac " ".
idtac " ".
idtac "Max points - standard: 24".
idtac "Max points - advanced: 32".
idtac "".
idtac "********** Summary **********".
idtac "".
idtac "********** Standard **********".
idtac "---------- AExp.optimize_0plus_b_sound ---------".
Print Assumptions AExp.optimize_0plus_b_sound.
idtac "---------- AExp.beval_iff_bevalR ---------".
Print Assumptions AExp.beval_iff_bevalR.
idtac "---------- ceval_example2 ---------".
Print Assumptions ceval_example2.
idtac "---------- XtimesYinZ_spec ---------".
idtac "MANUAL".
idtac "---------- loop_never_stops ---------".
Print Assumptions loop_never_stops.
idtac "---------- no_whiles_eqv ---------".
Print Assumptions no_whiles_eqv.
idtac "---------- no_whiles_terminating ---------".
idtac "MANUAL".
idtac "---------- s_execute1 ---------".
Print Assumptions s_execute1.
idtac "---------- s_execute2 ---------".
Print Assumptions s_execute2.
idtac "---------- s_compile1 ---------".
Print Assumptions s_compile1.
idtac "".
idtac "********** Advanced **********".
idtac "---------- s_compile_correct ---------".
Print Assumptions s_compile_correct.
idtac "---------- BreakImp.break_ignore ---------".
Print Assumptions BreakImp.break_ignore.
idtac "---------- BreakImp.while_continue ---------".
Print Assumptions BreakImp.while_continue.
idtac "---------- BreakImp.while_stops_on_break ---------".
Print Assumptions BreakImp.while_stops_on_break.
Abort.
(* Wed Jan 9 12:02:24 EST 2019 *)

714
IndPrinciples.v Normal file
View file

@ -0,0 +1,714 @@
(** * IndPrinciples: Induction Principles *)
(** With the Curry-Howard correspondence and its realization in Coq in
mind, we can now take a deeper look at induction principles. *)
Set Warnings "-notation-overridden,-parsing".
From LF Require Export ProofObjects.
(* ################################################################# *)
(** * Basics *)
(** Every time we declare a new [Inductive] datatype, Coq
automatically generates an _induction principle_ for this type.
This induction principle is a theorem like any other: If [t] is
defined inductively, the corresponding induction principle is
called [t_ind]. Here is the one for natural numbers: *)
Check nat_ind.
(* ===> nat_ind :
forall P : nat -> Prop,
P 0 ->
(forall n : nat, P n -> P (S n)) ->
forall n : nat, P n *)
(** The [induction] tactic is a straightforward wrapper that, at its
core, simply performs [apply t_ind]. To see this more clearly,
let's experiment with directly using [apply nat_ind], instead of
the [induction] tactic, to carry out some proofs. Here, for
example, is an alternate proof of a theorem that we saw in the
[Basics] chapter. *)
Theorem mult_0_r' : forall n:nat,
n * 0 = 0.
Proof.
apply nat_ind.
- (* n = O *) reflexivity.
- (* n = S n' *) simpl. intros n' IHn'. rewrite -> IHn'.
reflexivity. Qed.
(** This proof is basically the same as the earlier one, but a
few minor differences are worth noting.
First, in the induction step of the proof (the ["S"] case), we
have to do a little bookkeeping manually (the [intros]) that
[induction] does automatically.
Second, we do not introduce [n] into the context before applying
[nat_ind] -- the conclusion of [nat_ind] is a quantified formula,
and [apply] needs this conclusion to exactly match the shape of
the goal state, including the quantifier. By contrast, the
[induction] tactic works either with a variable in the context or
a quantified variable in the goal.
These conveniences make [induction] nicer to use in practice than
applying induction principles like [nat_ind] directly. But it is
important to realize that, modulo these bits of bookkeeping,
applying [nat_ind] is what we are really doing. *)
(** **** Exercise: 2 stars, standard, optional (plus_one_r')
Complete this proof without using the [induction] tactic. *)
Theorem plus_one_r' : forall n:nat,
n + 1 = S n.
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** Coq generates induction principles for every datatype defined with
[Inductive], including those that aren't recursive. Although of
course we don't need induction to prove properties of
non-recursive datatypes, the idea of an induction principle still
makes sense for them: it gives a way to prove that a property
holds for all values of the type.
These generated principles follow a similar pattern. If we define
a type [t] with constructors [c1] ... [cn], Coq generates a
theorem with this shape:
t_ind : forall P : t -> Prop,
... case for c1 ... ->
... case for c2 ... -> ...
... case for cn ... ->
forall n : t, P n
The specific shape of each case depends on the arguments to the
corresponding constructor. Before trying to write down a general
rule, let's look at some more examples. First, an example where
the constructors take no arguments: *)
Inductive yesno : Type :=
| yes
| no.
Check yesno_ind.
(* ===> yesno_ind : forall P : yesno -> Prop,
P yes ->
P no ->
forall y : yesno, P y *)
(** **** Exercise: 1 star, standard, optional (rgb)
Write out the induction principle that Coq will generate for the
following datatype. Write down your answer on paper or type it
into a comment, and then compare it with what Coq prints. *)
Inductive rgb : Type :=
| red
| green
| blue.
Check rgb_ind.
(** [] *)
(** Here's another example, this time with one of the constructors
taking some arguments. *)
Inductive natlist : Type :=
| nnil
| ncons (n : nat) (l : natlist).
Check natlist_ind.
(* ===> (modulo a little variable renaming)
natlist_ind :
forall P : natlist -> Prop,
P nnil ->
(forall (n : nat) (l : natlist),
P l -> P (ncons n l)) ->
forall n : natlist, P n *)
(** **** Exercise: 1 star, standard, optional (natlist1)
Suppose we had written the above definition a little
differently: *)
Inductive natlist1 : Type :=
| nnil1
| nsnoc1 (l : natlist1) (n : nat).
(** Now what will the induction principle look like?
[] *)
(** From these examples, we can extract this general rule:
- The type declaration gives several constructors; each
corresponds to one clause of the induction principle.
- Each constructor [c] takes argument types [a1] ... [an].
- Each [ai] can be either [t] (the datatype we are defining) or
some other type [s].
- The corresponding case of the induction principle says:
- "For all values [x1]...[xn] of types [a1]...[an], if [P]
holds for each of the inductive arguments (each [xi] of type
[t]), then [P] holds for [c x1 ... xn]".
*)
(** **** Exercise: 1 star, standard, optional (byntree_ind)
Write out the induction principle that Coq will generate for the
following datatype. (Again, write down your answer on paper or
type it into a comment, and then compare it with what Coq
prints.) *)
Inductive byntree : Type :=
| bempty
| bleaf (yn : yesno)
| nbranch (yn : yesno) (t1 t2 : byntree).
(** [] *)
(** **** Exercise: 1 star, standard, optional (ex_set)
Here is an induction principle for an inductively defined
set.
ExSet_ind :
forall P : ExSet -> Prop,
(forall b : bool, P (con1 b)) ->
(forall (n : nat) (e : ExSet), P e -> P (con2 n e)) ->
forall e : ExSet, P e
Give an [Inductive] definition of [ExSet]: *)
Inductive ExSet : Type :=
(* FILL IN HERE *)
.
(** [] *)
(* ################################################################# *)
(** * Polymorphism *)
(** Next, what about polymorphic datatypes?
The inductive definition of polymorphic lists
Inductive list (X:Type) : Type :=
| nil : list X
| cons : X -> list X -> list X.
is very similar to that of [natlist]. The main difference is
that, here, the whole definition is _parameterized_ on a set [X]:
that is, we are defining a _family_ of inductive types [list X],
one for each [X]. (Note that, wherever [list] appears in the body
of the declaration, it is always applied to the parameter [X].)
The induction principle is likewise parameterized on [X]:
list_ind :
forall (X : Type) (P : list X -> Prop),
P [] ->
(forall (x : X) (l : list X), P l -> P (x :: l)) ->
forall l : list X, P l
Note that the _whole_ induction principle is parameterized on
[X]. That is, [list_ind] can be thought of as a polymorphic
function that, when applied to a type [X], gives us back an
induction principle specialized to the type [list X]. *)
(** **** Exercise: 1 star, standard, optional (tree)
Write out the induction principle that Coq will generate for
the following datatype. Compare your answer with what Coq
prints. *)
Inductive tree (X:Type) : Type :=
| leaf (x : X)
| node (t1 t2 : tree X).
Check tree_ind.
(** [] *)
(** **** Exercise: 1 star, standard, optional (mytype)
Find an inductive definition that gives rise to the
following induction principle:
mytype_ind :
forall (X : Type) (P : mytype X -> Prop),
(forall x : X, P (constr1 X x)) ->
(forall n : nat, P (constr2 X n)) ->
(forall m : mytype X, P m ->
forall n : nat, P (constr3 X m n)) ->
forall m : mytype X, P m
*)
(** [] *)
(** **** Exercise: 1 star, standard, optional (foo)
Find an inductive definition that gives rise to the
following induction principle:
foo_ind :
forall (X Y : Type) (P : foo X Y -> Prop),
(forall x : X, P (bar X Y x)) ->
(forall y : Y, P (baz X Y y)) ->
(forall f1 : nat -> foo X Y,
(forall n : nat, P (f1 n)) -> P (quux X Y f1)) ->
forall f2 : foo X Y, P f2
*)
(** [] *)
(** **** Exercise: 1 star, standard, optional (foo')
Consider the following inductive definition: *)
Inductive foo' (X:Type) : Type :=
| C1 (l : list X) (f : foo' X)
| C2.
(** What induction principle will Coq generate for [foo']? Fill
in the blanks, then check your answer with Coq.)
foo'_ind :
forall (X : Type) (P : foo' X -> Prop),
(forall (l : list X) (f : foo' X),
_______________________ ->
_______________________ ) ->
___________________________________________ ->
forall f : foo' X, ________________________
*)
(** [] *)
(* ################################################################# *)
(** * Induction Hypotheses *)
(** Where does the phrase "induction hypothesis" fit into this story?
The induction principle for numbers
forall P : nat -> Prop,
P 0 ->
(forall n : nat, P n -> P (S n)) ->
forall n : nat, P n
is a generic statement that holds for all propositions
[P] (or rather, strictly speaking, for all families of
propositions [P] indexed by a number [n]). Each time we
use this principle, we are choosing [P] to be a particular
expression of type [nat->Prop].
We can make proofs by induction more explicit by giving
this expression a name. For example, instead of stating
the theorem [mult_0_r] as "[forall n, n * 0 = 0]," we can
write it as "[forall n, P_m0r n]", where [P_m0r] is defined
as... *)
Definition P_m0r (n:nat) : Prop :=
n * 0 = 0.
(** ... or equivalently: *)
Definition P_m0r' : nat->Prop :=
fun n => n * 0 = 0.
(** Now it is easier to see where [P_m0r] appears in the proof. *)
Theorem mult_0_r'' : forall n:nat,
P_m0r n.
Proof.
apply nat_ind.
- (* n = O *) reflexivity.
- (* n = S n' *)
(* Note the proof state at this point! *)
intros n IHn.
unfold P_m0r in IHn. unfold P_m0r. simpl. apply IHn. Qed.
(** This extra naming step isn't something that we do in
normal proofs, but it is useful to do it explicitly for an example
or two, because it allows us to see exactly what the induction
hypothesis is. If we prove [forall n, P_m0r n] by induction on
[n] (using either [induction] or [apply nat_ind]), we see that the
first subgoal requires us to prove [P_m0r 0] ("[P] holds for
zero"), while the second subgoal requires us to prove [forall n',
P_m0r n' -> P_m0r (S n')] (that is "[P] holds of [S n'] if it
holds of [n']" or, more elegantly, "[P] is preserved by [S]").
The _induction hypothesis_ is the premise of this latter
implication -- the assumption that [P] holds of [n'], which we are
allowed to use in proving that [P] holds for [S n']. *)
(* ################################################################# *)
(** * More on the [induction] Tactic *)
(** The [induction] tactic actually does even more low-level
bookkeeping for us than we discussed above.
Recall the informal statement of the induction principle for
natural numbers:
- If [P n] is some proposition involving a natural number n, and
we want to show that P holds for _all_ numbers n, we can
reason like this:
- show that [P O] holds
- show that, if [P n'] holds, then so does [P (S n')]
- conclude that [P n] holds for all n.
So, when we begin a proof with [intros n] and then [induction n],
we are first telling Coq to consider a _particular_ [n] (by
introducing it into the context) and then telling it to prove
something about _all_ numbers (by using induction).
What Coq actually does in this situation, internally, is to
"re-generalize" the variable we perform induction on. For
example, in our original proof that [plus] is associative... *)
Theorem plus_assoc' : forall n m p : nat,
n + (m + p) = (n + m) + p.
Proof.
(* ...we first introduce all 3 variables into the context,
which amounts to saying "Consider an arbitrary [n], [m], and
[p]..." *)
intros n m p.
(* ...We now use the [induction] tactic to prove [P n] (that
is, [n + (m + p) = (n + m) + p]) for _all_ [n],
and hence also for the particular [n] that is in the context
at the moment. *)
induction n as [| n'].
- (* n = O *) reflexivity.
- (* n = S n' *)
(* In the second subgoal generated by [induction] -- the
"inductive step" -- we must prove that [P n'] implies
[P (S n')] for all [n']. The [induction] tactic
automatically introduces [n'] and [P n'] into the context
for us, leaving just [P (S n')] as the goal. *)
simpl. rewrite -> IHn'. reflexivity. Qed.
(** It also works to apply [induction] to a variable that is
quantified in the goal. *)
Theorem plus_comm' : forall n m : nat,
n + m = m + n.
Proof.
induction n as [| n'].
- (* n = O *) intros m. rewrite <- plus_n_O. reflexivity.
- (* n = S n' *) intros m. simpl. rewrite -> IHn'.
rewrite <- plus_n_Sm. reflexivity. Qed.
(** Note that [induction n] leaves [m] still bound in the goal --
i.e., what we are proving inductively is a statement beginning
with [forall m].
If we do [induction] on a variable that is quantified in the goal
_after_ some other quantifiers, the [induction] tactic will
automatically introduce the variables bound by these quantifiers
into the context. *)
Theorem plus_comm'' : forall n m : nat,
n + m = m + n.
Proof.
(* Let's do induction on [m] this time, instead of [n]... *)
induction m as [| m'].
- (* m = O *) simpl. rewrite <- plus_n_O. reflexivity.
- (* m = S m' *) simpl. rewrite <- IHm'.
rewrite <- plus_n_Sm. reflexivity. Qed.
(** **** Exercise: 1 star, standard, optional (plus_explicit_prop)
Rewrite both [plus_assoc'] and [plus_comm'] and their proofs in
the same style as [mult_0_r''] above -- that is, for each theorem,
give an explicit [Definition] of the proposition being proved by
induction, and state the theorem and proof in terms of this
defined proposition. *)
(* FILL IN HERE
[] *)
(* ################################################################# *)
(** * Induction Principles in [Prop] *)
(** Earlier, we looked in detail at the induction principles that Coq
generates for inductively defined _sets_. The induction
principles for inductively defined _propositions_ like [even] are a
tiny bit more complicated. As with all induction principles, we
want to use the induction principle on [even] to prove things by
inductively considering the possible shapes that something in [even]
can have. Intuitively speaking, however, what we want to prove
are not statements about _evidence_ but statements about
_numbers_: accordingly, we want an induction principle that lets
us prove properties of numbers by induction on evidence.
For example, from what we've said so far, you might expect the
inductive definition of [even]...
Inductive even : nat -> Prop :=
| ev_0 : even 0
| ev_SS : forall n : nat, even n -> even (S (S n)).
...to give rise to an induction principle that looks like this...
ev_ind_max : forall P : (forall n : nat, even n -> Prop),
P O ev_0 ->
(forall (m : nat) (E : even m),
P m E ->
P (S (S m)) (ev_SS m E)) ->
forall (n : nat) (E : even n),
P n E
... because:
- Since [even] is indexed by a number [n] (every [even] object [E] is
a piece of evidence that some particular number [n] is even),
the proposition [P] is parameterized by both [n] and [E] --
that is, the induction principle can be used to prove
assertions involving both an even number and the evidence that
it is even.
- Since there are two ways of giving evidence of evenness ([even]
has two constructors), applying the induction principle
generates two subgoals:
- We must prove that [P] holds for [O] and [ev_0].
- We must prove that, whenever [n] is an even number and [E]
is an evidence of its evenness, if [P] holds of [n] and
[E], then it also holds of [S (S n)] and [ev_SS n E].
- If these subgoals can be proved, then the induction principle
tells us that [P] is true for _all_ even numbers [n] and
evidence [E] of their evenness.
This is more flexibility than we normally need or want: it is
giving us a way to prove logical assertions where the assertion
involves properties of some piece of _evidence_ of evenness, while
all we really care about is proving properties of _numbers_ that
are even -- we are interested in assertions about numbers, not
about evidence. It would therefore be more convenient to have an
induction principle for proving propositions [P] that are
parameterized just by [n] and whose conclusion establishes [P] for
all even numbers [n]:
forall P : nat -> Prop,
... ->
forall n : nat,
even n -> P n
For this reason, Coq actually generates the following simplified
induction principle for [even]: *)
Check even_ind.
(* ===> ev_ind
: forall P : nat -> Prop,
P 0 ->
(forall n : nat, even n -> P n -> P (S (S n))) ->
forall n : nat,
even n -> P n *)
(** In particular, Coq has dropped the evidence term [E] as a
parameter of the the proposition [P]. *)
(** In English, [ev_ind] says:
- Suppose, [P] is a property of natural numbers (that is, [P n] is
a [Prop] for every [n]). To show that [P n] holds whenever [n]
is even, it suffices to show:
- [P] holds for [0],
- for any [n], if [n] is even and [P] holds for [n], then [P]
holds for [S (S n)]. *)
(** As expected, we can apply [ev_ind] directly instead of using
[induction]. For example, we can use it to show that [even'] (the
slightly awkward alternate definition of evenness that we saw in
an exercise in the \chap{IndProp} chapter) is equivalent to the
cleaner inductive definition [even]: *)
Theorem ev_ev' : forall n, even n -> even' n.
Proof.
apply even_ind.
- (* ev_0 *)
apply even'_0.
- (* ev_SS *)
intros m Hm IH.
apply (even'_sum 2 m).
+ apply even'_2.
+ apply IH.
Qed.
(** The precise form of an [Inductive] definition can affect the
induction principle Coq generates.
For example, in chapter [IndProp], we defined [<=] as: *)
(* Inductive le : nat -> nat -> Prop :=
| le_n : forall n, le n n
| le_S : forall n m, (le n m) -> (le n (S m)). *)
(** This definition can be streamlined a little by observing that the
left-hand argument [n] is the same everywhere in the definition,
so we can actually make it a "general parameter" to the whole
definition, rather than an argument to each constructor. *)
Inductive le (n:nat) : nat -> Prop :=
| le_n : le n n
| le_S m (H : le n m) : le n (S m).
Notation "m <= n" := (le m n).
(** The second one is better, even though it looks less symmetric.
Why? Because it gives us a simpler induction principle. *)
Check le_ind.
(* ===> forall (n : nat) (P : nat -> Prop),
P n ->
(forall m : nat, n <= m -> P m -> P (S m)) ->
forall n0 : nat, n <= n0 -> P n0 *)
(* ################################################################# *)
(** * Formal vs. Informal Proofs by Induction *)
(** Question: What is the relation between a formal proof of a
proposition [P] and an informal proof of the same proposition [P]?
Answer: The latter should _teach_ the reader how to produce the
former.
Question: How much detail is needed??
Unfortunately, there is no single right answer; rather, there is a
range of choices.
At one end of the spectrum, we can essentially give the reader the
whole formal proof (i.e., the "informal" proof will amount to just
transcribing the formal one into words). This may give the reader
the ability to reproduce the formal one for themselves, but it
probably doesn't _teach_ them anything much.
At the other end of the spectrum, we can say "The theorem is true
and you can figure out why for yourself if you think about it hard
enough." This is also not a good teaching strategy, because often
writing the proof requires one or more significant insights into
the thing we're proving, and most readers will give up before they
rediscover all the same insights as we did.
In the middle is the golden mean -- a proof that includes all of
the essential insights (saving the reader the hard work that we
went through to find the proof in the first place) plus high-level
suggestions for the more routine parts to save the reader from
spending too much time reconstructing these (e.g., what the IH says
and what must be shown in each case of an inductive proof), but not
so much detail that the main ideas are obscured.
Since we've spent much of this chapter looking "under the hood" at
formal proofs by induction, now is a good moment to talk a little
about _informal_ proofs by induction.
In the real world of mathematical communication, written proofs
range from extremely longwinded and pedantic to extremely brief and
telegraphic. Although the ideal is somewhere in between, while one
is getting used to the style it is better to start out at the
pedantic end. Also, during the learning phase, it is probably
helpful to have a clear standard to compare against. With this in
mind, we offer two templates -- one for proofs by induction over
_data_ (i.e., where the thing we're doing induction on lives in
[Type]) and one for proofs by induction over _evidence_ (i.e.,
where the inductively defined thing lives in [Prop]). *)
(* ================================================================= *)
(** ** Induction Over an Inductively Defined Set *)
(** _Template_:
- _Theorem_: <Universally quantified proposition of the form
"For all [n:S], [P(n)]," where [S] is some inductively defined
set.>
_Proof_: By induction on [n].
<one case for each constructor [c] of [S]...>
- Suppose [n = c a1 ... ak], where <...and here we state
the IH for each of the [a]'s that has type [S], if any>.
We must show <...and here we restate [P(c a1 ... ak)]>.
<go on and prove [P(n)] to finish the case...>
- <other cases similarly...> []
_Example_:
- _Theorem_: For all sets [X], lists [l : list X], and numbers
[n], if [length l = n] then [index (S n) l = None].
_Proof_: By induction on [l].
- Suppose [l = []]. We must show, for all numbers [n],
that, if [length [] = n], then [index (S n) [] =
None].
This follows immediately from the definition of [index].
- Suppose [l = x :: l'] for some [x] and [l'], where
[length l' = n'] implies [index (S n') l' = None], for
any number [n']. We must show, for all [n], that, if
[length (x::l') = n] then [index (S n) (x::l') =
None].
Let [n] be a number with [length l = n]. Since
length l = length (x::l') = S (length l'),
it suffices to show that
index (S (length l')) l' = None.
But this follows directly from the induction hypothesis,
picking [n'] to be [length l']. [] *)
(* ================================================================= *)
(** ** Induction Over an Inductively Defined Proposition *)
(** Since inductively defined proof objects are often called
"derivation trees," this form of proof is also known as _induction
on derivations_.
_Template_:
- _Theorem_: <Proposition of the form "[Q -> P]," where [Q] is
some inductively defined proposition (more generally,
"For all [x] [y] [z], [Q x y z -> P x y z]")>
_Proof_: By induction on a derivation of [Q]. <Or, more
generally, "Suppose we are given [x], [y], and [z]. We
show that [Q x y z] implies [P x y z], by induction on a
derivation of [Q x y z]"...>
<one case for each constructor [c] of [Q]...>
- Suppose the final rule used to show [Q] is [c]. Then
<...and here we state the types of all of the [a]'s
together with any equalities that follow from the
definition of the constructor and the IH for each of
the [a]'s that has type [Q], if there are any>. We must
show <...and here we restate [P]>.
<go on and prove [P] to finish the case...>
- <other cases similarly...> []
_Example_
- _Theorem_: The [<=] relation is transitive -- i.e., for all
numbers [n], [m], and [o], if [n <= m] and [m <= o], then
[n <= o].
_Proof_: By induction on a derivation of [m <= o].
- Suppose the final rule used to show [m <= o] is
[le_n]. Then [m = o] and we must show that [n <= m],
which is immediate by hypothesis.
- Suppose the final rule used to show [m <= o] is
[le_S]. Then [o = S o'] for some [o'] with [m <= o'].
We must show that [n <= S o'].
By induction hypothesis, [n <= o'].
But then, by [le_S], [n <= S o']. [] *)
(* Wed Jan 9 12:02:46 EST 2019 *)

47
IndPrinciplesTest.v Normal file
View file

@ -0,0 +1,47 @@
Set Warnings "-notation-overridden,-parsing".
From Coq Require Export String.
From LF Require Import IndPrinciples.
Parameter MISSING: Type.
Module Check.
Ltac check_type A B :=
match type of A with
| context[MISSING] => idtac "Missing:" A
| ?T => first [unify T B; idtac "Type: ok" | idtac "Type: wrong - should be (" B ")"]
end.
Ltac print_manual_grade A :=
match eval compute in A with
| Some (_ ?S ?C) =>
idtac "Score:" S;
match eval compute in C with
| ""%string => idtac "Comment: None"
| _ => idtac "Comment:" C
end
| None =>
idtac "Score: Ungraded";
idtac "Comment: None"
end.
End Check.
From LF Require Import IndPrinciples.
Import Check.
Goal True.
idtac " ".
idtac "Max points - standard: 0".
idtac "Max points - advanced: 0".
idtac "".
idtac "********** Summary **********".
idtac "".
idtac "********** Standard **********".
idtac "".
idtac "********** Advanced **********".
Abort.
(* Wed Jan 9 12:02:19 EST 2019 *)

2089
IndProp.v Normal file

File diff suppressed because it is too large Load diff

295
IndPropTest.v Normal file
View file

@ -0,0 +1,295 @@
Set Warnings "-notation-overridden,-parsing".
From Coq Require Export String.
From LF Require Import IndProp.
Parameter MISSING: Type.
Module Check.
Ltac check_type A B :=
match type of A with
| context[MISSING] => idtac "Missing:" A
| ?T => first [unify T B; idtac "Type: ok" | idtac "Type: wrong - should be (" B ")"]
end.
Ltac print_manual_grade A :=
match eval compute in A with
| Some (_ ?S ?C) =>
idtac "Score:" S;
match eval compute in C with
| ""%string => idtac "Comment: None"
| _ => idtac "Comment:" C
end
| None =>
idtac "Score: Ungraded";
idtac "Comment: None"
end.
End Check.
From LF Require Import IndProp.
Import Check.
Goal True.
idtac "------------------- ev_double --------------------".
idtac " ".
idtac "#> ev_double".
idtac "Possible points: 1".
check_type @ev_double ((forall n : nat, even (double n))).
idtac "Assumptions:".
Abort.
Print Assumptions ev_double.
Goal True.
idtac " ".
idtac "------------------- inversion_practice --------------------".
idtac " ".
idtac "#> SSSSev__even".
idtac "Possible points: 1".
check_type @SSSSev__even ((forall n : nat, even (S (S (S (S n)))) -> even n)).
idtac "Assumptions:".
Abort.
Print Assumptions SSSSev__even.
Goal True.
idtac " ".
idtac "------------------- even5_nonsense --------------------".
idtac " ".
idtac "#> even5_nonsense".
idtac "Possible points: 1".
check_type @even5_nonsense ((even 5 -> 2 + 2 = 9)).
idtac "Assumptions:".
Abort.
Print Assumptions even5_nonsense.
Goal True.
idtac " ".
idtac "------------------- ev_sum --------------------".
idtac " ".
idtac "#> ev_sum".
idtac "Possible points: 2".
check_type @ev_sum ((forall n m : nat, even n -> even m -> even (n + m))).
idtac "Assumptions:".
Abort.
Print Assumptions ev_sum.
Goal True.
idtac " ".
idtac "------------------- ev_ev__ev --------------------".
idtac " ".
idtac "#> ev_ev__ev".
idtac "Advanced".
idtac "Possible points: 3".
check_type @ev_ev__ev ((forall n m : nat, even (n + m) -> even n -> even m)).
idtac "Assumptions:".
Abort.
Print Assumptions ev_ev__ev.
Goal True.
idtac " ".
idtac "------------------- R_provability --------------------".
idtac " ".
idtac "#> Manually graded: R.R_provability".
idtac "Possible points: 3".
print_manual_grade R.manual_grade_for_R_provability.
idtac " ".
idtac "------------------- subsequence --------------------".
idtac " ".
idtac "#> subseq_refl".
idtac "Advanced".
idtac "Possible points: 1".
check_type @subseq_refl ((forall l : list nat, subseq l l)).
idtac "Assumptions:".
Abort.
Print Assumptions subseq_refl.
Goal True.
idtac " ".
idtac "#> subseq_app".
idtac "Advanced".
idtac "Possible points: 1".
check_type @subseq_app (
(forall l1 l2 l3 : list nat, subseq l1 l2 -> subseq l1 (l2 ++ l3))).
idtac "Assumptions:".
Abort.
Print Assumptions subseq_app.
Goal True.
idtac " ".
idtac "------------------- exp_match_ex1 --------------------".
idtac " ".
idtac "#> empty_is_empty".
idtac "Possible points: 1".
check_type @empty_is_empty ((forall (T : Type) (s : list T), ~ (s =~ @EmptySet T))).
idtac "Assumptions:".
Abort.
Print Assumptions empty_is_empty.
Goal True.
idtac " ".
idtac "#> MUnion'".
idtac "Possible points: 1".
check_type @MUnion' (
(forall (T : Type) (s : list T) (re1 re2 : @reg_exp T),
s =~ re1 \/ s =~ re2 -> s =~ @Union T re1 re2)).
idtac "Assumptions:".
Abort.
Print Assumptions MUnion'.
Goal True.
idtac " ".
idtac "#> MStar'".
idtac "Possible points: 1".
check_type @MStar' (
(forall (T : Type) (ss : list (list T)) (re : @reg_exp T),
(forall s : list T, @In (list T) s ss -> s =~ re) ->
@fold (list T) (list T) (@app T) ss [ ] =~ @Star T re)).
idtac "Assumptions:".
Abort.
Print Assumptions MStar'.
Goal True.
idtac " ".
idtac "------------------- re_not_empty --------------------".
idtac " ".
idtac "#> re_not_empty".
idtac "Possible points: 2".
check_type @re_not_empty ((forall T : Type, @reg_exp T -> bool)).
idtac "Assumptions:".
Abort.
Print Assumptions re_not_empty.
Goal True.
idtac " ".
idtac "#> re_not_empty_correct".
idtac "Possible points: 2".
check_type @re_not_empty_correct (
(forall (T : Type) (re : @reg_exp T),
(exists s : list T, s =~ re) <-> @re_not_empty T re = true)).
idtac "Assumptions:".
Abort.
Print Assumptions re_not_empty_correct.
Goal True.
idtac " ".
idtac "------------------- pumping --------------------".
idtac " ".
idtac "#> Pumping.pumping".
idtac "Advanced".
idtac "Possible points: 5".
check_type @Pumping.pumping (
(forall (T : Type) (re : @reg_exp T) (s : list T),
s =~ re ->
@Pumping.pumping_constant T re <= @length T s ->
exists s1 s2 s3 : list T,
s = s1 ++ s2 ++ s3 /\
s2 <> [ ] /\ (forall m : nat, s1 ++ @Pumping.napp T m s2 ++ s3 =~ re))).
idtac "Assumptions:".
Abort.
Print Assumptions Pumping.pumping.
Goal True.
idtac " ".
idtac "------------------- reflect_iff --------------------".
idtac " ".
idtac "#> reflect_iff".
idtac "Possible points: 2".
check_type @reflect_iff ((forall (P : Prop) (b : bool), reflect P b -> P <-> b = true)).
idtac "Assumptions:".
Abort.
Print Assumptions reflect_iff.
Goal True.
idtac " ".
idtac "------------------- eqbP_practice --------------------".
idtac " ".
idtac "#> eqbP_practice".
idtac "Possible points: 3".
check_type @eqbP_practice (
(forall (n : nat) (l : list nat), count n l = 0 -> ~ @In nat n l)).
idtac "Assumptions:".
Abort.
Print Assumptions eqbP_practice.
Goal True.
idtac " ".
idtac "------------------- nostutter_defn --------------------".
idtac " ".
idtac "#> Manually graded: nostutter".
idtac "Possible points: 3".
print_manual_grade manual_grade_for_nostutter.
idtac " ".
idtac "------------------- filter_challenge --------------------".
idtac " ".
idtac "#> Manually graded: filter_challenge".
idtac "Advanced".
idtac "Possible points: 4".
print_manual_grade manual_grade_for_filter_challenge.
idtac " ".
idtac " ".
idtac "Max points - standard: 23".
idtac "Max points - advanced: 37".
idtac "".
idtac "********** Summary **********".
idtac "".
idtac "********** Standard **********".
idtac "---------- ev_double ---------".
Print Assumptions ev_double.
idtac "---------- SSSSev__even ---------".
Print Assumptions SSSSev__even.
idtac "---------- even5_nonsense ---------".
Print Assumptions even5_nonsense.
idtac "---------- ev_sum ---------".
Print Assumptions ev_sum.
idtac "---------- R_provability ---------".
idtac "MANUAL".
idtac "---------- empty_is_empty ---------".
Print Assumptions empty_is_empty.
idtac "---------- MUnion' ---------".
Print Assumptions MUnion'.
idtac "---------- MStar' ---------".
Print Assumptions MStar'.
idtac "---------- re_not_empty ---------".
Print Assumptions re_not_empty.
idtac "---------- re_not_empty_correct ---------".
Print Assumptions re_not_empty_correct.
idtac "---------- reflect_iff ---------".
Print Assumptions reflect_iff.
idtac "---------- eqbP_practice ---------".
Print Assumptions eqbP_practice.
idtac "---------- nostutter ---------".
idtac "MANUAL".
idtac "".
idtac "********** Advanced **********".
idtac "---------- ev_ev__ev ---------".
Print Assumptions ev_ev__ev.
idtac "---------- subseq_refl ---------".
Print Assumptions subseq_refl.
idtac "---------- subseq_app ---------".
Print Assumptions subseq_app.
idtac "---------- Pumping.pumping ---------".
Print Assumptions Pumping.pumping.
idtac "---------- filter_challenge ---------".
idtac "MANUAL".
Abort.
(* Wed Jan 9 12:02:15 EST 2019 *)

680
Induction.v Normal file
View file

@ -0,0 +1,680 @@
(** * Induction: Proof by Induction *)
(** Before getting started, we need to import all of our
definitions from the previous chapter: *)
From LF Require Export Basics.
(** For the [Require Export] to work, Coq needs to be able to
find a compiled version of [Basics.v], called [Basics.vo], in a directory
associated with the prefix [LF]. This file is analogous to the [.class]
files compiled from [.java] source files and the [.o] files compiled from
[.c] files.
First create a file named [_CoqProject] containing the following line
(if you obtained the whole volume "Logical Foundations" as a single
archive, a [_CoqProject] should already exist and you can skip this step):
[-Q . LF]
This maps the current directory ("[.]", which contains [Basics.v],
[Induction.v], etc.) to the prefix (or "logical directory") "[LF]".
PG and CoqIDE read [_CoqProject] automatically, so they know to where to
look for the file [Basics.vo] corresponding to the library [LF.Basics].
Once [_CoqProject] is thus created, there are various ways to build
[Basics.vo]:
- In Proof General: The compilation can be made to happen automatically
when you submit the [Require] line above to PG, by setting the emacs
variable [coq-compile-before-require] to [t].
- In CoqIDE: Open [Basics.v]; then, in the "Compile" menu, click
on "Compile Buffer".
- From the command line: Generate a [Makefile] using the [coq_makefile]
utility, that comes installed with Coq (if you obtained the whole
volume as a single archive, a [Makefile] should already exist
and you can skip this step):
[coq_makefile -f _CoqProject *.v -o Makefile]
Note: You should rerun that command whenever you add or remove Coq files
to the directory.
Then you can compile [Basics.v] by running [make] with the corresponding
[.vo] file as a target:
[make Basics.vo]
All files in the directory can be compiled by giving no arguments:
[make]
Under the hood, [make] uses the Coq compiler, [coqc]. You can also
run [coqc] directly:
[coqc -Q . LF Basics.v]
But [make] also calculates dependencies between source files to compile
them in the right order, so [make] should generally be prefered over
explicit [coqc].
If you have trouble (e.g., if you get complaints about missing
identifiers later in the file), it may be because the "load path"
for Coq is not set up correctly. The [Print LoadPath.] command
may be helpful in sorting out such issues.
In particular, if you see a message like
[Compiled library Foo makes inconsistent assumptions over
library Bar]
check whether you have multiple installations of Coq on your machine.
It may be that commands (like [coqc]) that you execute in a terminal
window are getting a different version of Coq than commands executed by
Proof General or CoqIDE.
- Another common reason is that the library [Bar] was modified and
recompiled without also recompiling [Foo] which depends on it. Recompile
[Foo], or everything if too many files are affected. (Using the third
solution above: [make clean; make].)
One more tip for CoqIDE users: If you see messages like [Error:
Unable to locate library Basics], a likely reason is
inconsistencies between compiling things _within CoqIDE_ vs _using
[coqc] from the command line_. This typically happens when there
are two incompatible versions of [coqc] installed on your
system (one associated with CoqIDE, and one associated with [coqc]
from the terminal). The workaround for this situation is
compiling using CoqIDE only (i.e. choosing "make" from the menu),
and avoiding using [coqc] directly at all. *)
(* ################################################################# *)
(** * Proof by Induction *)
(** We proved in the last chapter that [0] is a neutral element
for [+] on the left, using an easy argument based on
simplification. We also observed that proving the fact that it is
also a neutral element on the _right_... *)
Theorem plus_n_O_firsttry : forall n:nat,
n = n + 0.
(** ... can't be done in the same simple way. Just applying
[reflexivity] doesn't work, since the [n] in [n + 0] is an arbitrary
unknown number, so the [match] in the definition of [+] can't be
simplified. *)
Proof.
intros n.
simpl. (* Does nothing! *)
Abort.
(** And reasoning by cases using [destruct n] doesn't get us much
further: the branch of the case analysis where we assume [n = 0]
goes through fine, but in the branch where [n = S n'] for some [n'] we
get stuck in exactly the same way. *)
Theorem plus_n_O_secondtry : forall n:nat,
n = n + 0.
Proof.
intros n. destruct n as [| n'] eqn:E.
- (* n = 0 *)
reflexivity. (* so far so good... *)
- (* n = S n' *)
simpl. (* ...but here we are stuck again *)
Abort.
(** We could use [destruct n'] to get one step further, but,
since [n] can be arbitrarily large, if we just go on like this
we'll never finish. *)
(** To prove interesting facts about numbers, lists, and other
inductively defined sets, we usually need a more powerful
reasoning principle: _induction_.
Recall (from high school, a discrete math course, etc.) the
_principle of induction over natural numbers_: If [P(n)] is some
proposition involving a natural number [n] and we want to show
that [P] holds for all numbers [n], we can reason like this:
- show that [P(O)] holds;
- show that, for any [n'], if [P(n')] holds, then so does
[P(S n')];
- conclude that [P(n)] holds for all [n].
In Coq, the steps are the same: we begin with the goal of proving
[P(n)] for all [n] and break it down (by applying the [induction]
tactic) into two separate subgoals: one where we must show [P(O)]
and another where we must show [P(n') -> P(S n')]. Here's how
this works for the theorem at hand: *)
Theorem plus_n_O : forall n:nat, n = n + 0.
Proof.
intros n. induction n as [| n' IHn'].
- (* n = 0 *) reflexivity.
- (* n = S n' *) simpl. rewrite <- IHn'. reflexivity. Qed.
(** Like [destruct], the [induction] tactic takes an [as...]
clause that specifies the names of the variables to be introduced
in the subgoals. Since there are two subgoals, the [as...] clause
has two parts, separated by [|]. (Strictly speaking, we can omit
the [as...] clause and Coq will choose names for us. In practice,
this is a bad idea, as Coq's automatic choices tend to be
confusing.)
In the first subgoal, [n] is replaced by [0]. No new variables
are introduced (so the first part of the [as...] is empty), and
the goal becomes [0 = 0 + 0], which follows by simplification.
In the second subgoal, [n] is replaced by [S n'], and the
assumption [n' + 0 = n'] is added to the context with the name
[IHn'] (i.e., the Induction Hypothesis for [n']). These two names
are specified in the second part of the [as...] clause. The goal
in this case becomes [S n' = (S n') + 0], which simplifies to
[S n' = S (n' + 0)], which in turn follows from [IHn']. *)
Theorem minus_diag : forall n,
minus n n = 0.
Proof.
(* WORKED IN CLASS *)
intros n. induction n as [| n' IHn'].
- (* n = 0 *)
simpl. reflexivity.
- (* n = S n' *)
simpl. rewrite -> IHn'. reflexivity. Qed.
(** (The use of the [intros] tactic in these proofs is actually
redundant. When applied to a goal that contains quantified
variables, the [induction] tactic will automatically move them
into the context as needed.) *)
(** **** Exercise: 2 stars, standard, recommended (basic_induction)
Prove the following using induction. You might need previously
proven results. *)
Theorem mult_0_r : forall n:nat,
n * 0 = 0.
Proof.
(* FILL IN HERE *) Admitted.
Theorem plus_n_Sm : forall n m : nat,
S (n + m) = n + (S m).
Proof.
(* FILL IN HERE *) Admitted.
Theorem plus_comm : forall n m : nat,
n + m = m + n.
Proof.
(* FILL IN HERE *) Admitted.
Theorem plus_assoc : forall n m p : nat,
n + (m + p) = (n + m) + p.
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** **** Exercise: 2 stars, standard (double_plus)
Consider the following function, which doubles its argument: *)
Fixpoint double (n:nat) :=
match n with
| O => O
| S n' => S (S (double n'))
end.
(** Use induction to prove this simple fact about [double]: *)
Lemma double_plus : forall n, double n = n + n .
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** **** Exercise: 2 stars, standard, optional (evenb_S)
One inconvenient aspect of our definition of [evenb n] is the
recursive call on [n - 2]. This makes proofs about [evenb n]
harder when done by induction on [n], since we may need an
induction hypothesis about [n - 2]. The following lemma gives an
alternative characterization of [evenb (S n)] that works better
with induction: *)
Theorem evenb_S : forall n : nat,
evenb (S n) = negb (evenb n).
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** **** Exercise: 1 star, standard (destruct_induction)
Briefly explain the difference between the tactics [destruct]
and [induction].
(* FILL IN HERE *)
*)
(* Do not modify the following line: *)
Definition manual_grade_for_destruct_induction : option (nat*string) := None.
(** [] *)
(* ################################################################# *)
(** * Proofs Within Proofs *)
(** In Coq, as in informal mathematics, large proofs are often
broken into a sequence of theorems, with later proofs referring to
earlier theorems. But sometimes a proof will require some
miscellaneous fact that is too trivial and of too little general
interest to bother giving it its own top-level name. In such
cases, it is convenient to be able to simply state and prove the
needed "sub-theorem" right at the point where it is used. The
[assert] tactic allows us to do this. For example, our earlier
proof of the [mult_0_plus] theorem referred to a previous theorem
named [plus_O_n]. We could instead use [assert] to state and
prove [plus_O_n] in-line: *)
Theorem mult_0_plus' : forall n m : nat,
(0 + n) * m = n * m.
Proof.
intros n m.
assert (H: 0 + n = n). { reflexivity. }
rewrite -> H.
reflexivity. Qed.
(** The [assert] tactic introduces two sub-goals. The first is
the assertion itself; by prefixing it with [H:] we name the
assertion [H]. (We can also name the assertion with [as] just as
we did above with [destruct] and [induction], i.e., [assert (0 + n
= n) as H].) Note that we surround the proof of this assertion
with curly braces [{ ... }], both for readability and so that,
when using Coq interactively, we can see more easily when we have
finished this sub-proof. The second goal is the same as the one
at the point where we invoke [assert] except that, in the context,
we now have the assumption [H] that [0 + n = n]. That is,
[assert] generates one subgoal where we must prove the asserted
fact and a second subgoal where we can use the asserted fact to
make progress on whatever we were trying to prove in the first
place. *)
(** Another example of [assert]... *)
(** For example, suppose we want to prove that [(n + m) + (p + q)
= (m + n) + (p + q)]. The only difference between the two sides of
the [=] is that the arguments [m] and [n] to the first inner [+]
are swapped, so it seems we should be able to use the
commutativity of addition ([plus_comm]) to rewrite one into the
other. However, the [rewrite] tactic is not very smart about
_where_ it applies the rewrite. There are three uses of [+] here,
and it turns out that doing [rewrite -> plus_comm] will affect
only the _outer_ one... *)
Theorem plus_rearrange_firsttry : forall n m p q : nat,
(n + m) + (p + q) = (m + n) + (p + q).
Proof.
intros n m p q.
(* We just need to swap (n + m) for (m + n)... seems
like plus_comm should do the trick! *)
rewrite -> plus_comm.
(* Doesn't work...Coq rewrites the wrong plus! *)
Abort.
(** To use [plus_comm] at the point where we need it, we can introduce
a local lemma stating that [n + m = m + n] (for the particular [m]
and [n] that we are talking about here), prove this lemma using
[plus_comm], and then use it to do the desired rewrite. *)
Theorem plus_rearrange : forall n m p q : nat,
(n + m) + (p + q) = (m + n) + (p + q).
Proof.
intros n m p q.
assert (H: n + m = m + n).
{ rewrite -> plus_comm. reflexivity. }
rewrite -> H. reflexivity. Qed.
(* ################################################################# *)
(** * Formal vs. Informal Proof *)
(** "_Informal proofs are algorithms; formal proofs are code_." *)
(** What constitutes a successful proof of a mathematical claim?
The question has challenged philosophers for millennia, but a
rough and ready definition could be this: A proof of a
mathematical proposition [P] is a written (or spoken) text that
instills in the reader or hearer the certainty that [P] is true --
an unassailable argument for the truth of [P]. That is, a proof
is an act of communication.
Acts of communication may involve different sorts of readers. On
one hand, the "reader" can be a program like Coq, in which case
the "belief" that is instilled is that [P] can be mechanically
derived from a certain set of formal logical rules, and the proof
is a recipe that guides the program in checking this fact. Such
recipes are _formal_ proofs.
Alternatively, the reader can be a human being, in which case the
proof will be written in English or some other natural language,
and will thus necessarily be _informal_. Here, the criteria for
success are less clearly specified. A "valid" proof is one that
makes the reader believe [P]. But the same proof may be read by
many different readers, some of whom may be convinced by a
particular way of phrasing the argument, while others may not be.
Some readers may be particularly pedantic, inexperienced, or just
plain thick-headed; the only way to convince them will be to make
the argument in painstaking detail. But other readers, more
familiar in the area, may find all this detail so overwhelming
that they lose the overall thread; all they want is to be told the
main ideas, since it is easier for them to fill in the details for
themselves than to wade through a written presentation of them.
Ultimately, there is no universal standard, because there is no
single way of writing an informal proof that is guaranteed to
convince every conceivable reader.
In practice, however, mathematicians have developed a rich set of
conventions and idioms for writing about complex mathematical
objects that -- at least within a certain community -- make
communication fairly reliable. The conventions of this stylized
form of communication give a fairly clear standard for judging
proofs good or bad.
Because we are using Coq in this course, we will be working
heavily with formal proofs. But this doesn't mean we can
completely forget about informal ones! Formal proofs are useful
in many ways, but they are _not_ very efficient ways of
communicating ideas between human beings. *)
(** For example, here is a proof that addition is associative: *)
Theorem plus_assoc' : forall n m p : nat,
n + (m + p) = (n + m) + p.
Proof. intros n m p. induction n as [| n' IHn']. reflexivity.
simpl. rewrite -> IHn'. reflexivity. Qed.
(** Coq is perfectly happy with this. For a human, however, it
is difficult to make much sense of it. We can use comments and
bullets to show the structure a little more clearly... *)
Theorem plus_assoc'' : forall n m p : nat,
n + (m + p) = (n + m) + p.
Proof.
intros n m p. induction n as [| n' IHn'].
- (* n = 0 *)
reflexivity.
- (* n = S n' *)
simpl. rewrite -> IHn'. reflexivity. Qed.
(** ... and if you're used to Coq you may be able to step
through the tactics one after the other in your mind and imagine
the state of the context and goal stack at each point, but if the
proof were even a little bit more complicated this would be next
to impossible.
A (pedantic) mathematician might write the proof something like
this: *)
(** - _Theorem_: For any [n], [m] and [p],
n + (m + p) = (n + m) + p.
_Proof_: By induction on [n].
- First, suppose [n = 0]. We must show
0 + (m + p) = (0 + m) + p.
This follows directly from the definition of [+].
- Next, suppose [n = S n'], where
n' + (m + p) = (n' + m) + p.
We must show
(S n') + (m + p) = ((S n') + m) + p.
By the definition of [+], this follows from
S (n' + (m + p)) = S ((n' + m) + p),
which is immediate from the induction hypothesis. _Qed_. *)
(** The overall form of the proof is basically similar, and of
course this is no accident: Coq has been designed so that its
[induction] tactic generates the same sub-goals, in the same
order, as the bullet points that a mathematician would write. But
there are significant differences of detail: the formal proof is
much more explicit in some ways (e.g., the use of [reflexivity])
but much less explicit in others (in particular, the "proof state"
at any given point in the Coq proof is completely implicit,
whereas the informal proof reminds the reader several times where
things stand). *)
(** **** Exercise: 2 stars, advanced, recommended (plus_comm_informal)
Translate your solution for [plus_comm] into an informal proof:
Theorem: Addition is commutative.
Proof: (* FILL IN HERE *)
*)
(* Do not modify the following line: *)
Definition manual_grade_for_plus_comm_informal : option (nat*string) := None.
(** [] *)
(** **** Exercise: 2 stars, standard, optional (eqb_refl_informal)
Write an informal proof of the following theorem, using the
informal proof of [plus_assoc] as a model. Don't just
paraphrase the Coq tactics into English!
Theorem: [true = n =? n] for any [n].
Proof: (* FILL IN HERE *)
[] *)
(* ################################################################# *)
(** * More Exercises *)
(** **** Exercise: 3 stars, standard, recommended (mult_comm)
Use [assert] to help prove this theorem. You shouldn't need to
use induction on [plus_swap]. *)
Theorem plus_swap : forall n m p : nat,
n + (m + p) = m + (n + p).
Proof.
(* FILL IN HERE *) Admitted.
(** Now prove commutativity of multiplication. (You will probably
need to define and prove a separate subsidiary theorem to be used
in the proof of this one. You may find that [plus_swap] comes in
handy.) *)
Theorem mult_comm : forall m n : nat,
m * n = n * m.
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** **** Exercise: 3 stars, standard, optional (more_exercises)
Take a piece of paper. For each of the following theorems, first
_think_ about whether (a) it can be proved using only
simplification and rewriting, (b) it also requires case
analysis ([destruct]), or (c) it also requires induction. Write
down your prediction. Then fill in the proof. (There is no need
to turn in your piece of paper; this is just to encourage you to
reflect before you hack!) *)
Check leb.
Theorem leb_refl : forall n:nat,
true = (n <=? n).
Proof.
(* FILL IN HERE *) Admitted.
Theorem zero_nbeq_S : forall n:nat,
0 =? (S n) = false.
Proof.
(* FILL IN HERE *) Admitted.
Theorem andb_false_r : forall b : bool,
andb b false = false.
Proof.
(* FILL IN HERE *) Admitted.
Theorem plus_ble_compat_l : forall n m p : nat,
n <=? m = true -> (p + n) <=? (p + m) = true.
Proof.
(* FILL IN HERE *) Admitted.
Theorem S_nbeq_0 : forall n:nat,
(S n) =? 0 = false.
Proof.
(* FILL IN HERE *) Admitted.
Theorem mult_1_l : forall n:nat, 1 * n = n.
Proof.
(* FILL IN HERE *) Admitted.
Theorem all3_spec : forall b c : bool,
orb
(andb b c)
(orb (negb b)
(negb c))
= true.
Proof.
(* FILL IN HERE *) Admitted.
Theorem mult_plus_distr_r : forall n m p : nat,
(n + m) * p = (n * p) + (m * p).
Proof.
(* FILL IN HERE *) Admitted.
Theorem mult_assoc : forall n m p : nat,
n * (m * p) = (n * m) * p.
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** **** Exercise: 2 stars, standard, optional (eqb_refl)
Prove the following theorem. (Putting the [true] on the left-hand
side of the equality may look odd, but this is how the theorem is
stated in the Coq standard library, so we follow suit. Rewriting
works equally well in either direction, so we will have no problem
using the theorem no matter which way we state it.) *)
Theorem eqb_refl : forall n : nat,
true = (n =? n).
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** **** Exercise: 2 stars, standard, optional (plus_swap')
The [replace] tactic allows you to specify a particular subterm to
rewrite and what you want it rewritten to: [replace (t) with (u)]
replaces (all copies of) expression [t] in the goal by expression
[u], and generates [t = u] as an additional subgoal. This is often
useful when a plain [rewrite] acts on the wrong part of the goal.
Use the [replace] tactic to do a proof of [plus_swap'], just like
[plus_swap] but without needing [assert (n + m = m + n)]. *)
Theorem plus_swap' : forall n m p : nat,
n + (m + p) = m + (n + p).
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** **** Exercise: 3 stars, standard, recommended (binary_commute)
Recall the [incr] and [bin_to_nat] functions that you
wrote for the [binary] exercise in the [Basics] chapter. Prove
that the following diagram commutes:
incr
bin ----------------------> bin
| |
bin_to_nat | | bin_to_nat
| |
v v
nat ----------------------> nat
S
That is, incrementing a binary number and then converting it to
a (unary) natural number yields the same result as first converting
it to a natural number and then incrementing.
Name your theorem [bin_to_nat_pres_incr] ("pres" for "preserves").
Before you start working on this exercise, copy the definitions
from your solution to the [binary] exercise here so that this file
can be graded on its own. If you want to change your original
definitions to make the property easier to prove, feel free to
do so! *)
(* FILL IN HERE *)
(* Do not modify the following line: *)
Definition manual_grade_for_binary_commute : option (nat*string) := None.
(** [] *)
(** **** Exercise: 5 stars, advanced (binary_inverse)
This is a further continuation of the previous exercises about
binary numbers. You may find you need to go back and change your
earlier definitions to get things to work here.
(a) First, write a function to convert natural numbers to binary
numbers. *)
Fixpoint nat_to_bin (n:nat) : bin
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
(** Prove that, if we start with any [nat], convert it to binary, and
convert it back, we get the same [nat] we started with. (Hint: If
your definition of [nat_to_bin] involved any extra functions, you
may need to prove a subsidiary lemma showing how such functions
relate to [nat_to_bin].) *)
Theorem nat_bin_nat : forall n, bin_to_nat (nat_to_bin n) = n.
Proof.
(* FILL IN HERE *) Admitted.
(* Do not modify the following line: *)
Definition manual_grade_for_binary_inverse_a : option (nat*string) := None.
(** (b) One might naturally expect that we should also prove the
opposite direction -- that starting with a binary number,
converting to a natural, and then back to binary should yield
the same number we started with. However, this is not the
case! Explain (in a comment) what the problem is. *)
(* FILL IN HERE *)
(* Do not modify the following line: *)
Definition manual_grade_for_binary_inverse_b : option (nat*string) := None.
(** (c) Define a normalization function -- i.e., a function
[normalize] going directly from [bin] to [bin] (i.e., _not_ by
converting to [nat] and back) such that, for any binary number
[b], converting [b] to a natural and then back to binary yields
[(normalize b)]. Prove it. (Warning: This part is a bit
tricky -- you may end up defining several auxiliary lemmas.
One good way to find out what you need is to start by trying
to prove the main statement, see where you get stuck, and see
if you can find a lemma -- perhaps requiring its own inductive
proof -- that will allow the main proof to make progress.) Don't
define thi using nat_to_bin and bin_to_nat! *)
(* FILL IN HERE *)
(* Do not modify the following line: *)
Definition manual_grade_for_binary_inverse_c : option (nat*string) := None.
(** [] *)
(* Wed Jan 9 12:02:44 EST 2019 *)

180
InductionTest.v Normal file
View file

@ -0,0 +1,180 @@
Set Warnings "-notation-overridden,-parsing".
From Coq Require Export String.
From LF Require Import Induction.
Parameter MISSING: Type.
Module Check.
Ltac check_type A B :=
match type of A with
| context[MISSING] => idtac "Missing:" A
| ?T => first [unify T B; idtac "Type: ok" | idtac "Type: wrong - should be (" B ")"]
end.
Ltac print_manual_grade A :=
match eval compute in A with
| Some (_ ?S ?C) =>
idtac "Score:" S;
match eval compute in C with
| ""%string => idtac "Comment: None"
| _ => idtac "Comment:" C
end
| None =>
idtac "Score: Ungraded";
idtac "Comment: None"
end.
End Check.
From LF Require Import Induction.
Import Check.
Goal True.
idtac "------------------- basic_induction --------------------".
idtac " ".
idtac "#> mult_0_r".
idtac "Possible points: 0.5".
check_type @mult_0_r ((forall n : nat, n * 0 = 0)).
idtac "Assumptions:".
Abort.
Print Assumptions mult_0_r.
Goal True.
idtac " ".
idtac "#> plus_n_Sm".
idtac "Possible points: 0.5".
check_type @plus_n_Sm ((forall n m : nat, S (n + m) = n + S m)).
idtac "Assumptions:".
Abort.
Print Assumptions plus_n_Sm.
Goal True.
idtac " ".
idtac "#> plus_comm".
idtac "Possible points: 0.5".
check_type @plus_comm ((forall n m : nat, n + m = m + n)).
idtac "Assumptions:".
Abort.
Print Assumptions plus_comm.
Goal True.
idtac " ".
idtac "#> plus_assoc".
idtac "Possible points: 0.5".
check_type @plus_assoc ((forall n m p : nat, n + (m + p) = n + m + p)).
idtac "Assumptions:".
Abort.
Print Assumptions plus_assoc.
Goal True.
idtac " ".
idtac "------------------- double_plus --------------------".
idtac " ".
idtac "#> double_plus".
idtac "Possible points: 2".
check_type @double_plus ((forall n : nat, double n = n + n)).
idtac "Assumptions:".
Abort.
Print Assumptions double_plus.
Goal True.
idtac " ".
idtac "------------------- destruct_induction --------------------".
idtac " ".
idtac "#> Manually graded: destruct_induction".
idtac "Possible points: 1".
print_manual_grade manual_grade_for_destruct_induction.
idtac " ".
idtac "------------------- plus_comm_informal --------------------".
idtac " ".
idtac "#> Manually graded: plus_comm_informal".
idtac "Advanced".
idtac "Possible points: 2".
print_manual_grade manual_grade_for_plus_comm_informal.
idtac " ".
idtac "------------------- mult_comm --------------------".
idtac " ".
idtac "#> mult_comm".
idtac "Possible points: 3".
check_type @mult_comm ((forall m n : nat, m * n = n * m)).
idtac "Assumptions:".
Abort.
Print Assumptions mult_comm.
Goal True.
idtac " ".
idtac "------------------- binary_commute --------------------".
idtac " ".
idtac "#> Manually graded: binary_commute".
idtac "Possible points: 3".
print_manual_grade manual_grade_for_binary_commute.
idtac " ".
idtac "------------------- binary_inverse --------------------".
idtac " ".
idtac "#> Manually graded: binary_inverse_a".
idtac "Advanced".
idtac "Possible points: 2".
print_manual_grade manual_grade_for_binary_inverse_a.
idtac " ".
idtac "#> Manually graded: binary_inverse_b".
idtac "Advanced".
idtac "Possible points: 1".
print_manual_grade manual_grade_for_binary_inverse_b.
idtac " ".
idtac "#> Manually graded: binary_inverse_c".
idtac "Advanced".
idtac "Possible points: 2".
print_manual_grade manual_grade_for_binary_inverse_c.
idtac " ".
idtac " ".
idtac "Max points - standard: 11".
idtac "Max points - advanced: 18".
idtac "".
idtac "********** Summary **********".
idtac "".
idtac "********** Standard **********".
idtac "---------- mult_0_r ---------".
Print Assumptions mult_0_r.
idtac "---------- plus_n_Sm ---------".
Print Assumptions plus_n_Sm.
idtac "---------- plus_comm ---------".
Print Assumptions plus_comm.
idtac "---------- plus_assoc ---------".
Print Assumptions plus_assoc.
idtac "---------- double_plus ---------".
Print Assumptions double_plus.
idtac "---------- destruct_induction ---------".
idtac "MANUAL".
idtac "---------- mult_comm ---------".
Print Assumptions mult_comm.
idtac "---------- binary_commute ---------".
idtac "MANUAL".
idtac "".
idtac "********** Advanced **********".
idtac "---------- plus_comm_informal ---------".
idtac "MANUAL".
idtac "---------- binary_inverse_a ---------".
idtac "MANUAL".
idtac "---------- binary_inverse_b ---------".
idtac "MANUAL".
idtac "---------- binary_inverse_c ---------".
idtac "MANUAL".
Abort.
(* Wed Jan 9 12:02:06 EST 2019 *)

1136
Lists.v Normal file

File diff suppressed because it is too large Load diff

476
ListsTest.v Normal file
View file

@ -0,0 +1,476 @@
Set Warnings "-notation-overridden,-parsing".
From Coq Require Export String.
From LF Require Import Lists.
Parameter MISSING: Type.
Module Check.
Ltac check_type A B :=
match type of A with
| context[MISSING] => idtac "Missing:" A
| ?T => first [unify T B; idtac "Type: ok" | idtac "Type: wrong - should be (" B ")"]
end.
Ltac print_manual_grade A :=
match eval compute in A with
| Some (_ ?S ?C) =>
idtac "Score:" S;
match eval compute in C with
| ""%string => idtac "Comment: None"
| _ => idtac "Comment:" C
end
| None =>
idtac "Score: Ungraded";
idtac "Comment: None"
end.
End Check.
From LF Require Import Lists.
Import Check.
Goal True.
idtac "------------------- snd_fst_is_swap --------------------".
idtac " ".
idtac "#> NatList.snd_fst_is_swap".
idtac "Possible points: 1".
check_type @NatList.snd_fst_is_swap (
(forall p : NatList.natprod,
NatList.pair (NatList.snd p) (NatList.fst p) = NatList.swap_pair p)).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.snd_fst_is_swap.
Goal True.
idtac " ".
idtac "------------------- list_funs --------------------".
idtac " ".
idtac "#> NatList.test_nonzeros".
idtac "Possible points: 0.5".
check_type @NatList.test_nonzeros (
(NatList.nonzeros
(NatList.cons 0
(NatList.cons 1
(NatList.cons 0
(NatList.cons 2
(NatList.cons 3 (NatList.cons 0 (NatList.cons 0 NatList.nil))))))) =
NatList.cons 1 (NatList.cons 2 (NatList.cons 3 NatList.nil)))).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.test_nonzeros.
Goal True.
idtac " ".
idtac "#> NatList.test_oddmembers".
idtac "Possible points: 0.5".
check_type @NatList.test_oddmembers (
(NatList.oddmembers
(NatList.cons 0
(NatList.cons 1
(NatList.cons 0
(NatList.cons 2
(NatList.cons 3 (NatList.cons 0 (NatList.cons 0 NatList.nil))))))) =
NatList.cons 1 (NatList.cons 3 NatList.nil))).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.test_oddmembers.
Goal True.
idtac " ".
idtac "#> NatList.test_countoddmembers2".
idtac "Possible points: 0.5".
check_type @NatList.test_countoddmembers2 (
(NatList.countoddmembers
(NatList.cons 0 (NatList.cons 2 (NatList.cons 4 NatList.nil))) = 0)).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.test_countoddmembers2.
Goal True.
idtac " ".
idtac "#> NatList.test_countoddmembers3".
idtac "Possible points: 0.5".
check_type @NatList.test_countoddmembers3 ((NatList.countoddmembers NatList.nil = 0)).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.test_countoddmembers3.
Goal True.
idtac " ".
idtac "------------------- alternate --------------------".
idtac " ".
idtac "#> NatList.test_alternate1".
idtac "Advanced".
idtac "Possible points: 1".
check_type @NatList.test_alternate1 (
(NatList.alternate
(NatList.cons 1 (NatList.cons 2 (NatList.cons 3 NatList.nil)))
(NatList.cons 4 (NatList.cons 5 (NatList.cons 6 NatList.nil))) =
NatList.cons 1
(NatList.cons 4
(NatList.cons 2
(NatList.cons 5 (NatList.cons 3 (NatList.cons 6 NatList.nil))))))).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.test_alternate1.
Goal True.
idtac " ".
idtac "#> NatList.test_alternate2".
idtac "Advanced".
idtac "Possible points: 1".
check_type @NatList.test_alternate2 (
(NatList.alternate (NatList.cons 1 NatList.nil)
(NatList.cons 4 (NatList.cons 5 (NatList.cons 6 NatList.nil))) =
NatList.cons 1
(NatList.cons 4 (NatList.cons 5 (NatList.cons 6 NatList.nil))))).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.test_alternate2.
Goal True.
idtac " ".
idtac "#> NatList.test_alternate4".
idtac "Advanced".
idtac "Possible points: 1".
check_type @NatList.test_alternate4 (
(NatList.alternate NatList.nil
(NatList.cons 20 (NatList.cons 30 NatList.nil)) =
NatList.cons 20 (NatList.cons 30 NatList.nil))).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.test_alternate4.
Goal True.
idtac " ".
idtac "------------------- bag_functions --------------------".
idtac " ".
idtac "#> NatList.test_count2".
idtac "Possible points: 0.5".
check_type @NatList.test_count2 (
(NatList.count 6
(NatList.cons 1
(NatList.cons 2
(NatList.cons 3
(NatList.cons 1 (NatList.cons 4 (NatList.cons 1 NatList.nil)))))) =
0)).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.test_count2.
Goal True.
idtac " ".
idtac "#> NatList.test_sum1".
idtac "Possible points: 0.5".
check_type @NatList.test_sum1 (
(NatList.count 1
(NatList.sum
(NatList.cons 1 (NatList.cons 2 (NatList.cons 3 NatList.nil)))
(NatList.cons 1 (NatList.cons 4 (NatList.cons 1 NatList.nil)))) = 3)).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.test_sum1.
Goal True.
idtac " ".
idtac "#> NatList.test_add1".
idtac "Possible points: 0.5".
check_type @NatList.test_add1 (
(NatList.count 1
(NatList.add 1
(NatList.cons 1 (NatList.cons 4 (NatList.cons 1 NatList.nil)))) = 3)).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.test_add1.
Goal True.
idtac " ".
idtac "#> NatList.test_add2".
idtac "Possible points: 0.5".
check_type @NatList.test_add2 (
(NatList.count 5
(NatList.add 1
(NatList.cons 1 (NatList.cons 4 (NatList.cons 1 NatList.nil)))) = 0)).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.test_add2.
Goal True.
idtac " ".
idtac "#> NatList.test_member1".
idtac "Possible points: 0.5".
check_type @NatList.test_member1 (
(NatList.member 1
(NatList.cons 1 (NatList.cons 4 (NatList.cons 1 NatList.nil))) = true)).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.test_member1.
Goal True.
idtac " ".
idtac "#> NatList.test_member2".
idtac "Possible points: 0.5".
check_type @NatList.test_member2 (
(NatList.member 2
(NatList.cons 1 (NatList.cons 4 (NatList.cons 1 NatList.nil))) = false)).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.test_member2.
Goal True.
idtac " ".
idtac "------------------- bag_theorem --------------------".
idtac " ".
idtac "#> Manually graded: NatList.bag_theorem".
idtac "Possible points: 2".
print_manual_grade NatList.manual_grade_for_bag_theorem.
idtac " ".
idtac "------------------- list_exercises --------------------".
idtac " ".
idtac "#> NatList.app_nil_r".
idtac "Possible points: 0.5".
check_type @NatList.app_nil_r (
(forall l : NatList.natlist, NatList.app l NatList.nil = l)).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.app_nil_r.
Goal True.
idtac " ".
idtac "#> NatList.rev_app_distr".
idtac "Possible points: 0.5".
check_type @NatList.rev_app_distr (
(forall l1 l2 : NatList.natlist,
NatList.rev (NatList.app l1 l2) =
NatList.app (NatList.rev l2) (NatList.rev l1))).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.rev_app_distr.
Goal True.
idtac " ".
idtac "#> NatList.rev_involutive".
idtac "Possible points: 0.5".
check_type @NatList.rev_involutive (
(forall l : NatList.natlist, NatList.rev (NatList.rev l) = l)).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.rev_involutive.
Goal True.
idtac " ".
idtac "#> NatList.app_assoc4".
idtac "Possible points: 0.5".
check_type @NatList.app_assoc4 (
(forall l1 l2 l3 l4 : NatList.natlist,
NatList.app l1 (NatList.app l2 (NatList.app l3 l4)) =
NatList.app (NatList.app (NatList.app l1 l2) l3) l4)).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.app_assoc4.
Goal True.
idtac " ".
idtac "#> NatList.nonzeros_app".
idtac "Possible points: 1".
check_type @NatList.nonzeros_app (
(forall l1 l2 : NatList.natlist,
NatList.nonzeros (NatList.app l1 l2) =
NatList.app (NatList.nonzeros l1) (NatList.nonzeros l2))).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.nonzeros_app.
Goal True.
idtac " ".
idtac "------------------- eqblist --------------------".
idtac " ".
idtac "#> NatList.eqblist_refl".
idtac "Possible points: 2".
check_type @NatList.eqblist_refl (
(forall l : NatList.natlist, true = NatList.eqblist l l)).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.eqblist_refl.
Goal True.
idtac " ".
idtac "------------------- count_member_nonzero --------------------".
idtac " ".
idtac "#> NatList.count_member_nonzero".
idtac "Possible points: 1".
check_type @NatList.count_member_nonzero (
(forall s : NatList.bag, (1 <=? NatList.count 1 (NatList.cons 1 s)) = true)).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.count_member_nonzero.
Goal True.
idtac " ".
idtac "------------------- remove_does_not_increase_count --------------------".
idtac " ".
idtac "#> NatList.remove_does_not_increase_count".
idtac "Advanced".
idtac "Possible points: 3".
check_type @NatList.remove_does_not_increase_count (
(forall s : NatList.bag,
(NatList.count 0 (NatList.remove_one 0 s) <=? NatList.count 0 s) = true)).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.remove_does_not_increase_count.
Goal True.
idtac " ".
idtac "------------------- rev_injective --------------------".
idtac " ".
idtac "#> Manually graded: NatList.rev_injective".
idtac "Advanced".
idtac "Possible points: 4".
print_manual_grade NatList.manual_grade_for_rev_injective.
idtac " ".
idtac "------------------- hd_error --------------------".
idtac " ".
idtac "#> NatList.hd_error".
idtac "Possible points: 2".
check_type @NatList.hd_error ((NatList.natlist -> NatList.natoption)).
idtac "Assumptions:".
Abort.
Print Assumptions NatList.hd_error.
Goal True.
idtac " ".
idtac "------------------- eqb_id_refl --------------------".
idtac " ".
idtac "#> eqb_id_refl".
idtac "Possible points: 1".
check_type @eqb_id_refl ((forall x : id, true = eqb_id x x)).
idtac "Assumptions:".
Abort.
Print Assumptions eqb_id_refl.
Goal True.
idtac " ".
idtac "------------------- update_eq --------------------".
idtac " ".
idtac "#> PartialMap.update_eq".
idtac "Possible points: 1".
check_type @PartialMap.update_eq (
(forall (d : PartialMap.partial_map) (x : id) (v : nat),
PartialMap.find x (PartialMap.update d x v) = NatList.Some v)).
idtac "Assumptions:".
Abort.
Print Assumptions PartialMap.update_eq.
Goal True.
idtac " ".
idtac "------------------- update_neq --------------------".
idtac " ".
idtac "#> PartialMap.update_neq".
idtac "Possible points: 1".
check_type @PartialMap.update_neq (
(forall (d : PartialMap.partial_map) (x y : id) (o : nat),
eqb_id x y = false ->
PartialMap.find x (PartialMap.update d y o) = PartialMap.find x d)).
idtac "Assumptions:".
Abort.
Print Assumptions PartialMap.update_neq.
Goal True.
idtac " ".
idtac "------------------- baz_num_elts --------------------".
idtac " ".
idtac "#> Manually graded: baz_num_elts".
idtac "Possible points: 2".
print_manual_grade manual_grade_for_baz_num_elts.
idtac " ".
idtac " ".
idtac "Max points - standard: 21".
idtac "Max points - advanced: 31".
idtac "".
idtac "********** Summary **********".
idtac "".
idtac "********** Standard **********".
idtac "---------- NatList.snd_fst_is_swap ---------".
Print Assumptions NatList.snd_fst_is_swap.
idtac "---------- NatList.test_nonzeros ---------".
Print Assumptions NatList.test_nonzeros.
idtac "---------- NatList.test_oddmembers ---------".
Print Assumptions NatList.test_oddmembers.
idtac "---------- NatList.test_countoddmembers2 ---------".
Print Assumptions NatList.test_countoddmembers2.
idtac "---------- NatList.test_countoddmembers3 ---------".
Print Assumptions NatList.test_countoddmembers3.
idtac "---------- NatList.test_count2 ---------".
Print Assumptions NatList.test_count2.
idtac "---------- NatList.test_sum1 ---------".
Print Assumptions NatList.test_sum1.
idtac "---------- NatList.test_add1 ---------".
Print Assumptions NatList.test_add1.
idtac "---------- NatList.test_add2 ---------".
Print Assumptions NatList.test_add2.
idtac "---------- NatList.test_member1 ---------".
Print Assumptions NatList.test_member1.
idtac "---------- NatList.test_member2 ---------".
Print Assumptions NatList.test_member2.
idtac "---------- bag_theorem ---------".
idtac "MANUAL".
idtac "---------- NatList.app_nil_r ---------".
Print Assumptions NatList.app_nil_r.
idtac "---------- NatList.rev_app_distr ---------".
Print Assumptions NatList.rev_app_distr.
idtac "---------- NatList.rev_involutive ---------".
Print Assumptions NatList.rev_involutive.
idtac "---------- NatList.app_assoc4 ---------".
Print Assumptions NatList.app_assoc4.
idtac "---------- NatList.nonzeros_app ---------".
Print Assumptions NatList.nonzeros_app.
idtac "---------- NatList.eqblist_refl ---------".
Print Assumptions NatList.eqblist_refl.
idtac "---------- NatList.count_member_nonzero ---------".
Print Assumptions NatList.count_member_nonzero.
idtac "---------- NatList.hd_error ---------".
Print Assumptions NatList.hd_error.
idtac "---------- eqb_id_refl ---------".
Print Assumptions eqb_id_refl.
idtac "---------- PartialMap.update_eq ---------".
Print Assumptions PartialMap.update_eq.
idtac "---------- PartialMap.update_neq ---------".
Print Assumptions PartialMap.update_neq.
idtac "---------- baz_num_elts ---------".
idtac "MANUAL".
idtac "".
idtac "********** Advanced **********".
idtac "---------- NatList.test_alternate1 ---------".
Print Assumptions NatList.test_alternate1.
idtac "---------- NatList.test_alternate2 ---------".
Print Assumptions NatList.test_alternate2.
idtac "---------- NatList.test_alternate4 ---------".
Print Assumptions NatList.test_alternate4.
idtac "---------- NatList.remove_does_not_increase_count ---------".
Print Assumptions NatList.remove_does_not_increase_count.
idtac "---------- rev_injective ---------".
idtac "MANUAL".
Abort.
(* Wed Jan 9 12:02:08 EST 2019 *)

1603
Logic.v Normal file

File diff suppressed because it is too large Load diff

393
LogicTest.v Normal file
View file

@ -0,0 +1,393 @@
Set Warnings "-notation-overridden,-parsing".
From Coq Require Export String.
From LF Require Import Logic.
Parameter MISSING: Type.
Module Check.
Ltac check_type A B :=
match type of A with
| context[MISSING] => idtac "Missing:" A
| ?T => first [unify T B; idtac "Type: ok" | idtac "Type: wrong - should be (" B ")"]
end.
Ltac print_manual_grade A :=
match eval compute in A with
| Some (_ ?S ?C) =>
idtac "Score:" S;
match eval compute in C with
| ""%string => idtac "Comment: None"
| _ => idtac "Comment:" C
end
| None =>
idtac "Score: Ungraded";
idtac "Comment: None"
end.
End Check.
From LF Require Import Logic.
Import Check.
Goal True.
idtac "------------------- and_exercise --------------------".
idtac " ".
idtac "#> and_exercise".
idtac "Possible points: 2".
check_type @and_exercise ((forall n m : nat, n + m = 0 -> n = 0 /\ m = 0)).
idtac "Assumptions:".
Abort.
Print Assumptions and_exercise.
Goal True.
idtac " ".
idtac "------------------- and_assoc --------------------".
idtac " ".
idtac "#> and_assoc".
idtac "Possible points: 2".
check_type @and_assoc ((forall P Q R : Prop, P /\ Q /\ R -> (P /\ Q) /\ R)).
idtac "Assumptions:".
Abort.
Print Assumptions and_assoc.
Goal True.
idtac " ".
idtac "------------------- mult_eq_0 --------------------".
idtac " ".
idtac "#> mult_eq_0".
idtac "Possible points: 1".
check_type @mult_eq_0 ((forall n m : nat, n * m = 0 -> n = 0 \/ m = 0)).
idtac "Assumptions:".
Abort.
Print Assumptions mult_eq_0.
Goal True.
idtac " ".
idtac "------------------- or_commut --------------------".
idtac " ".
idtac "#> or_commut".
idtac "Possible points: 1".
check_type @or_commut ((forall P Q : Prop, P \/ Q -> Q \/ P)).
idtac "Assumptions:".
Abort.
Print Assumptions or_commut.
Goal True.
idtac " ".
idtac "------------------- double_neg_inf --------------------".
idtac " ".
idtac "#> Manually graded: double_neg_inf".
idtac "Advanced".
idtac "Possible points: 2".
print_manual_grade manual_grade_for_double_neg_inf.
idtac " ".
idtac "------------------- contrapositive --------------------".
idtac " ".
idtac "#> contrapositive".
idtac "Possible points: 2".
check_type @contrapositive ((forall P Q : Prop, (P -> Q) -> ~ Q -> ~ P)).
idtac "Assumptions:".
Abort.
Print Assumptions contrapositive.
Goal True.
idtac " ".
idtac "------------------- not_both_true_and_false --------------------".
idtac " ".
idtac "#> not_both_true_and_false".
idtac "Possible points: 1".
check_type @not_both_true_and_false ((forall P : Prop, ~ (P /\ ~ P))).
idtac "Assumptions:".
Abort.
Print Assumptions not_both_true_and_false.
Goal True.
idtac " ".
idtac "------------------- informal_not_PNP --------------------".
idtac " ".
idtac "#> Manually graded: informal_not_PNP".
idtac "Advanced".
idtac "Possible points: 1".
print_manual_grade manual_grade_for_informal_not_PNP.
idtac " ".
idtac "------------------- or_distributes_over_and --------------------".
idtac " ".
idtac "#> or_distributes_over_and".
idtac "Possible points: 3".
check_type @or_distributes_over_and (
(forall P Q R : Prop, P \/ Q /\ R <-> (P \/ Q) /\ (P \/ R))).
idtac "Assumptions:".
Abort.
Print Assumptions or_distributes_over_and.
Goal True.
idtac " ".
idtac "------------------- dist_not_exists --------------------".
idtac " ".
idtac "#> dist_not_exists".
idtac "Possible points: 1".
check_type @dist_not_exists (
(forall (X : Type) (P : X -> Prop),
(forall x : X, P x) -> ~ (exists x : X, ~ P x))).
idtac "Assumptions:".
Abort.
Print Assumptions dist_not_exists.
Goal True.
idtac " ".
idtac "------------------- dist_exists_or --------------------".
idtac " ".
idtac "#> dist_exists_or".
idtac "Possible points: 2".
check_type @dist_exists_or (
(forall (X : Type) (P Q : X -> Prop),
(exists x : X, P x \/ Q x) <-> (exists x : X, P x) \/ (exists x : X, Q x))).
idtac "Assumptions:".
Abort.
Print Assumptions dist_exists_or.
Goal True.
idtac " ".
idtac "------------------- In_map_iff --------------------".
idtac " ".
idtac "#> In_map_iff".
idtac "Possible points: 2".
check_type @In_map_iff (
(forall (A B : Type) (f : A -> B) (l : list A) (y : B),
@In B y (@map A B f l) <-> (exists x : A, f x = y /\ @In A x l))).
idtac "Assumptions:".
Abort.
Print Assumptions In_map_iff.
Goal True.
idtac " ".
idtac "------------------- In_app_iff --------------------".
idtac " ".
idtac "#> In_app_iff".
idtac "Possible points: 2".
check_type @In_app_iff (
(forall (A : Type) (l l' : list A) (a : A),
@In A a (l ++ l') <-> @In A a l \/ @In A a l')).
idtac "Assumptions:".
Abort.
Print Assumptions In_app_iff.
Goal True.
idtac " ".
idtac "------------------- All --------------------".
idtac " ".
idtac "#> All".
idtac "Possible points: 3".
check_type @All ((forall T : Type, (T -> Prop) -> list T -> Prop)).
idtac "Assumptions:".
Abort.
Print Assumptions All.
Goal True.
idtac " ".
idtac "------------------- combine_odd_even --------------------".
idtac " ".
idtac "#> combine_odd_even".
idtac "Possible points: 3".
check_type @combine_odd_even (((nat -> Prop) -> (nat -> Prop) -> nat -> Prop)).
idtac "Assumptions:".
Abort.
Print Assumptions combine_odd_even.
Goal True.
idtac " ".
idtac "------------------- tr_rev_correct --------------------".
idtac " ".
idtac "#> tr_rev_correct".
idtac "Possible points: 4".
check_type @tr_rev_correct ((forall X : Type, @tr_rev X = @rev X)).
idtac "Assumptions:".
Abort.
Print Assumptions tr_rev_correct.
Goal True.
idtac " ".
idtac "------------------- evenb_double_conv --------------------".
idtac " ".
idtac "#> evenb_double_conv".
idtac "Possible points: 3".
check_type @evenb_double_conv (
(forall n : nat,
exists k : nat, n = (if evenb n then double k else S (double k)))).
idtac "Assumptions:".
Abort.
Print Assumptions evenb_double_conv.
Goal True.
idtac " ".
idtac "------------------- logical_connectives --------------------".
idtac " ".
idtac "#> andb_true_iff".
idtac "Possible points: 1".
check_type @andb_true_iff (
(forall b1 b2 : bool, b1 && b2 = true <-> b1 = true /\ b2 = true)).
idtac "Assumptions:".
Abort.
Print Assumptions andb_true_iff.
Goal True.
idtac " ".
idtac "#> orb_true_iff".
idtac "Possible points: 1".
check_type @orb_true_iff (
(forall b1 b2 : bool, b1 || b2 = true <-> b1 = true \/ b2 = true)).
idtac "Assumptions:".
Abort.
Print Assumptions orb_true_iff.
Goal True.
idtac " ".
idtac "------------------- eqb_neq --------------------".
idtac " ".
idtac "#> eqb_neq".
idtac "Possible points: 1".
check_type @eqb_neq ((forall x y : nat, (x =? y) = false <-> x <> y)).
idtac "Assumptions:".
Abort.
Print Assumptions eqb_neq.
Goal True.
idtac " ".
idtac "------------------- eqb_list --------------------".
idtac " ".
idtac "#> eqb_list".
idtac "Possible points: 3".
check_type @eqb_list ((forall A : Type, (A -> A -> bool) -> list A -> list A -> bool)).
idtac "Assumptions:".
Abort.
Print Assumptions eqb_list.
Goal True.
idtac " ".
idtac "------------------- All_forallb --------------------".
idtac " ".
idtac "#> forallb_true_iff".
idtac "Possible points: 2".
check_type @forallb_true_iff (
(forall (X : Type) (test : X -> bool) (l : list X),
@forallb X test l = true <-> @All X (fun x : X => test x = true) l)).
idtac "Assumptions:".
Abort.
Print Assumptions forallb_true_iff.
Goal True.
idtac " ".
idtac "------------------- excluded_middle_irrefutable --------------------".
idtac " ".
idtac "#> excluded_middle_irrefutable".
idtac "Possible points: 3".
check_type @excluded_middle_irrefutable ((forall P : Prop, ~ ~ (P \/ ~ P))).
idtac "Assumptions:".
Abort.
Print Assumptions excluded_middle_irrefutable.
Goal True.
idtac " ".
idtac "------------------- not_exists_dist --------------------".
idtac " ".
idtac "#> not_exists_dist".
idtac "Advanced".
idtac "Possible points: 3".
check_type @not_exists_dist (
(excluded_middle ->
forall (X : Type) (P : X -> Prop),
~ (exists x : X, ~ P x) -> forall x : X, P x)).
idtac "Assumptions:".
Abort.
Print Assumptions not_exists_dist.
Goal True.
idtac " ".
idtac " ".
idtac "Max points - standard: 43".
idtac "Max points - advanced: 49".
idtac "".
idtac "********** Summary **********".
idtac "".
idtac "********** Standard **********".
idtac "---------- and_exercise ---------".
Print Assumptions and_exercise.
idtac "---------- and_assoc ---------".
Print Assumptions and_assoc.
idtac "---------- mult_eq_0 ---------".
Print Assumptions mult_eq_0.
idtac "---------- or_commut ---------".
Print Assumptions or_commut.
idtac "---------- contrapositive ---------".
Print Assumptions contrapositive.
idtac "---------- not_both_true_and_false ---------".
Print Assumptions not_both_true_and_false.
idtac "---------- or_distributes_over_and ---------".
Print Assumptions or_distributes_over_and.
idtac "---------- dist_not_exists ---------".
Print Assumptions dist_not_exists.
idtac "---------- dist_exists_or ---------".
Print Assumptions dist_exists_or.
idtac "---------- In_map_iff ---------".
Print Assumptions In_map_iff.
idtac "---------- In_app_iff ---------".
Print Assumptions In_app_iff.
idtac "---------- All ---------".
Print Assumptions All.
idtac "---------- combine_odd_even ---------".
Print Assumptions combine_odd_even.
idtac "---------- tr_rev_correct ---------".
Print Assumptions tr_rev_correct.
idtac "---------- evenb_double_conv ---------".
Print Assumptions evenb_double_conv.
idtac "---------- andb_true_iff ---------".
Print Assumptions andb_true_iff.
idtac "---------- orb_true_iff ---------".
Print Assumptions orb_true_iff.
idtac "---------- eqb_neq ---------".
Print Assumptions eqb_neq.
idtac "---------- eqb_list ---------".
Print Assumptions eqb_list.
idtac "---------- forallb_true_iff ---------".
Print Assumptions forallb_true_iff.
idtac "---------- excluded_middle_irrefutable ---------".
Print Assumptions excluded_middle_irrefutable.
idtac "".
idtac "********** Advanced **********".
idtac "---------- double_neg_inf ---------".
idtac "MANUAL".
idtac "---------- informal_not_PNP ---------".
idtac "MANUAL".
idtac "---------- not_exists_dist ---------".
Print Assumptions not_exists_dist.
Abort.
(* Wed Jan 9 12:02:13 EST 2019 *)

828
Makefile Normal file
View file

@ -0,0 +1,828 @@
###############################################################################
## v # The Coq Proof Assistant ##
## <O___,, # INRIA - CNRS - LIX - LRI - PPS ##
## \VV/ # ##
## // # ##
###############################################################################
## GNUMakefile for Coq 8.11.1
# For debugging purposes (must stay here, don't move below)
INITIAL_VARS := $(.VARIABLES)
# To implement recursion we save the name of the main Makefile
SELF := $(lastword $(MAKEFILE_LIST))
PARENT := $(firstword $(MAKEFILE_LIST))
# This file is generated by coq_makefile and contains many variable
# definitions, like the list of .v files or the path to Coq
include Makefile.conf
# Put in place old names
VFILES := $(COQMF_VFILES)
MLIFILES := $(COQMF_MLIFILES)
MLFILES := $(COQMF_MLFILES)
MLGFILES := $(COQMF_MLGFILES)
MLPACKFILES := $(COQMF_MLPACKFILES)
MLLIBFILES := $(COQMF_MLLIBFILES)
CMDLINE_VFILES := $(COQMF_CMDLINE_VFILES)
INSTALLCOQDOCROOT := $(COQMF_INSTALLCOQDOCROOT)
OTHERFLAGS := $(COQMF_OTHERFLAGS)
COQ_SRC_SUBDIRS := $(COQMF_COQ_SRC_SUBDIRS)
OCAMLLIBS := $(COQMF_OCAMLLIBS)
SRC_SUBDIRS := $(COQMF_SRC_SUBDIRS)
COQLIBS := $(COQMF_COQLIBS)
COQLIBS_NOML := $(COQMF_COQLIBS_NOML)
CMDLINE_COQLIBS := $(COQMF_CMDLINE_COQLIBS)
LOCAL := $(COQMF_LOCAL)
COQLIB := $(COQMF_COQLIB)
DOCDIR := $(COQMF_DOCDIR)
OCAMLFIND := $(COQMF_OCAMLFIND)
CAMLFLAGS := $(COQMF_CAMLFLAGS)
HASNATDYNLINK := $(COQMF_HASNATDYNLINK)
OCAMLWARN := $(COQMF_WARN)
Makefile.conf:
coq_makefile -Q . LF OTHERFLAGS = '-Q . LF ' COQLIBS = '' -install none Preface.v Basics.v Induction.v Lists.v Poly.v Tactics.v Logic.v IndProp.v Maps.v ProofObjects.v IndPrinciples.v Rel.v Imp.v ImpParser.v ImpCEvalFun.v Extraction.v Auto.v Postscript.v Bib.v PrefaceTest.v BasicsTest.v InductionTest.v ListsTest.v PolyTest.v TacticsTest.v LogicTest.v IndPropTest.v MapsTest.v ProofObjectsTest.v IndPrinciplesTest.v RelTest.v ImpTest.v ImpParserTest.v ImpCEvalFunTest.v ExtractionTest.v AutoTest.v PostscriptTest.v BibTest.v -o Makefile
# This file can be created by the user to hook into double colon rules or
# add any other Makefile code he may need
-include Makefile.local
# Parameters ##################################################################
#
# Parameters are make variable assignments.
# They can be passed to (each call to) make on the command line.
# They can also be put in Makefile.local once an for all.
# For retro-compatibility reasons they can be put in the _CoqProject, but this
# practice is discouraged since _CoqProject better not contain make specific
# code (be nice to user interfaces).
# Print shell commands (set to non empty)
VERBOSE ?=
# Time the Coq process (set to non empty), and how (see default value)
TIMED?=
TIMECMD?=
# Use command time on linux, gtime on Mac OS
TIMEFMT?="$* (real: %e, user: %U, sys: %S, mem: %M ko)"
ifneq (,$(TIMED))
ifeq (0,$(shell command time -f $(TIMEFMT) true >/dev/null 2>/dev/null; echo $$?))
STDTIME?=command time -f $(TIMEFMT)
else
ifeq (0,$(shell gtime -f $(TIMEFMT) true >/dev/null 2>/dev/null; echo $$?))
STDTIME?=gtime -f $(TIMEFMT)
else
STDTIME?=command time
endif
endif
else
STDTIME?=command time -f $(TIMEFMT)
endif
ifneq (,$(COQBIN))
# add an ending /
COQBIN:=$(COQBIN)/
endif
# Coq binaries
COQC ?= "$(COQBIN)coqc"
COQTOP ?= "$(COQBIN)coqtop"
COQCHK ?= "$(COQBIN)coqchk"
COQDEP ?= "$(COQBIN)coqdep"
COQDOC ?= "$(COQBIN)coqdoc"
COQPP ?= "$(COQBIN)coqpp"
COQMKFILE ?= "$(COQBIN)coq_makefile"
# Timing scripts
COQMAKE_ONE_TIME_FILE ?= "$(COQLIB)/tools/make-one-time-file.py"
COQMAKE_BOTH_TIME_FILES ?= "$(COQLIB)/tools/make-both-time-files.py"
COQMAKE_BOTH_SINGLE_TIMING_FILES ?= "$(COQLIB)/tools/make-both-single-timing-files.py"
BEFORE ?=
AFTER ?=
# FIXME this should be generated by Coq (modules already linked by Coq)
CAMLDONTLINK=unix,str
# OCaml binaries
CAMLC ?= "$(OCAMLFIND)" ocamlc -c
CAMLOPTC ?= "$(OCAMLFIND)" opt -c
CAMLLINK ?= "$(OCAMLFIND)" ocamlc -linkpkg -dontlink $(CAMLDONTLINK)
CAMLOPTLINK ?= "$(OCAMLFIND)" opt -linkpkg -dontlink $(CAMLDONTLINK)
CAMLDOC ?= "$(OCAMLFIND)" ocamldoc
CAMLDEP ?= "$(OCAMLFIND)" ocamldep -slash -ml-synonym .mlpack
# DESTDIR is prepended to all installation paths
DESTDIR ?=
# Debug builds, typically -g to OCaml, -debug to Coq.
CAMLDEBUG ?=
COQDEBUG ?=
# Extra packages to be linked in (as in findlib -package)
CAMLPKGS ?=
# Option for making timing files
TIMING?=
# Option for changing sorting of timing output file
TIMING_SORT_BY ?= auto
# Output file names for timed builds
TIME_OF_BUILD_FILE ?= time-of-build.log
TIME_OF_BUILD_BEFORE_FILE ?= time-of-build-before.log
TIME_OF_BUILD_AFTER_FILE ?= time-of-build-after.log
TIME_OF_PRETTY_BUILD_FILE ?= time-of-build-pretty.log
TIME_OF_PRETTY_BOTH_BUILD_FILE ?= time-of-build-both.log
TIME_OF_PRETTY_BUILD_EXTRA_FILES ?= - # also output to the command line
TGTS ?=
########## End of parameters ##################################################
# What follows may be relevant to you only if you need to
# extend this Makefile. If so, look for 'Extension point' here and
# put in Makefile.local double colon rules accordingly.
# E.g. to perform some work after the all target completes you can write
#
# post-all::
# echo "All done!"
#
# in Makefile.local
#
###############################################################################
# Flags #######################################################################
#
# We define a bunch of variables combining the parameters.
# To add additional flags to coq, coqchk or coqdoc, set the
# {COQ,COQCHK,COQDOC}EXTRAFLAGS variable to whatever you want to add.
# To overwrite the default choice and set your own flags entirely, set the
# {COQ,COQCHK,COQDOC}FLAGS variable.
SHOW := $(if $(VERBOSE),@true "",@echo "")
HIDE := $(if $(VERBOSE),,@)
TIMER=$(if $(TIMED), $(STDTIME), $(TIMECMD))
OPT?=
# The DYNOBJ and DYNLIB variables are used by "coqdep -dyndep var" in .v.d
ifeq '$(OPT)' '-byte'
USEBYTE:=true
DYNOBJ:=.cma
DYNLIB:=.cma
else
USEBYTE:=
DYNOBJ:=.cmxs
DYNLIB:=.cmxs
endif
# these variables are meant to be overridden if you want to add *extra* flags
COQEXTRAFLAGS?=
COQCHKEXTRAFLAGS?=
COQDOCEXTRAFLAGS?=
# these flags do NOT contain the libraries, to make them easier to overwrite
COQFLAGS?=-q $(OTHERFLAGS) $(COQEXTRAFLAGS)
COQCHKFLAGS?=-silent -o $(COQCHKEXTRAFLAGS)
COQDOCFLAGS?=-interpolate -utf8 $(COQDOCEXTRAFLAGS)
COQDOCLIBS?=$(COQLIBS_NOML)
# The version of Coq being run and the version of coq_makefile that
# generated this makefile
COQ_VERSION:=$(shell $(COQC) --print-version | cut -d " " -f 1)
COQMAKEFILE_VERSION:=8.11.1
COQSRCLIBS?= $(foreach d,$(COQ_SRC_SUBDIRS), -I "$(COQLIB)/$(d)")
CAMLFLAGS+=$(OCAMLLIBS) $(COQSRCLIBS)
# ocamldoc fails with unknown argument otherwise
CAMLDOCFLAGS:=$(filter-out -annot, $(filter-out -bin-annot, $(CAMLFLAGS)))
CAMLFLAGS+=$(OCAMLWARN)
ifneq (,$(TIMING))
TIMING_ARG=-time
ifeq (after,$(TIMING))
TIMING_EXT=after-timing
else
ifeq (before,$(TIMING))
TIMING_EXT=before-timing
else
TIMING_EXT=timing
endif
endif
else
TIMING_ARG=
endif
# Retro compatibility (DESTDIR is standard on Unix, DSTROOT is not)
ifdef DSTROOT
DESTDIR := $(DSTROOT)
endif
concat_path = $(if $(1),$(1)/$(if $(COQMF_WINDRIVE),$(subst $(COQMF_WINDRIVE),/,$(2)),$(2)),$(2))
COQLIBINSTALL = $(call concat_path,$(DESTDIR),$(COQLIB)/user-contrib)
COQDOCINSTALL = $(call concat_path,$(DESTDIR),$(DOCDIR)/user-contrib)
COQTOPINSTALL = $(call concat_path,$(DESTDIR),$(COQLIB)/toploop)
# Files #######################################################################
#
# We here define a bunch of variables about the files being part of the
# Coq project in order to ease the writing of build target and build rules
VDFILE := .Makefile.d
ALLSRCFILES := \
$(MLGFILES) \
$(MLFILES) \
$(MLPACKFILES) \
$(MLLIBFILES) \
$(MLIFILES)
# helpers
vo_to_obj = $(addsuffix .o,\
$(filter-out Warning: Error:,\
$(shell $(COQTOP) -q -noinit -batch -quiet -print-mod-uid $(1))))
strip_dotslash = $(patsubst ./%,%,$(1))
# without this we get undefined variables in the expansion for the
# targets of the [deprecated,use-mllib-or-mlpack] rule
with_undef = $(if $(filter-out undefined, $(origin $(1))),$($(1)))
VO = vo
VOS = vos
VOFILES = $(VFILES:.v=.$(VO))
GLOBFILES = $(VFILES:.v=.glob)
HTMLFILES = $(VFILES:.v=.html)
GHTMLFILES = $(VFILES:.v=.g.html)
BEAUTYFILES = $(addsuffix .beautified,$(VFILES))
TEXFILES = $(VFILES:.v=.tex)
GTEXFILES = $(VFILES:.v=.g.tex)
CMOFILES = \
$(MLGFILES:.mlg=.cmo) \
$(MLFILES:.ml=.cmo) \
$(MLPACKFILES:.mlpack=.cmo)
CMXFILES = $(CMOFILES:.cmo=.cmx)
OFILES = $(CMXFILES:.cmx=.o)
CMAFILES = $(MLLIBFILES:.mllib=.cma) $(MLPACKFILES:.mlpack=.cma)
CMXAFILES = $(CMAFILES:.cma=.cmxa)
CMIFILES = \
$(CMOFILES:.cmo=.cmi) \
$(MLIFILES:.mli=.cmi)
# the /if/ is because old _CoqProject did not list a .ml(pack|lib) but just
# a .mlg file
CMXSFILES = \
$(MLPACKFILES:.mlpack=.cmxs) \
$(CMXAFILES:.cmxa=.cmxs) \
$(if $(MLPACKFILES)$(CMXAFILES),,\
$(MLGFILES:.mlg=.cmxs) $(MLFILES:.ml=.cmxs))
# files that are packed into a plugin (no extension)
PACKEDFILES = \
$(call strip_dotslash, \
$(foreach lib, \
$(call strip_dotslash, \
$(MLPACKFILES:.mlpack=_MLPACK_DEPENDENCIES)),$(call with_undef,$(lib))))
# files that are archived into a .cma (mllib)
LIBEDFILES = \
$(call strip_dotslash, \
$(foreach lib, \
$(call strip_dotslash, \
$(MLLIBFILES:.mllib=_MLLIB_DEPENDENCIES)),$(call with_undef,$(lib))))
CMIFILESTOINSTALL = $(filter-out $(addsuffix .cmi,$(PACKEDFILES)),$(CMIFILES))
CMOFILESTOINSTALL = $(filter-out $(addsuffix .cmo,$(PACKEDFILES)),$(CMOFILES))
OBJFILES = $(call vo_to_obj,$(VOFILES))
ALLNATIVEFILES = \
$(OBJFILES:.o=.cmi) \
$(OBJFILES:.o=.cmx) \
$(OBJFILES:.o=.cmxs)
# trick: wildcard filters out non-existing files, so that `install` doesn't show
# warnings and `clean` doesn't pass to rm a list of files that is too long for
# the shell.
NATIVEFILES = $(wildcard $(ALLNATIVEFILES))
FILESTOINSTALL = \
$(VOFILES) \
$(VFILES) \
$(GLOBFILES) \
$(NATIVEFILES) \
$(CMIFILESTOINSTALL)
BYTEFILESTOINSTALL = \
$(CMOFILESTOINSTALL) \
$(CMAFILES)
ifeq '$(HASNATDYNLINK)' 'true'
DO_NATDYNLINK = yes
FILESTOINSTALL += $(CMXSFILES) $(CMXAFILES) $(CMOFILESTOINSTALL:.cmo=.cmx)
else
DO_NATDYNLINK =
endif
ALLDFILES = $(addsuffix .d,$(ALLSRCFILES)) $(VDFILE)
# Compilation targets #########################################################
all:
$(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" pre-all
$(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" real-all
$(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" post-all
.PHONY: all
all.timing.diff:
$(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" pre-all
$(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" real-all.timing.diff TIME_OF_PRETTY_BUILD_EXTRA_FILES=""
$(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" post-all
.PHONY: all.timing.diff
make-pretty-timed-before:: TIME_OF_BUILD_FILE=$(TIME_OF_BUILD_BEFORE_FILE)
make-pretty-timed-after:: TIME_OF_BUILD_FILE=$(TIME_OF_BUILD_AFTER_FILE)
make-pretty-timed make-pretty-timed-before make-pretty-timed-after::
$(HIDE)rm -f pretty-timed-success.ok
$(HIDE)($(MAKE) --no-print-directory -f "$(PARENT)" $(TGTS) TIMED=1 2>&1 && touch pretty-timed-success.ok) | tee -a $(TIME_OF_BUILD_FILE)
$(HIDE)rm pretty-timed-success.ok # must not be -f; must fail if the touch failed
print-pretty-timed::
$(HIDE)$(COQMAKE_ONE_TIME_FILE) $(TIME_OF_BUILD_FILE) $(TIME_OF_PRETTY_BUILD_FILE) $(TIME_OF_PRETTY_BUILD_EXTRA_FILES)
print-pretty-timed-diff::
$(HIDE)$(COQMAKE_BOTH_TIME_FILES) --sort-by=$(TIMING_SORT_BY) $(TIME_OF_BUILD_AFTER_FILE) $(TIME_OF_BUILD_BEFORE_FILE) $(TIME_OF_PRETTY_BOTH_BUILD_FILE) $(TIME_OF_PRETTY_BUILD_EXTRA_FILES)
ifeq (,$(BEFORE))
print-pretty-single-time-diff::
@echo 'Error: Usage: $(MAKE) print-pretty-single-time-diff AFTER=path/to/file.v.after-timing BEFORE=path/to/file.v.before-timing'
$(HIDE)false
else
ifeq (,$(AFTER))
print-pretty-single-time-diff::
@echo 'Error: Usage: $(MAKE) print-pretty-single-time-diff AFTER=path/to/file.v.after-timing BEFORE=path/to/file.v.before-timing'
$(HIDE)false
else
print-pretty-single-time-diff::
$(HIDE)$(COQMAKE_BOTH_SINGLE_TIMING_FILES) --sort-by=$(TIMING_SORT_BY) $(AFTER) $(BEFORE) $(TIME_OF_PRETTY_BUILD_FILE) $(TIME_OF_PRETTY_BUILD_EXTRA_FILES)
endif
endif
pretty-timed:
$(HIDE)$(MAKE) --no-print-directory -f "$(PARENT)" make-pretty-timed
$(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" print-pretty-timed
.PHONY: pretty-timed make-pretty-timed make-pretty-timed-before make-pretty-timed-after print-pretty-timed print-pretty-timed-diff print-pretty-single-time-diff
# Extension points for actions to be performed before/after the all target
pre-all::
@# Extension point
$(HIDE)if [ "$(COQMAKEFILE_VERSION)" != "$(COQ_VERSION)" ]; then\
echo "W: This Makefile was generated by Coq $(COQMAKEFILE_VERSION)";\
echo "W: while the current Coq version is $(COQ_VERSION)";\
fi
.PHONY: pre-all
post-all::
@# Extension point
.PHONY: post-all
real-all: $(VOFILES) $(if $(USEBYTE),bytefiles,optfiles)
.PHONY: real-all
real-all.timing.diff: $(VOFILES:.vo=.v.timing.diff)
.PHONY: real-all.timing.diff
bytefiles: $(CMOFILES) $(CMAFILES)
.PHONY: bytefiles
optfiles: $(if $(DO_NATDYNLINK),$(CMXSFILES))
.PHONY: optfiles
# FIXME, see Ralf's bugreport
# quick is deprecated, now renamed vio
vio: $(VOFILES:.vo=.vio)
.PHONY: vio
quick: vio
$(warning "'make quick' is deprecated, use 'make vio' or consider using 'vos' files")
.PHONY: quick
vio2vo:
$(TIMER) $(COQC) $(COQDEBUG) $(COQFLAGS) $(COQLIBS) \
-schedule-vio2vo $(J) $(VOFILES:%.vo=%.vio)
.PHONY: vio2vo
# quick2vo is undocumented
quick2vo:
$(HIDE)make -j $(J) vio
$(HIDE)VIOFILES=$$(for vofile in $(VOFILES); do \
viofile="$$(echo "$$vofile" | sed "s/\.vo$$/.vio/")"; \
if [ "$$vofile" -ot "$$viofile" -o ! -e "$$vofile" ]; then printf "$$viofile "; fi; \
done); \
echo "VIO2VO: $$VIOFILES"; \
if [ -n "$$VIOFILES" ]; then \
$(TIMER) $(COQC) $(COQDEBUG) $(COQFLAGS) $(COQLIBS) -schedule-vio2vo $(J) $$VIOFILES; \
fi
.PHONY: quick2vo
checkproofs:
$(TIMER) $(COQC) $(COQDEBUG) $(COQFLAGS) $(COQLIBS) \
-schedule-vio-checking $(J) $(VOFILES:%.vo=%.vio)
.PHONY: checkproofs
vos: $(VOFILES:%.vo=%.vos)
.PHONY: vos
vok: $(VOFILES:%.vo=%.vok)
.PHONY: vok
validate: $(VOFILES)
$(TIMER) $(COQCHK) $(COQCHKFLAGS) $(COQLIBS) $^
.PHONY: validate
only: $(TGTS)
.PHONY: only
# Documentation targets #######################################################
html: $(GLOBFILES) $(VFILES)
$(SHOW)'COQDOC -d html $(GAL)'
$(HIDE)mkdir -p html
$(HIDE)$(COQDOC) \
-toc $(COQDOCFLAGS) -html $(GAL) $(COQDOCLIBS) -d html $(VFILES)
mlihtml: $(MLIFILES:.mli=.cmi)
$(SHOW)'CAMLDOC -d $@'
$(HIDE)mkdir $@ || rm -rf $@/*
$(HIDE)$(CAMLDOC) -html \
-d $@ -m A $(CAMLDEBUG) $(CAMLDOCFLAGS) $(MLIFILES)
all-mli.tex: $(MLIFILES:.mli=.cmi)
$(SHOW)'CAMLDOC -latex $@'
$(HIDE)$(CAMLDOC) -latex \
-o $@ -m A $(CAMLDEBUG) $(CAMLDOCFLAGS) $(MLIFILES)
all.ps: $(VFILES)
$(SHOW)'COQDOC -ps $(GAL)'
$(HIDE)$(COQDOC) \
-toc $(COQDOCFLAGS) -ps $(GAL) $(COQDOCLIBS) \
-o $@ `$(COQDEP) -sort -suffix .v $(VFILES)`
all.pdf: $(VFILES)
$(SHOW)'COQDOC -pdf $(GAL)'
$(HIDE)$(COQDOC) \
-toc $(COQDOCFLAGS) -pdf $(GAL) $(COQDOCLIBS) \
-o $@ `$(COQDEP) -sort -suffix .v $(VFILES)`
# FIXME: not quite right, since the output name is different
gallinahtml: GAL=-g
gallinahtml: html
all-gal.ps: GAL=-g
all-gal.ps: all.ps
all-gal.pdf: GAL=-g
all-gal.pdf: all.pdf
# ?
beautify: $(BEAUTYFILES)
for file in $^; do mv $${file%.beautified} $${file%beautified}old && mv $${file} $${file%.beautified}; done
@echo 'Do not do "make clean" until you are sure that everything went well!'
@echo 'If there were a problem, execute "for file in $$(find . -name \*.v.old -print); do mv $${file} $${file%.old}; done" in your shell/'
.PHONY: beautify
# Installation targets ########################################################
#
# There rules can be extended in Makefile.local
# Extensions can't assume when they run.
install:
$(HIDE)code=0; for f in $(FILESTOINSTALL); do\
if ! [ -f "$$f" ]; then >&2 echo $$f does not exist; code=1; fi \
done; exit $$code
$(HIDE)for f in $(FILESTOINSTALL); do\
df="`$(COQMKFILE) -destination-of "$$f" $(COQLIBS)`";\
if [ "$$?" != "0" -o -z "$$df" ]; then\
echo SKIP "$$f" since it has no logical path;\
else\
install -d "$(COQLIBINSTALL)/$$df" &&\
install -m 0644 "$$f" "$(COQLIBINSTALL)/$$df" &&\
echo INSTALL "$$f" "$(COQLIBINSTALL)/$$df";\
fi;\
done
$(HIDE)$(MAKE) install-extra -f "$(SELF)"
install-extra::
@# Extension point
.PHONY: install install-extra
install-byte:
$(HIDE)for f in $(BYTEFILESTOINSTALL); do\
df="`$(COQMKFILE) -destination-of "$$f" $(COQLIBS)`";\
if [ "$$?" != "0" -o -z "$$df" ]; then\
echo SKIP "$$f" since it has no logical path;\
else\
install -d "$(COQLIBINSTALL)/$$df" &&\
install -m 0644 "$$f" "$(COQLIBINSTALL)/$$df" &&\
echo INSTALL "$$f" "$(COQLIBINSTALL)/$$df";\
fi;\
done
install-doc:: html mlihtml
@# Extension point
$(HIDE)install -d "$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/html"
$(HIDE)for i in html/*; do \
dest="$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/$$i";\
install -m 0644 "$$i" "$$dest";\
echo INSTALL "$$i" "$$dest";\
done
$(HIDE)install -d \
"$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/mlihtml"
$(HIDE)for i in mlihtml/*; do \
dest="$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/$$i";\
install -m 0644 "$$i" "$$dest";\
echo INSTALL "$$i" "$$dest";\
done
.PHONY: install-doc
uninstall::
@# Extension point
$(HIDE)for f in $(FILESTOINSTALL); do \
df="`$(COQMKFILE) -destination-of "$$f" $(COQLIBS)`" &&\
instf="$(COQLIBINSTALL)/$$df/`basename $$f`" &&\
rm -f "$$instf" &&\
echo RM "$$instf" &&\
(rmdir "$(call concat_path,,$(COQLIBINSTALL)/$$df/)" 2>/dev/null || true); \
done
.PHONY: uninstall
uninstall-doc::
@# Extension point
$(SHOW)'RM $(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/html'
$(HIDE)rm -rf "$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/html"
$(SHOW)'RM $(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/mlihtml'
$(HIDE)rm -rf "$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/mlihtml"
$(HIDE) rmdir "$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/" || true
.PHONY: uninstall-doc
# Cleaning ####################################################################
#
# There rules can be extended in Makefile.local
# Extensions can't assume when they run.
clean::
@# Extension point
$(SHOW)'CLEAN'
$(HIDE)rm -f $(CMOFILES)
$(HIDE)rm -f $(CMIFILES)
$(HIDE)rm -f $(CMAFILES)
$(HIDE)rm -f $(CMOFILES:.cmo=.cmx)
$(HIDE)rm -f $(CMXAFILES)
$(HIDE)rm -f $(CMXSFILES)
$(HIDE)rm -f $(CMOFILES:.cmo=.o)
$(HIDE)rm -f $(CMXAFILES:.cmxa=.a)
$(HIDE)rm -f $(MLGFILES:.mlg=.ml)
$(HIDE)rm -f $(ALLDFILES)
$(HIDE)rm -f $(NATIVEFILES)
$(HIDE)find . -name .coq-native -type d -empty -delete
$(HIDE)rm -f $(VOFILES)
$(HIDE)rm -f $(VOFILES:.vo=.vio)
$(HIDE)rm -f $(VOFILES:.vo=.vos)
$(HIDE)rm -f $(VOFILES:.vo=.vok)
$(HIDE)rm -f $(BEAUTYFILES) $(VFILES:=.old)
$(HIDE)rm -f all.ps all-gal.ps all.pdf all-gal.pdf all.glob all-mli.tex
$(HIDE)rm -f $(VFILES:.v=.glob)
$(HIDE)rm -f $(VFILES:.v=.tex)
$(HIDE)rm -f $(VFILES:.v=.g.tex)
$(HIDE)rm -f pretty-timed-success.ok
$(HIDE)rm -rf html mlihtml
.PHONY: clean
cleanall:: clean
@# Extension point
$(SHOW)'CLEAN *.aux *.timing'
$(HIDE)rm -f $(foreach f,$(VFILES:.v=),$(dir $(f)).$(notdir $(f)).aux)
$(HIDE)rm -f $(TIME_OF_BUILD_FILE) $(TIME_OF_BUILD_BEFORE_FILE) $(TIME_OF_BUILD_AFTER_FILE) $(TIME_OF_PRETTY_BUILD_FILE) $(TIME_OF_PRETTY_BOTH_BUILD_FILE)
$(HIDE)rm -f $(VOFILES:.vo=.v.timing)
$(HIDE)rm -f $(VOFILES:.vo=.v.before-timing)
$(HIDE)rm -f $(VOFILES:.vo=.v.after-timing)
$(HIDE)rm -f $(VOFILES:.vo=.v.timing.diff)
.PHONY: cleanall
archclean::
@# Extension point
$(SHOW)'CLEAN *.cmx *.o'
$(HIDE)rm -f $(NATIVEFILES)
$(HIDE)rm -f $(CMOFILES:%.cmo=%.cmx)
.PHONY: archclean
# Compilation rules ###########################################################
$(MLIFILES:.mli=.cmi): %.cmi: %.mli
$(SHOW)'CAMLC -c $<'
$(HIDE)$(CAMLC) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) $<
$(MLGFILES:.mlg=.ml): %.ml: %.mlg
$(SHOW)'COQPP $<'
$(HIDE)$(COQPP) $<
# Stupid hack around a deficient syntax: we cannot concatenate two expansions
$(filter %.cmo, $(MLFILES:.ml=.cmo) $(MLGFILES:.mlg=.cmo)): %.cmo: %.ml
$(SHOW)'CAMLC -c $<'
$(HIDE)$(CAMLC) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) $<
# Same hack
$(filter %.cmx, $(MLFILES:.ml=.cmx) $(MLGFILES:.mlg=.cmx)): %.cmx: %.ml
$(SHOW)'CAMLOPT -c $(FOR_PACK) $<'
$(HIDE)$(CAMLOPTC) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) $(FOR_PACK) $<
$(MLLIBFILES:.mllib=.cmxs): %.cmxs: %.cmxa
$(SHOW)'CAMLOPT -shared -o $@'
$(HIDE)$(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) \
-linkall -shared -o $@ $<
$(MLLIBFILES:.mllib=.cma): %.cma: | %.mllib
$(SHOW)'CAMLC -a -o $@'
$(HIDE)$(CAMLLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) -a -o $@ $^
$(MLLIBFILES:.mllib=.cmxa): %.cmxa: | %.mllib
$(SHOW)'CAMLOPT -a -o $@'
$(HIDE)$(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) -a -o $@ $^
$(MLPACKFILES:.mlpack=.cmxs): %.cmxs: %.cmxa
$(SHOW)'CAMLOPT -shared -o $@'
$(HIDE)$(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) \
-shared -linkall -o $@ $<
$(MLPACKFILES:.mlpack=.cmxa): %.cmxa: %.cmx
$(SHOW)'CAMLOPT -a -o $@'
$(HIDE)$(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) -a -o $@ $<
$(MLPACKFILES:.mlpack=.cma): %.cma: %.cmo | %.mlpack
$(SHOW)'CAMLC -a -o $@'
$(HIDE)$(CAMLLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) -a -o $@ $^
$(MLPACKFILES:.mlpack=.cmo): %.cmo: | %.mlpack
$(SHOW)'CAMLC -pack -o $@'
$(HIDE)$(CAMLLINK) $(CAMLDEBUG) $(CAMLFLAGS) -pack -o $@ $^
$(MLPACKFILES:.mlpack=.cmx): %.cmx: | %.mlpack
$(SHOW)'CAMLOPT -pack -o $@'
$(HIDE)$(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) -pack -o $@ $^
# This rule is for _CoqProject with no .mllib nor .mlpack
$(filter-out $(MLLIBFILES:.mllib=.cmxs) $(MLPACKFILES:.mlpack=.cmxs) $(addsuffix .cmxs,$(PACKEDFILES)) $(addsuffix .cmxs,$(LIBEDFILES)),$(MLFILES:.ml=.cmxs) $(MLGFILES:.mlg=.cmxs)): %.cmxs: %.cmx
$(SHOW)'[deprecated,use-mllib-or-mlpack] CAMLOPT -shared -o $@'
$(HIDE)$(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) \
-shared -o $@ $<
ifneq (,$(TIMING))
TIMING_EXTRA = > $<.$(TIMING_EXT)
else
TIMING_EXTRA =
endif
$(VOFILES): %.vo: %.v
$(SHOW)COQC $<
$(HIDE)$(TIMER) $(COQC) $(COQDEBUG) $(TIMING_ARG) $(COQFLAGS) $(COQLIBS) $< $(TIMING_EXTRA)
# FIXME ?merge with .vo / .vio ?
$(GLOBFILES): %.glob: %.v
$(TIMER) $(COQC) $(COQDEBUG) $(COQFLAGS) $(COQLIBS) $<
$(VFILES:.v=.vio): %.vio: %.v
$(SHOW)COQC -vio $<
$(HIDE)$(TIMER) $(COQC) -vio $(COQDEBUG) $(COQFLAGS) $(COQLIBS) $<
$(VFILES:.v=.vos): %.vos: %.v
$(SHOW)COQC -vos $<
$(HIDE)$(TIMER) $(COQC) -vos $(COQDEBUG) $(COQFLAGS) $(COQLIBS) $<
$(VFILES:.v=.vok): %.vok: %.v
$(SHOW)COQC -vok $<
$(HIDE)$(TIMER) $(COQC) -vok $(COQDEBUG) $(COQFLAGS) $(COQLIBS) $<
$(addsuffix .timing.diff,$(VFILES)): %.timing.diff : %.before-timing %.after-timing
$(SHOW)PYTHON TIMING-DIFF $<
$(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" print-pretty-single-time-diff BEFORE=$*.before-timing AFTER=$*.after-timing TIME_OF_PRETTY_BUILD_FILE="$@"
$(BEAUTYFILES): %.v.beautified: %.v
$(SHOW)'BEAUTIFY $<'
$(HIDE)$(TIMER) $(COQC) $(COQDEBUG) $(COQFLAGS) $(COQLIBS) -beautify $<
$(TEXFILES): %.tex: %.v
$(SHOW)'COQDOC -latex $<'
$(HIDE)$(COQDOC) $(COQDOCFLAGS) -latex $< -o $@
$(GTEXFILES): %.g.tex: %.v
$(SHOW)'COQDOC -latex -g $<'
$(HIDE)$(COQDOC) $(COQDOCFLAGS) -latex -g $< -o $@
$(HTMLFILES): %.html: %.v %.glob
$(SHOW)'COQDOC -html $<'
$(HIDE)$(COQDOC) $(COQDOCFLAGS) -html $< -o $@
$(GHTMLFILES): %.g.html: %.v %.glob
$(SHOW)'COQDOC -html -g $<'
$(HIDE)$(COQDOC) $(COQDOCFLAGS) -html -g $< -o $@
# Dependency files ############################################################
ifndef MAKECMDGOALS
-include $(ALLDFILES)
else
ifneq ($(filter-out archclean clean cleanall printenv make-pretty-timed make-pretty-timed-before make-pretty-timed-after print-pretty-timed print-pretty-timed-diff print-pretty-single-time-diff,$(MAKECMDGOALS)),)
-include $(ALLDFILES)
endif
endif
.SECONDARY: $(ALLDFILES)
redir_if_ok = > "$@" || ( RV=$$?; rm -f "$@"; exit $$RV )
GENMLFILES:=$(MLGFILES:.mlg=.ml)
$(addsuffix .d,$(ALLSRCFILES)): $(GENMLFILES)
$(addsuffix .d,$(MLIFILES)): %.mli.d: %.mli
$(SHOW)'CAMLDEP $<'
$(HIDE)$(CAMLDEP) $(OCAMLLIBS) "$<" $(redir_if_ok)
$(addsuffix .d,$(MLGFILES)): %.mlg.d: %.ml
$(SHOW)'CAMLDEP $<'
$(HIDE)$(CAMLDEP) $(OCAMLLIBS) "$<" $(redir_if_ok)
$(addsuffix .d,$(MLFILES)): %.ml.d: %.ml
$(SHOW)'CAMLDEP $<'
$(HIDE)$(CAMLDEP) $(OCAMLLIBS) "$<" $(redir_if_ok)
$(addsuffix .d,$(MLLIBFILES)): %.mllib.d: %.mllib
$(SHOW)'COQDEP $<'
$(HIDE)$(COQDEP) $(OCAMLLIBS) -c "$<" $(redir_if_ok)
$(addsuffix .d,$(MLPACKFILES)): %.mlpack.d: %.mlpack
$(SHOW)'COQDEP $<'
$(HIDE)$(COQDEP) $(OCAMLLIBS) -c "$<" $(redir_if_ok)
# If this makefile is created using a _CoqProject we have coqdep get
# options from it. This avoids argument length limits for pathological
# projects. Note that extra options might be on the command line.
VDFILE_FLAGS:=$(if ,-f ,) $(CMDLINE_COQLIBS) $(CMDLINE_VFILES)
$(VDFILE): $(VFILES)
$(SHOW)'COQDEP VFILES'
$(HIDE)$(COQDEP) -vos -dyndep var $(VDFILE_FLAGS) $(redir_if_ok)
# Misc ########################################################################
byte:
$(HIDE)$(MAKE) all "OPT:=-byte" -f "$(SELF)"
.PHONY: byte
opt:
$(HIDE)$(MAKE) all "OPT:=-opt" -f "$(SELF)"
.PHONY: opt
# This is deprecated. To extend this makefile use
# extension points and Makefile.local
printenv::
$(warning printenv is deprecated)
$(warning write extensions in Makefile.local or include Makefile.conf)
@echo 'LOCAL = $(LOCAL)'
@echo 'COQLIB = $(COQLIB)'
@echo 'DOCDIR = $(DOCDIR)'
@echo 'OCAMLFIND = $(OCAMLFIND)'
@echo 'HASNATDYNLINK = $(HASNATDYNLINK)'
@echo 'SRC_SUBDIRS = $(SRC_SUBDIRS)'
@echo 'COQ_SRC_SUBDIRS = $(COQ_SRC_SUBDIRS)'
@echo 'OCAMLFIND = $(OCAMLFIND)'
@echo 'PP = $(PP)'
@echo 'COQFLAGS = $(COQFLAGS)'
@echo 'COQLIB = $(COQLIBS)'
@echo 'COQLIBINSTALL = $(COQLIBINSTALL)'
@echo 'COQDOCINSTALL = $(COQDOCINSTALL)'
.PHONY: printenv
# Generate a .merlin file. If you need to append directives to this
# file you can extend the merlin-hook target in Makefile.local
.merlin:
$(SHOW)'FILL .merlin'
$(HIDE)echo 'FLG $(COQMF_CAMLFLAGS)' > .merlin
$(HIDE)echo 'B $(COQLIB)' >> .merlin
$(HIDE)echo 'S $(COQLIB)' >> .merlin
$(HIDE)$(foreach d,$(COQ_SRC_SUBDIRS), \
echo 'B $(COQLIB)$(d)' >> .merlin;)
$(HIDE)$(foreach d,$(COQ_SRC_SUBDIRS), \
echo 'S $(COQLIB)$(d)' >> .merlin;)
$(HIDE)$(foreach d,$(SRC_SUBDIRS), echo 'B $(d)' >> .merlin;)
$(HIDE)$(foreach d,$(SRC_SUBDIRS), echo 'S $(d)' >> .merlin;)
$(HIDE)$(MAKE) merlin-hook -f "$(SELF)"
.PHONY: merlin
merlin-hook::
@# Extension point
.PHONY: merlin-hook
# prints all variables
debug:
$(foreach v,\
$(sort $(filter-out $(INITIAL_VARS) INITIAL_VARS,\
$(.VARIABLES))),\
$(info $(v) = $($(v))))
.PHONY: debug
.DEFAULT_GOAL := all
# Local Variables:
# mode: makefile-gmake
# End:

56
Makefile.conf Normal file
View file

@ -0,0 +1,56 @@
# This configuration file was generated by running:
# coq_makefile -Q . LF OTHERFLAGS = '-Q . LF ' COQLIBS = '' -install none Preface.v Basics.v Induction.v Lists.v Poly.v Tactics.v Logic.v IndProp.v Maps.v ProofObjects.v IndPrinciples.v Rel.v Imp.v ImpParser.v ImpCEvalFun.v Extraction.v Auto.v Postscript.v Bib.v PrefaceTest.v BasicsTest.v InductionTest.v ListsTest.v PolyTest.v TacticsTest.v LogicTest.v IndPropTest.v MapsTest.v ProofObjectsTest.v IndPrinciplesTest.v RelTest.v ImpTest.v ImpParserTest.v ImpCEvalFunTest.v ExtractionTest.v AutoTest.v PostscriptTest.v BibTest.v -o Makefile
###############################################################################
# #
# Project files. #
# #
###############################################################################
COQMF_VFILES = Preface.v Basics.v Induction.v Lists.v Poly.v Tactics.v Logic.v IndProp.v Maps.v ProofObjects.v IndPrinciples.v Rel.v Imp.v ImpParser.v ImpCEvalFun.v Extraction.v Auto.v Postscript.v Bib.v PrefaceTest.v BasicsTest.v InductionTest.v ListsTest.v PolyTest.v TacticsTest.v LogicTest.v IndPropTest.v MapsTest.v ProofObjectsTest.v IndPrinciplesTest.v RelTest.v ImpTest.v ImpParserTest.v ImpCEvalFunTest.v ExtractionTest.v AutoTest.v PostscriptTest.v BibTest.v
COQMF_MLIFILES =
COQMF_MLFILES =
COQMF_MLGFILES =
COQMF_MLPACKFILES =
COQMF_MLLIBFILES =
COQMF_CMDLINE_VFILES = Preface.v Basics.v Induction.v Lists.v Poly.v Tactics.v Logic.v IndProp.v Maps.v ProofObjects.v IndPrinciples.v Rel.v Imp.v ImpParser.v ImpCEvalFun.v Extraction.v Auto.v Postscript.v Bib.v PrefaceTest.v BasicsTest.v InductionTest.v ListsTest.v PolyTest.v TacticsTest.v LogicTest.v IndPropTest.v MapsTest.v ProofObjectsTest.v IndPrinciplesTest.v RelTest.v ImpTest.v ImpParserTest.v ImpCEvalFunTest.v ExtractionTest.v AutoTest.v PostscriptTest.v BibTest.v
###############################################################################
# #
# Path directives (-I, -R, -Q). #
# #
###############################################################################
COQMF_OCAMLLIBS =
COQMF_SRC_SUBDIRS =
COQMF_COQLIBS = -Q . LF
COQMF_COQLIBS_NOML = -Q . LF
COQMF_CMDLINE_COQLIBS = -Q . LF
###############################################################################
# #
# Coq configuration. #
# #
###############################################################################
COQMF_LOCAL=0
COQMF_COQLIB=/home/michael/.opam/default/lib/coq/
COQMF_DOCDIR=/home/michael/.opam/default/doc/
COQMF_OCAMLFIND=/home/michael/.opam/default/bin/ocamlfind
COQMF_CAMLFLAGS=-thread -rectypes -w +a-4-9-27-41-42-44-45-48-58-67 -safe-string -strict-sequence
COQMF_WARN=-warn-error +a-3
COQMF_HASNATDYNLINK=true
COQMF_COQ_SRC_SUBDIRS=config lib clib kernel library engine pretyping interp gramlib gramlib/.pack parsing proofs tactics toplevel printing ide stm vernac plugins/btauto plugins/cc plugins/derive plugins/extraction plugins/firstorder plugins/fourier plugins/funind plugins/ltac plugins/micromega plugins/nsatz plugins/omega plugins/rtauto plugins/setoid_ring plugins/ssr plugins/ssrmatching plugins/syntax
COQMF_WINDRIVE=
###############################################################################
# #
# Extra variables. #
# #
###############################################################################
OTHERFLAGS = -Q . LF
COQLIBS =
COQMF_OTHERFLAGS =
COQMF_INSTALLCOQDOCROOT = LF

382
Maps.v Normal file
View file

@ -0,0 +1,382 @@
(** * Maps: Total and Partial Maps *)
(** _Maps_ (or _dictionaries_) are ubiquitous data structures both
generally and in the theory of programming languages in
particular; we're going to need them in many places in the coming
chapters. They also make a nice case study using ideas we've seen
in previous chapters, including building data structures out of
higher-order functions (from [Basics] and [Poly]) and the use of
reflection to streamline proofs (from [IndProp]).
We'll define two flavors of maps: _total_ maps, which include a
"default" element to be returned when a key being looked up
doesn't exist, and _partial_ maps, which return an [option] to
indicate success or failure. The latter is defined in terms of
the former, using [None] as the default element. *)
(* ################################################################# *)
(** * The Coq Standard Library *)
(** One small digression before we begin...
Unlike the chapters we have seen so far, this one does not
[Require Import] the chapter before it (and, transitively, all the
earlier chapters). Instead, in this chapter and from now, on
we're going to import the definitions and theorems we need
directly from Coq's standard library stuff. You should not notice
much difference, though, because we've been careful to name our
own definitions and theorems the same as their counterparts in the
standard library, wherever they overlap. *)
From Coq Require Import Arith.Arith.
From Coq Require Import Bool.Bool.
Require Export Coq.Strings.String.
From Coq Require Import Logic.FunctionalExtensionality.
From Coq Require Import Lists.List.
Import ListNotations.
(** Documentation for the standard library can be found at
http://coq.inria.fr/library/.
The [Search] command is a good way to look for theorems involving
objects of specific types. Take a minute now to experiment with it. *)
(* ################################################################# *)
(** * Identifiers *)
(** First, we need a type for the keys that we use to index into our
maps. In [Lists.v] we introduced a fresh type [id] for a similar
purpose; here and for the rest of _Software Foundations_ we will
use the [string] type from Coq's standard library. *)
(** To compare strings, we define the function [eqb_string], which
internally uses the function [string_dec] from Coq's string
library. *)
Definition eqb_string (x y : string) : bool :=
if string_dec x y then true else false.
(** (The function [string_dec] comes from Coq's string library.
If you check the result type of [string_dec], you'll see that it
does not actually return a [bool], but rather a type that looks
like [{x = y} + {x <> y}], called a [sumbool], which can be
thought of as an "evidence-carrying boolean." Formally, an
element of [sumbool] is either a proof that two things are equal
or a proof that they are unequal, together with a tag indicating
which. But for present purposes you can think of it as just a
fancy [bool].) *)
(** Now we need a few basic properties of string equality... *)
Theorem eqb_string_refl : forall s : string, true = eqb_string s s.
Proof. intros s. unfold eqb_string. destruct (string_dec s s) as [|Hs].
- reflexivity.
- destruct Hs. reflexivity.
Qed.
(** The following useful property follows from an analogous
lemma about strings: *)
Theorem eqb_string_true_iff : forall x y : string,
eqb_string x y = true <-> x = y.
Proof.
intros x y.
unfold eqb_string.
destruct (string_dec x y) as [|Hs].
- subst. split. reflexivity. reflexivity.
- split.
+ intros contra. discriminate contra.
+ intros H. rewrite H in Hs. destruct Hs. reflexivity.
Qed.
(** Similarly: *)
Theorem eqb_string_false_iff : forall x y : string,
eqb_string x y = false <-> x <> y.
Proof.
intros x y. rewrite <- eqb_string_true_iff.
rewrite not_true_iff_false. reflexivity. Qed.
(** This handy variant follows just by rewriting: *)
Theorem false_eqb_string : forall x y : string,
x <> y -> eqb_string x y = false.
Proof.
intros x y. rewrite eqb_string_false_iff.
intros H. apply H. Qed.
(* ################################################################# *)
(** * Total Maps *)
(** Our main job in this chapter will be to build a definition of
partial maps that is similar in behavior to the one we saw in the
[Lists] chapter, plus accompanying lemmas about its behavior.
This time around, though, we're going to use _functions_, rather
than lists of key-value pairs, to build maps. The advantage of
this representation is that it offers a more _extensional_ view of
maps, where two maps that respond to queries in the same way will
be represented as literally the same thing (the very same function),
rather than just "equivalent" data structures. This, in turn,
simplifies proofs that use maps. *)
(** We build partial maps in two steps. First, we define a type of
_total maps_ that return a default value when we look up a key
that is not present in the map. *)
Definition total_map (A : Type) := string -> A.
(** Intuitively, a total map over an element type [A] is just a
function that can be used to look up [string]s, yielding [A]s. *)
(** The function [t_empty] yields an empty total map, given a default
element; this map always returns the default element when applied
to any string. *)
Definition t_empty {A : Type} (v : A) : total_map A :=
(fun _ => v).
(** More interesting is the [update] function, which (as before) takes
a map [m], a key [x], and a value [v] and returns a new map that
takes [x] to [v] and takes every other key to whatever [m] does. *)
Definition t_update {A : Type} (m : total_map A)
(x : string) (v : A) :=
fun x' => if eqb_string x x' then v else m x'.
(** This definition is a nice example of higher-order programming:
[t_update] takes a _function_ [m] and yields a new function
[fun x' => ...] that behaves like the desired map. *)
(** For example, we can build a map taking [string]s to [bool]s, where
["foo"] and ["bar"] are mapped to [true] and every other key is
mapped to [false], like this: *)
Definition examplemap :=
t_update (t_update (t_empty false) "foo" true)
"bar" true.
(** Next, let's introduce some new notations to facilitate working
with maps. *)
(** First, we will use the following notation to create an empty
total map with a default value. *)
Notation "'_' '!->' v" := (t_empty v)
(at level 100, right associativity).
Example example_empty := (_ !-> false).
(** We then introduce a convenient notation for extending an existing
map with some bindings. *)
Notation "x '!->' v ';' m" := (t_update m x v)
(at level 100, v at next level, right associativity).
(** The [examplemap] above can now be defined as follows: *)
Definition examplemap' :=
( "bar" !-> true;
"foo" !-> true;
_ !-> false
).
(** This completes the definition of total maps. Note that we
don't need to define a [find] operation because it is just
function application! *)
Example update_example1 : examplemap' "baz" = false.
Proof. reflexivity. Qed.
Example update_example2 : examplemap' "foo" = true.
Proof. reflexivity. Qed.
Example update_example3 : examplemap' "quux" = false.
Proof. reflexivity. Qed.
Example update_example4 : examplemap' "bar" = true.
Proof. reflexivity. Qed.
(** To use maps in later chapters, we'll need several fundamental
facts about how they behave. *)
(** Even if you don't work the following exercises, make sure
you thoroughly understand the statements of the lemmas! *)
(** (Some of the proofs require the functional extensionality axiom,
which is discussed in the [Logic] chapter.) *)
(** **** Exercise: 1 star, standard, optional (t_apply_empty)
First, the empty map returns its default element for all keys: *)
Lemma t_apply_empty : forall (A : Type) (x : string) (v : A),
(_ !-> v) x = v.
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** **** Exercise: 2 stars, standard, optional (t_update_eq)
Next, if we update a map [m] at a key [x] with a new value [v]
and then look up [x] in the map resulting from the [update], we
get back [v]: *)
Lemma t_update_eq : forall (A : Type) (m : total_map A) x v,
(x !-> v ; m) x = v.
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** **** Exercise: 2 stars, standard, optional (t_update_neq)
On the other hand, if we update a map [m] at a key [x1] and then
look up a _different_ key [x2] in the resulting map, we get the
same result that [m] would have given: *)
Theorem t_update_neq : forall (A : Type) (m : total_map A) x1 x2 v,
x1 <> x2 ->
(x1 !-> v ; m) x2 = m x2.
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** **** Exercise: 2 stars, standard, optional (t_update_shadow)
If we update a map [m] at a key [x] with a value [v1] and then
update again with the same key [x] and another value [v2], the
resulting map behaves the same (gives the same result when applied
to any key) as the simpler map obtained by performing just
the second [update] on [m]: *)
Lemma t_update_shadow : forall (A : Type) (m : total_map A) x v1 v2,
(x !-> v2 ; x !-> v1 ; m) = (x !-> v2 ; m).
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** For the final two lemmas about total maps, it's convenient to use
the reflection idioms introduced in chapter [IndProp]. We begin
by proving a fundamental _reflection lemma_ relating the equality
proposition on [id]s with the boolean function [eqb_id]. *)
(** **** Exercise: 2 stars, standard, optional (eqb_stringP)
Use the proof of [eqbP] in chapter [IndProp] as a template to
prove the following: *)
Lemma eqb_stringP : forall x y : string,
reflect (x = y) (eqb_string x y).
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** Now, given [string]s [x1] and [x2], we can use the tactic
[destruct (eqb_stringP x1 x2)] to simultaneously perform case
analysis on the result of [eqb_string x1 x2] and generate
hypotheses about the equality (in the sense of [=]) of [x1]
and [x2]. *)
(** **** Exercise: 2 stars, standard (t_update_same)
With the example in chapter [IndProp] as a template, use
[eqb_stringP] to prove the following theorem, which states that
if we update a map to assign key [x] the same value as it already
has in [m], then the result is equal to [m]: *)
Theorem t_update_same : forall (A : Type) (m : total_map A) x,
(x !-> m x ; m) = m.
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** **** Exercise: 3 stars, standard, recommended (t_update_permute)
Use [eqb_stringP] to prove one final property of the [update]
function: If we update a map [m] at two distinct keys, it doesn't
matter in which order we do the updates. *)
Theorem t_update_permute : forall (A : Type) (m : total_map A)
v1 v2 x1 x2,
x2 <> x1 ->
(x1 !-> v1 ; x2 !-> v2 ; m)
=
(x2 !-> v2 ; x1 !-> v1 ; m).
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(* ################################################################# *)
(** * Partial maps *)
(** Finally, we define _partial maps_ on top of total maps. A partial
map with elements of type [A] is simply a total map with elements
of type [option A] and default element [None]. *)
Definition partial_map (A : Type) := total_map (option A).
Definition empty {A : Type} : partial_map A :=
t_empty None.
Definition update {A : Type} (m : partial_map A)
(x : string) (v : A) :=
(x !-> Some v ; m).
(** We introduce a similar notation for partial maps: *)
Notation "x '|->' v ';' m" := (update m x v)
(at level 100, v at next level, right associativity).
(** We can also hide the last case when it is empty. *)
Notation "x '|->' v" := (update empty x v)
(at level 100).
Example examplepmap :=
("Church" |-> true ; "Turing" |-> false).
(** We now straightforwardly lift all of the basic lemmas about total
maps to partial maps. *)
Lemma apply_empty : forall (A : Type) (x : string),
@empty A x = None.
Proof.
intros. unfold empty. rewrite t_apply_empty.
reflexivity.
Qed.
Lemma update_eq : forall (A : Type) (m : partial_map A) x v,
(x |-> v ; m) x = Some v.
Proof.
intros. unfold update. rewrite t_update_eq.
reflexivity.
Qed.
Theorem update_neq : forall (A : Type) (m : partial_map A) x1 x2 v,
x2 <> x1 ->
(x2 |-> v ; m) x1 = m x1.
Proof.
intros A m x1 x2 v H.
unfold update. rewrite t_update_neq. reflexivity.
apply H. Qed.
Lemma update_shadow : forall (A : Type) (m : partial_map A) x v1 v2,
(x |-> v2 ; x |-> v1 ; m) = (x |-> v2 ; m).
Proof.
intros A m x v1 v2. unfold update. rewrite t_update_shadow.
reflexivity.
Qed.
Theorem update_same : forall (A : Type) (m : partial_map A) x v,
m x = Some v ->
(x |-> v ; m) = m.
Proof.
intros A m x v H. unfold update. rewrite <- H.
apply t_update_same.
Qed.
Theorem update_permute : forall (A : Type) (m : partial_map A)
x1 x2 v1 v2,
x2 <> x1 ->
(x1 |-> v1 ; x2 |-> v2 ; m) = (x2 |-> v2 ; x1 |-> v1 ; m).
Proof.
intros A m x1 x2 v1 v2. unfold update.
apply t_update_permute.
Qed.
(* Wed Jan 9 12:02:45 EST 2019 *)

78
MapsTest.v Normal file
View file

@ -0,0 +1,78 @@
Set Warnings "-notation-overridden,-parsing".
From Coq Require Export String.
From LF Require Import Maps.
Parameter MISSING: Type.
Module Check.
Ltac check_type A B :=
match type of A with
| context[MISSING] => idtac "Missing:" A
| ?T => first [unify T B; idtac "Type: ok" | idtac "Type: wrong - should be (" B ")"]
end.
Ltac print_manual_grade A :=
match eval compute in A with
| Some (_ ?S ?C) =>
idtac "Score:" S;
match eval compute in C with
| ""%string => idtac "Comment: None"
| _ => idtac "Comment:" C
end
| None =>
idtac "Score: Ungraded";
idtac "Comment: None"
end.
End Check.
From LF Require Import Maps.
Import Check.
Goal True.
idtac "------------------- t_update_same --------------------".
idtac " ".
idtac "#> t_update_same".
idtac "Possible points: 2".
check_type @t_update_same (
(forall (A : Type) (m : total_map A) (x : string), (x !-> m x; m) = m)).
idtac "Assumptions:".
Abort.
Print Assumptions t_update_same.
Goal True.
idtac " ".
idtac "------------------- t_update_permute --------------------".
idtac " ".
idtac "#> t_update_permute".
idtac "Possible points: 3".
check_type @t_update_permute (
(forall (A : Type) (m : total_map A) (v1 v2 : A) (x1 x2 : string),
x2 <> x1 -> (x1 !-> v1; x2 !-> v2; m) = (x2 !-> v2; x1 !-> v1; m))).
idtac "Assumptions:".
Abort.
Print Assumptions t_update_permute.
Goal True.
idtac " ".
idtac " ".
idtac "Max points - standard: 5".
idtac "Max points - advanced: 5".
idtac "".
idtac "********** Summary **********".
idtac "".
idtac "********** Standard **********".
idtac "---------- t_update_same ---------".
Print Assumptions t_update_same.
idtac "---------- t_update_permute ---------".
Print Assumptions t_update_permute.
idtac "".
idtac "********** Advanced **********".
Abort.
(* Wed Jan 9 12:02:17 EST 2019 *)

1179
Poly.v Normal file

File diff suppressed because it is too large Load diff

424
PolyTest.v Normal file
View file

@ -0,0 +1,424 @@
Set Warnings "-notation-overridden,-parsing".
From Coq Require Export String.
From LF Require Import Poly.
Parameter MISSING: Type.
Module Check.
Ltac check_type A B :=
match type of A with
| context[MISSING] => idtac "Missing:" A
| ?T => first [unify T B; idtac "Type: ok" | idtac "Type: wrong - should be (" B ")"]
end.
Ltac print_manual_grade A :=
match eval compute in A with
| Some (_ ?S ?C) =>
idtac "Score:" S;
match eval compute in C with
| ""%string => idtac "Comment: None"
| _ => idtac "Comment:" C
end
| None =>
idtac "Score: Ungraded";
idtac "Comment: None"
end.
End Check.
From LF Require Import Poly.
Import Check.
Goal True.
idtac "------------------- mumble_grumble --------------------".
idtac " ".
idtac "#> Manually graded: mumble_grumble".
idtac "Possible points: 2".
print_manual_grade manual_grade_for_mumble_grumble.
idtac " ".
idtac "------------------- split --------------------".
idtac " ".
idtac "#> split".
idtac "Possible points: 1".
check_type @split ((forall X Y : Type, list (X * Y) -> list X * list Y)).
idtac "Assumptions:".
Abort.
Print Assumptions split.
Goal True.
idtac " ".
idtac "#> test_split".
idtac "Possible points: 1".
check_type @test_split (
(@split nat bool [(1, false); (2, false)] = ([1; 2], [false; false]))).
idtac "Assumptions:".
Abort.
Print Assumptions test_split.
Goal True.
idtac " ".
idtac "------------------- filter_even_gt7 --------------------".
idtac " ".
idtac "#> test_filter_even_gt7_1".
idtac "Possible points: 1".
check_type @test_filter_even_gt7_1 (
(filter_even_gt7 [1; 2; 6; 9; 10; 3; 12; 8] = [10; 12; 8])).
idtac "Assumptions:".
Abort.
Print Assumptions test_filter_even_gt7_1.
Goal True.
idtac " ".
idtac "#> test_filter_even_gt7_2".
idtac "Possible points: 1".
check_type @test_filter_even_gt7_2 ((filter_even_gt7 [5; 2; 6; 19; 129] = [ ])).
idtac "Assumptions:".
Abort.
Print Assumptions test_filter_even_gt7_2.
Goal True.
idtac " ".
idtac "------------------- partition --------------------".
idtac " ".
idtac "#> partition".
idtac "Possible points: 1".
check_type @partition ((forall X : Type, (X -> bool) -> list X -> list X * list X)).
idtac "Assumptions:".
Abort.
Print Assumptions partition.
Goal True.
idtac " ".
idtac "#> test_partition1".
idtac "Possible points: 1".
check_type @test_partition1 ((@partition nat oddb [1; 2; 3; 4; 5] = ([1; 3; 5], [2; 4]))).
idtac "Assumptions:".
Abort.
Print Assumptions test_partition1.
Goal True.
idtac " ".
idtac "#> test_partition2".
idtac "Possible points: 1".
check_type @test_partition2 (
(@partition nat (fun _ : nat => false) [5; 9; 0] = ([ ], [5; 9; 0]))).
idtac "Assumptions:".
Abort.
Print Assumptions test_partition2.
Goal True.
idtac " ".
idtac "------------------- map_rev --------------------".
idtac " ".
idtac "#> map_rev".
idtac "Possible points: 3".
check_type @map_rev (
(forall (X Y : Type) (f : X -> Y) (l : list X),
@map X Y f (@rev X l) = @rev Y (@map X Y f l))).
idtac "Assumptions:".
Abort.
Print Assumptions map_rev.
Goal True.
idtac " ".
idtac "------------------- flat_map --------------------".
idtac " ".
idtac "#> flat_map".
idtac "Possible points: 1".
check_type @flat_map ((forall X Y : Type, (X -> list Y) -> list X -> list Y)).
idtac "Assumptions:".
Abort.
Print Assumptions flat_map.
Goal True.
idtac " ".
idtac "#> test_flat_map1".
idtac "Possible points: 1".
check_type @test_flat_map1 (
(@flat_map nat nat (fun n : nat => [n; n; n]) [1; 5; 4] =
[1; 1; 1; 5; 5; 5; 4; 4; 4])).
idtac "Assumptions:".
Abort.
Print Assumptions test_flat_map1.
Goal True.
idtac " ".
idtac "------------------- fold_types_different --------------------".
idtac " ".
idtac "#> Manually graded: fold_types_different".
idtac "Advanced".
idtac "Possible points: 1".
print_manual_grade manual_grade_for_fold_types_different.
idtac " ".
idtac "------------------- fold_length --------------------".
idtac " ".
idtac "#> Exercises.fold_length_correct".
idtac "Possible points: 2".
check_type @Exercises.fold_length_correct (
(forall (X : Type) (l : list X), @Exercises.fold_length X l = @length X l)).
idtac "Assumptions:".
Abort.
Print Assumptions Exercises.fold_length_correct.
Goal True.
idtac " ".
idtac "------------------- fold_map --------------------".
idtac " ".
idtac "#> Manually graded: Exercises.fold_map".
idtac "Possible points: 3".
print_manual_grade Exercises.manual_grade_for_fold_map.
idtac " ".
idtac "------------------- currying --------------------".
idtac " ".
idtac "#> Exercises.uncurry_curry".
idtac "Advanced".
idtac "Possible points: 1".
check_type @Exercises.uncurry_curry (
(forall (X Y Z : Type) (f : X -> Y -> Z) (x : X) (y : Y),
@Exercises.prod_curry X Y Z (@Exercises.prod_uncurry X Y Z f) x y = f x y)).
idtac "Assumptions:".
Abort.
Print Assumptions Exercises.uncurry_curry.
Goal True.
idtac " ".
idtac "#> Exercises.curry_uncurry".
idtac "Advanced".
idtac "Possible points: 1".
check_type @Exercises.curry_uncurry (
(forall (X Y Z : Type) (f : X * Y -> Z) (p : X * Y),
@Exercises.prod_uncurry X Y Z (@Exercises.prod_curry X Y Z f) p = f p)).
idtac "Assumptions:".
Abort.
Print Assumptions Exercises.curry_uncurry.
Goal True.
idtac " ".
idtac "------------------- nth_error_informal --------------------".
idtac " ".
idtac "#> Manually graded: Exercises.informal_proof".
idtac "Advanced".
idtac "Possible points: 2".
print_manual_grade Exercises.manual_grade_for_informal_proof.
idtac " ".
idtac "------------------- church_succ --------------------".
idtac " ".
idtac "#> Exercises.Church.succ_2".
idtac "Advanced".
idtac "Possible points: 0.5".
check_type @Exercises.Church.succ_2 (
(Exercises.Church.succ Exercises.Church.one = Exercises.Church.two)).
idtac "Assumptions:".
Abort.
Print Assumptions Exercises.Church.succ_2.
Goal True.
idtac " ".
idtac "#> Exercises.Church.succ_3".
idtac "Advanced".
idtac "Possible points: 0.5".
check_type @Exercises.Church.succ_3 (
(Exercises.Church.succ Exercises.Church.two = Exercises.Church.three)).
idtac "Assumptions:".
Abort.
Print Assumptions Exercises.Church.succ_3.
Goal True.
idtac " ".
idtac "------------------- church_plus --------------------".
idtac " ".
idtac "#> Exercises.Church.plus_2".
idtac "Advanced".
idtac "Possible points: 0.5".
check_type @Exercises.Church.plus_2 (
(Exercises.Church.plus Exercises.Church.two Exercises.Church.three =
Exercises.Church.plus Exercises.Church.three Exercises.Church.two)).
idtac "Assumptions:".
Abort.
Print Assumptions Exercises.Church.plus_2.
Goal True.
idtac " ".
idtac "#> Exercises.Church.plus_3".
idtac "Advanced".
idtac "Possible points: 0.5".
check_type @Exercises.Church.plus_3 (
(Exercises.Church.plus
(Exercises.Church.plus Exercises.Church.two Exercises.Church.two)
Exercises.Church.three =
Exercises.Church.plus Exercises.Church.one
(Exercises.Church.plus Exercises.Church.three Exercises.Church.three))).
idtac "Assumptions:".
Abort.
Print Assumptions Exercises.Church.plus_3.
Goal True.
idtac " ".
idtac "------------------- church_mult --------------------".
idtac " ".
idtac "#> Exercises.Church.mult_1".
idtac "Advanced".
idtac "Possible points: 0.5".
check_type @Exercises.Church.mult_1 (
(Exercises.Church.mult Exercises.Church.one Exercises.Church.one =
Exercises.Church.one)).
idtac "Assumptions:".
Abort.
Print Assumptions Exercises.Church.mult_1.
Goal True.
idtac " ".
idtac "#> Exercises.Church.mult_2".
idtac "Advanced".
idtac "Possible points: 0.5".
check_type @Exercises.Church.mult_2 (
(Exercises.Church.mult Exercises.Church.zero
(Exercises.Church.plus Exercises.Church.three Exercises.Church.three) =
Exercises.Church.zero)).
idtac "Assumptions:".
Abort.
Print Assumptions Exercises.Church.mult_2.
Goal True.
idtac " ".
idtac "#> Exercises.Church.mult_3".
idtac "Advanced".
idtac "Possible points: 1".
check_type @Exercises.Church.mult_3 (
(Exercises.Church.mult Exercises.Church.two Exercises.Church.three =
Exercises.Church.plus Exercises.Church.three Exercises.Church.three)).
idtac "Assumptions:".
Abort.
Print Assumptions Exercises.Church.mult_3.
Goal True.
idtac " ".
idtac "------------------- church_exp --------------------".
idtac " ".
idtac "#> Exercises.Church.exp_1".
idtac "Advanced".
idtac "Possible points: 0.5".
check_type @Exercises.Church.exp_1 (
(Exercises.Church.exp Exercises.Church.two Exercises.Church.two =
Exercises.Church.plus Exercises.Church.two Exercises.Church.two)).
idtac "Assumptions:".
Abort.
Print Assumptions Exercises.Church.exp_1.
Goal True.
idtac " ".
idtac "#> Exercises.Church.exp_2".
idtac "Advanced".
idtac "Possible points: 0.5".
check_type @Exercises.Church.exp_2 (
(Exercises.Church.exp Exercises.Church.three Exercises.Church.zero =
Exercises.Church.one)).
idtac "Assumptions:".
Abort.
Print Assumptions Exercises.Church.exp_2.
Goal True.
idtac " ".
idtac "#> Exercises.Church.exp_3".
idtac "Advanced".
idtac "Possible points: 1".
check_type @Exercises.Church.exp_3 (
(Exercises.Church.exp Exercises.Church.three Exercises.Church.two =
Exercises.Church.plus
(Exercises.Church.mult Exercises.Church.two
(Exercises.Church.mult Exercises.Church.two Exercises.Church.two))
Exercises.Church.one)).
idtac "Assumptions:".
Abort.
Print Assumptions Exercises.Church.exp_3.
Goal True.
idtac " ".
idtac " ".
idtac "Max points - standard: 19".
idtac "Max points - advanced: 30".
idtac "".
idtac "********** Summary **********".
idtac "".
idtac "********** Standard **********".
idtac "---------- mumble_grumble ---------".
idtac "MANUAL".
idtac "---------- split ---------".
Print Assumptions split.
idtac "---------- test_split ---------".
Print Assumptions test_split.
idtac "---------- test_filter_even_gt7_1 ---------".
Print Assumptions test_filter_even_gt7_1.
idtac "---------- test_filter_even_gt7_2 ---------".
Print Assumptions test_filter_even_gt7_2.
idtac "---------- partition ---------".
Print Assumptions partition.
idtac "---------- test_partition1 ---------".
Print Assumptions test_partition1.
idtac "---------- test_partition2 ---------".
Print Assumptions test_partition2.
idtac "---------- map_rev ---------".
Print Assumptions map_rev.
idtac "---------- flat_map ---------".
Print Assumptions flat_map.
idtac "---------- test_flat_map1 ---------".
Print Assumptions test_flat_map1.
idtac "---------- Exercises.fold_length_correct ---------".
Print Assumptions Exercises.fold_length_correct.
idtac "---------- fold_map ---------".
idtac "MANUAL".
idtac "".
idtac "********** Advanced **********".
idtac "---------- fold_types_different ---------".
idtac "MANUAL".
idtac "---------- Exercises.uncurry_curry ---------".
Print Assumptions Exercises.uncurry_curry.
idtac "---------- Exercises.curry_uncurry ---------".
Print Assumptions Exercises.curry_uncurry.
idtac "---------- informal_proof ---------".
idtac "MANUAL".
idtac "---------- Exercises.Church.succ_2 ---------".
Print Assumptions Exercises.Church.succ_2.
idtac "---------- Exercises.Church.succ_3 ---------".
Print Assumptions Exercises.Church.succ_3.
idtac "---------- Exercises.Church.plus_2 ---------".
Print Assumptions Exercises.Church.plus_2.
idtac "---------- Exercises.Church.plus_3 ---------".
Print Assumptions Exercises.Church.plus_3.
idtac "---------- Exercises.Church.mult_1 ---------".
Print Assumptions Exercises.Church.mult_1.
idtac "---------- Exercises.Church.mult_2 ---------".
Print Assumptions Exercises.Church.mult_2.
idtac "---------- Exercises.Church.mult_3 ---------".
Print Assumptions Exercises.Church.mult_3.
idtac "---------- Exercises.Church.exp_1 ---------".
Print Assumptions Exercises.Church.exp_1.
idtac "---------- Exercises.Church.exp_2 ---------".
Print Assumptions Exercises.Church.exp_2.
idtac "---------- Exercises.Church.exp_3 ---------".
Print Assumptions Exercises.Church.exp_3.
Abort.
(* Wed Jan 9 12:02:09 EST 2019 *)

91
Postscript.v Normal file
View file

@ -0,0 +1,91 @@
(** * Postscript *)
(** Congratulations: We've made it to the end! *)
(* ################################################################# *)
(** * Looking Back *)
(** We've covered quite a bit of ground so far. Here's a quick review...
- _Functional programming_:
- "declarative" programming style (recursion over immutable
data structures, rather than looping over mutable arrays
or pointer structures)
- higher-order functions
- polymorphism *)
(**
- _Logic_, the mathematical basis for software engineering:
logic calculus
-------------------- ~ ----------------------------
software engineering mechanical/civil engineering
- inductively defined sets and relations
- inductive proofs
- proof objects *)
(**
- _Coq_, an industrial-strength proof assistant
- functional core language
- core tactics
- automation
*)
(* ################################################################# *)
(** * Looking Forward *)
(** If what you've seen so far has whetted your interest, you have two
choices for further reading in the _Software Foundations_ series:
- _Programming Language Foundations_ (volume 2, by a set of
authors similar to this book's) covers material that
might be found in a graduate course on the theory of
programming languages, including Hoare logic, operational
semantics, and type systems.
- _Verified Functional Algorithms_ (volume 3, by Andrew
Appel) builds on the themes of functional programming and
program verification in Coq, addressing a range of topics
that might be found in a standard data structures course,
with an eye to formal verification. *)
(* ################################################################# *)
(** * Other sources *)
(** Here are some other good places to learn more...
- This book includes some optional chapters covering topics
that you may find useful. Take a look at the table of contents and the chapter dependency diagram to find
them.
- For questions about Coq, the [#coq] area of Stack
Overflow (https://stackoverflow.com/questions/tagged/coq)
is an excellent community resource.
- Here are some great books on functional programming
- Learn You a Haskell for Great Good, by Miran Lipovaca
[Lipovaca 2011] (in Bib.v).
- Real World Haskell, by Bryan O'Sullivan, John Goerzen,
and Don Stewart [O'Sullivan 2008] (in Bib.v)
- ...and many other excellent books on Haskell, OCaml,
Scheme, Racket, Scala, F sharp, etc., etc.
- And some further resources for Coq:
- Certified Programming with Dependent Types, by Adam
Chlipala [Chlipala 2013] (in Bib.v).
- Interactive Theorem Proving and Program Development:
Coq'Art: The Calculus of Inductive Constructions, by Yves
Bertot and Pierre Casteran [Bertot 2004] (in Bib.v).
- If you're interested in real-world applications of formal
verification to critical software, see the Postscript chapter
of _Programming Language Foundations_.
- For applications of Coq in building verified systems, the
lectures and course materials for the 2017 DeepSpec Summer
School are a great resource.
https://deepspec.org/event/dsss17/index.html
*)
(* Wed Jan 9 12:02:47 EST 2019 *)

47
PostscriptTest.v Normal file
View file

@ -0,0 +1,47 @@
Set Warnings "-notation-overridden,-parsing".
From Coq Require Export String.
From LF Require Import Postscript.
Parameter MISSING: Type.
Module Check.
Ltac check_type A B :=
match type of A with
| context[MISSING] => idtac "Missing:" A
| ?T => first [unify T B; idtac "Type: ok" | idtac "Type: wrong - should be (" B ")"]
end.
Ltac print_manual_grade A :=
match eval compute in A with
| Some (_ ?S ?C) =>
idtac "Score:" S;
match eval compute in C with
| ""%string => idtac "Comment: None"
| _ => idtac "Comment:" C
end
| None =>
idtac "Score: Ungraded";
idtac "Comment: None"
end.
End Check.
From LF Require Import Postscript.
Import Check.
Goal True.
idtac " ".
idtac "Max points - standard: 0".
idtac "Max points - advanced: 0".
idtac "".
idtac "********** Summary **********".
idtac "".
idtac "********** Standard **********".
idtac "".
idtac "********** Advanced **********".
Abort.
(* Wed Jan 9 12:02:31 EST 2019 *)

416
Preface.v Normal file
View file

@ -0,0 +1,416 @@
(** * Preface *)
(* ################################################################# *)
(** * Welcome *)
(** This is the entry point in a series of electronic textbooks on
various aspects of _Software Foundations_ -- the mathematical
underpinnings of reliable software. Topics in the series include
basic concepts of logic, computer-assisted theorem proving, the
Coq proof assistant, functional programming, operational
semantics, logics for reasoning about programs, and static type
systems. The exposition is intended for a broad range of readers,
from advanced undergraduates to PhD students and researchers. No
specific background in logic or programming languages is assumed,
though a degree of mathematical maturity will be helpful.
The principal novelty of the series is that it is one hundred
percent formalized and machine-checked: each text is literally a
script for Coq. The books are intended to be read alongside (or
inside) an interactive session with Coq. All the details in the
text are fully formalized in Coq, and most of the exercises are
designed to be worked using Coq.
The files in each book are organized into a sequence of core
chapters, covering about one semester's worth of material and
organized into a coherent linear narrative, plus a number of
"offshoot" chapters covering additional topics. All the core
chapters are suitable for both upper-level undergraduate and
graduate students.
This book, _Logical Foundations_, lays groundwork for the others,
introducing the reader to the basic ideas of functional
programming, constructive logic, and the Coq proof assistant. *)
(* ################################################################# *)
(** * Overview *)
(** Building reliable software is really hard. The scale and
complexity of modern systems, the number of people involved, and
the range of demands placed on them make it extremely difficult to
build software that is even more-or-less correct, much less 100%%
correct. At the same time, the increasing degree to which
information processing is woven into every aspect of society
greatly amplifies the cost of bugs and insecurities.
Computer scientists and software engineers have responded to these
challenges by developing a whole host of techniques for improving
software reliability, ranging from recommendations about managing
software projects teams (e.g., extreme programming) to design
philosophies for libraries (e.g., model-view-controller,
publish-subscribe, etc.) and programming languages (e.g.,
object-oriented programming, aspect-oriented programming,
functional programming, ...) to mathematical techniques for
specifying and reasoning about properties of software and tools
for helping validate these properties. The _Software Foundations_
series is focused on this last set of techniques.
The text is constructed around three conceptual threads:
(1) basic tools from _logic_ for making and justifying precise
claims about programs;
(2) the use of _proof assistants_ to construct rigorous logical
arguments;
(3) _functional programming_, both as a method of programming that
simplifies reasoning about programs and as a bridge between
programming and logic.
Some suggestions for further reading can be found in the
[Postscript] chapter. Bibliographic information for all
cited works can be found in the file [Bib]. *)
(* ================================================================= *)
(** ** Logic *)
(** Logic is the field of study whose subject matter is _proofs_ --
unassailable arguments for the truth of particular propositions.
Volumes have been written about the central role of logic in
computer science. Manna and Waldinger called it "the calculus of
computer science," while Halpern et al.'s paper _On the Unusual
Effectiveness of Logic in Computer Science_ catalogs scores of
ways in which logic offers critical tools and insights. Indeed,
they observe that, "As a matter of fact, logic has turned out to
be significantly more effective in computer science than it has
been in mathematics. This is quite remarkable, especially since
much of the impetus for the development of logic during the past
one hundred years came from mathematics."
In particular, the fundamental tools of _inductive proof_ are
ubiquitous in all of computer science. You have surely seen them
before, perhaps in a course on discrete math or analysis of
algorithms, but in this course we will examine them more deeply
than you have probably done so far. *)
(* ================================================================= *)
(** ** Proof Assistants *)
(** The flow of ideas between logic and computer science has not been
unidirectional: CS has also made important contributions to logic.
One of these has been the development of software tools for
helping construct proofs of logical propositions. These tools
fall into two broad categories:
- _Automated theorem provers_ provide "push-button" operation:
you give them a proposition and they return either _true_ or
_false_ (or, sometimes, _don't know: ran out of time_).
Although their capabilities are still limited to specific
domains, they have matured tremendously in recent years and
are used now in a multitude of settings. Examples of such
tools include SAT solvers, SMT solvers, and model checkers.
- _Proof assistants_ are hybrid tools that automate the more
routine aspects of building proofs while depending on human
guidance for more difficult aspects. Widely used proof
assistants include Isabelle, Agda, Twelf, ACL2, PVS, and Coq,
among many others.
This course is based around Coq, a proof assistant that has been
under development since 1983 and that in recent years has
attracted a large community of users in both research and
industry. Coq provides a rich environment for interactive
development of machine-checked formal reasoning. The kernel of
the Coq system is a simple proof-checker, which guarantees that
only correct deduction steps are ever performed. On top of this
kernel, the Coq environment provides high-level facilities for
proof development, including a large library of common definitions
and lemmas, powerful tactics for constructing complex proofs
semi-automatically, and a special-purpose programming language for
defining new proof-automation tactics for specific situations.
Coq has been a critical enabler for a huge variety of work across
computer science and mathematics:
- As a _platform for modeling programming languages_, it has
become a standard tool for researchers who need to describe and
reason about complex language definitions. It has been used,
for example, to check the security of the JavaCard platform,
obtaining the highest level of common criteria certification,
and for formal specifications of the x86 and LLVM instruction
sets and programming languages such as C.
- As an _environment for developing formally certified software
and hardware_, Coq has been used, for example, to build
CompCert, a fully-verified optimizing compiler for C, and
CertiKos, a fully verified hypervisor, for proving the
correctness of subtle algorithms involving floating point
numbers, and as the basis for CertiCrypt, an environment for
reasoning about the security of cryptographic algorithms. It is
also being used to build verified implementations of the
open-source RISC-V processor.
- As a _realistic environment for functional programming with
dependent types_, it has inspired numerous innovations. For
example, the Ynot system embeds "relational Hoare reasoning" (an
extension of the _Hoare Logic_ we will see later in this course)
in Coq.
- As a _proof assistant for higher-order logic_, it has been used
to validate a number of important results in mathematics. For
example, its ability to include complex computations inside
proofs made it possible to develop the first formally verified
proof of the 4-color theorem. This proof had previously been
controversial among mathematicians because part of it included
checking a large number of configurations using a program. In
the Coq formalization, everything is checked, including the
correctness of the computational part. More recently, an even
more massive effort led to a Coq formalization of the
Feit-Thompson Theorem -- the first major step in the
classification of finite simple groups.
By the way, in case you're wondering about the name, here's what
the official Coq web site at INRIA (the French national research
lab where Coq has mostly been developed) says about it: "Some
French computer scientists have a tradition of naming their
software as animal species: Caml, Elan, Foc or Phox are examples of
this tacit convention. In French, 'coq' means rooster, and it
sounds like the initials of the Calculus of Constructions (CoC) on
which it is based." The rooster is also the national symbol of
France, and C-o-q are the first three letters of the name of
Thierry Coquand, one of Coq's early developers. *)
(* ================================================================= *)
(** ** Functional Programming *)
(** The term _functional programming_ refers both to a collection of
programming idioms that can be used in almost any programming
language and to a family of programming languages designed to
emphasize these idioms, including Haskell, OCaml, Standard ML,
F##, Scala, Scheme, Racket, Common Lisp, Clojure, Erlang, and Coq.
Functional programming has been developed over many decades --
indeed, its roots go back to Church's lambda-calculus, which was
invented in the 1930s, well before the first computers (at least
the first electronic ones)! But since the early '90s it has
enjoyed a surge of interest among industrial engineers and
language designers, playing a key role in high-value systems at
companies like Jane St. Capital, Microsoft, Facebook, and
Ericsson.
The most basic tenet of functional programming is that, as much as
possible, computation should be _pure_, in the sense that the only
effect of execution should be to produce a result: it should be
free from _side effects_ such as I/O, assignments to mutable
variables, redirecting pointers, etc. For example, whereas an
_imperative_ sorting function might take a list of numbers and
rearrange its pointers to put the list in order, a pure sorting
function would take the original list and return a _new_ list
containing the same numbers in sorted order.
A significant benefit of this style of programming is that it
makes programs easier to understand and reason about. If every
operation on a data structure yields a new data structure, leaving
the old one intact, then there is no need to worry about how that
structure is being shared and whether a change by one part of the
program might break an invariant that another part of the program
relies on. These considerations are particularly critical in
concurrent systems, where every piece of mutable state that is
shared between threads is a potential source of pernicious bugs.
Indeed, a large part of the recent interest in functional
programming in industry is due to its simpler behavior in the
presence of concurrency.
Another reason for the current excitement about functional
programming is related to the first: functional programs are often
much easier to parallelize than their imperative counterparts. If
running a computation has no effect other than producing a result,
then it does not matter _where_ it is run. Similarly, if a data
structure is never modified destructively, then it can be copied
freely, across cores or across the network. Indeed, the
"Map-Reduce" idiom, which lies at the heart of massively
distributed query processors like Hadoop and is used by Google to
index the entire web is a classic example of functional
programming.
For purposes of this course, functional programming has yet
another significant attraction: it serves as a bridge between
logic and computer science. Indeed, Coq itself can be viewed as a
combination of a small but extremely expressive functional
programming language plus a set of tools for stating and proving
logical assertions. Moreover, when we come to look more closely,
we find that these two sides of Coq are actually aspects of the
very same underlying machinery -- i.e., _proofs are programs_. *)
(* ================================================================= *)
(** ** Further Reading *)
(** This text is intended to be self contained, but readers looking
for a deeper treatment of particular topics will find some
suggestions for further reading in the [Postscript]
chapter. *)
(* ################################################################# *)
(** * Practicalities *)
(* ================================================================= *)
(** ** Chapter Dependencies *)
(** A diagram of the dependencies between chapters and some suggested
paths through the material can be found in the file [deps.html]. *)
(* ================================================================= *)
(** ** System Requirements *)
(** Coq runs on Windows, Linux, and macOS. You will need:
- A current installation of Coq, available from the Coq home page.
These files have been tested with Coq 8.8.1.
- An IDE for interacting with Coq. Currently, there are two
choices:
- Proof General is an Emacs-based IDE. It tends to be
preferred by users who are already comfortable with Emacs.
It requires a separate installation (google "Proof
General").
Adventurous users of Coq within Emacs may also want to check
out extensions such as [company-coq] and [control-lock].
- CoqIDE is a simpler stand-alone IDE. It is distributed with
Coq, so it should be available once you have Coq installed.
It can also be compiled from scratch, but on some platforms
this may involve installing additional packages for GUI
libraries and such.
Users who like CoqIDE should consider running it with the
"asynchronous" and "error resilience" modes disabled:
coqide -async-proofs off -async-proofs-command-error-resilience off Foo.v &
*)
(* ================================================================= *)
(** ** Exercises *)
(** Each chapter includes numerous exercises. Each is marked with a
"star rating," which can be interpreted as follows:
- One star: easy exercises that underscore points in the text
and that, for most readers, should take only a minute or two.
Get in the habit of working these as you reach them.
- Two stars: straightforward exercises (five or ten minutes).
- Three stars: exercises requiring a bit of thought (ten
minutes to half an hour).
- Four and five stars: more difficult exercises (half an hour
and up).
Also, some exercises are marked "advanced," and some are marked
"optional." Doing just the non-optional, non-advanced exercises
should provide good coverage of the core material. Optional
exercises provide a bit of extra practice with key concepts and
introduce secondary themes that may be of interest to some
readers. Advanced exercises are for readers who want an extra
challenge and a deeper cut at the material.
_Please do not post solutions to the exercises in a public place_.
Software Foundations is widely used both for self-study and for
university courses. Having solutions easily available makes it
much less useful for courses, which typically have graded homework
assignments. We especially request that readers not post
solutions to the exercises anyplace where they can be found by
search engines. *)
(* ================================================================= *)
(** ** Downloading the Coq Files *)
(** A tar file containing the full sources for the "release version"
of this book (as a collection of Coq scripts and HTML files) is
available at http://softwarefoundations.cis.upenn.edu.
If you are using the book as part of a class, your professor may
give you access to a locally modified version of the files; you
should use this one instead of the public release version, so that
you get any local updates during the semester. *)
(* ################################################################# *)
(** * Resources *)
(* ================================================================= *)
(** ** Sample Exams *)
(** A large compendium of exams from many offerings of
CIS500 ("Software Foundations") at the University of Pennsylvania
can be found at
https://www.seas.upenn.edu/~cis500/current/exams/index.html.
There has been some drift of notations over the years, but most of
the problems are still relevant to the current text. *)
(* ================================================================= *)
(** ** Lecture Videos *)
(** Lectures for two intensive summer courses based on _Logical
Foundations_ (part of the DeepSpec summer school series) can be
found at https://deepspec.org/event/dsss17 and
https://deepspec.org/event/dsss18/. The video quality in the
2017 lectures is poor at the beginning but gets better in the
later lectures. *)
(* ################################################################# *)
(** * Note for Instructors *)
(** If you plan to use these materials in your own course, you will
undoubtedly find things you'd like to change, improve, or add.
Your contributions are welcome!
In order to keep the legalities simple and to have a single point
of responsibility in case the need should ever arise to adjust the
license terms, sublicense, etc., we ask all contributors (i.e.,
everyone with access to the developers' repository) to assign
copyright in their contributions to the appropriate "author of
record," as follows:
- I hereby assign copyright in my past and future contributions
to the Software Foundations project to the Author of Record of
each volume or component, to be licensed under the same terms
as the rest of Software Foundations. I understand that, at
present, the Authors of Record are as follows: For Volumes 1
and 2, known until 2016 as "Software Foundations" and from
2016 as (respectively) "Logical Foundations" and "Programming
Foundations," and for Volume 4, "QuickChick: Property-Based
Testing in Coq," the Author of Record is Benjamin C. Pierce.
For Volume 3, "Verified Functional Algorithms", the Author of
Record is Andrew W. Appel. For components outside of
designated volumes (e.g., typesetting and grading tools and
other software infrastructure), the Author of Record is
Benjamin Pierce.
To get started, please send an email to Benjamin Pierce,
describing yourself and how you plan to use the materials and
including (1) the above copyright transfer text and (2) your
github username.
We'll set you up with access to the git repository and developers'
mailing lists. In the repository you'll find a file [INSTRUCTORS]
with further instructions. *)
(* ################################################################# *)
(** * Translations *)
(** Thanks to the efforts of a team of volunteer translators,
_Software Foundations_ can be enjoyed in Japanese at
http://proofcafe.org/sf. A Chinese translation is also underway;
you can preview it at https://coq-zh.github.io/SF-zh/. *)
(* ################################################################# *)
(** * Thanks *)
(** Development of the _Software Foundations_ series has been
supported, in part, by the National Science Foundation under the
NSF Expeditions grant 1521523, _The Science of Deep
Specification_. *)
(* Wed Jan 9 12:02:43 EST 2019 *)

47
PrefaceTest.v Normal file
View file

@ -0,0 +1,47 @@
Set Warnings "-notation-overridden,-parsing".
From Coq Require Export String.
From LF Require Import Preface.
Parameter MISSING: Type.
Module Check.
Ltac check_type A B :=
match type of A with
| context[MISSING] => idtac "Missing:" A
| ?T => first [unify T B; idtac "Type: ok" | idtac "Type: wrong - should be (" B ")"]
end.
Ltac print_manual_grade A :=
match eval compute in A with
| Some (_ ?S ?C) =>
idtac "Score:" S;
match eval compute in C with
| ""%string => idtac "Comment: None"
| _ => idtac "Comment:" C
end
| None =>
idtac "Score: Ungraded";
idtac "Comment: None"
end.
End Check.
From LF Require Import Preface.
Import Check.
Goal True.
idtac " ".
idtac "Max points - standard: 0".
idtac "Max points - advanced: 0".
idtac "".
idtac "********** Summary **********".
idtac "".
idtac "********** Standard **********".
idtac "".
idtac "********** Advanced **********".
Abort.
(* Wed Jan 9 12:02:04 EST 2019 *)

618
ProofObjects.v Normal file
View file

@ -0,0 +1,618 @@
(** * ProofObjects: The Curry-Howard Correspondence *)
Set Warnings "-notation-overridden,-parsing".
From LF Require Export IndProp.
(** "_Algorithms are the computational content of proofs_." --Robert Harper *)
(** We have seen that Coq has mechanisms both for _programming_,
using inductive data types like [nat] or [list] and functions over
these types, and for _proving_ properties of these programs, using
inductive propositions (like [even]), implication, universal
quantification, and the like. So far, we have mostly treated
these mechanisms as if they were quite separate, and for many
purposes this is a good way to think. But we have also seen hints
that Coq's programming and proving facilities are closely related.
For example, the keyword [Inductive] is used to declare both data
types and propositions, and [->] is used both to describe the type
of functions on data and logical implication. This is not just a
syntactic accident! In fact, programs and proofs in Coq are
almost the same thing. In this chapter we will study how this
works.
We have already seen the fundamental idea: provability in Coq is
represented by concrete _evidence_. When we construct the proof
of a basic proposition, we are actually building a tree of
evidence, which can be thought of as a data structure.
If the proposition is an implication like [A -> B], then its proof
will be an evidence _transformer_: a recipe for converting
evidence for A into evidence for B. So at a fundamental level,
proofs are simply programs that manipulate evidence. *)
(** Question: If evidence is data, what are propositions themselves?
Answer: They are types! *)
(** Look again at the formal definition of the [even] property. *)
Print even.
(* ==>
Inductive even : nat -> Prop :=
| ev_0 : even 0
| ev_SS : forall n, even n -> even (S (S n)).
*)
(** Suppose we introduce an alternative pronunciation of "[:]".
Instead of "has type," we can say "is a proof of." For example,
the second line in the definition of [even] declares that [ev_0 : even
0]. Instead of "[ev_0] has type [even 0]," we can say that "[ev_0]
is a proof of [even 0]." *)
(** This pun between types and propositions -- between [:] as "has type"
and [:] as "is a proof of" or "is evidence for" -- is called the
_Curry-Howard correspondence_. It proposes a deep connection
between the world of logic and the world of computation:
propositions ~ types
proofs ~ data values
See [Wadler 2015] (in Bib.v) for a brief history and up-to-date exposition. *)
(** Many useful insights follow from this connection. To begin with,
it gives us a natural interpretation of the type of the [ev_SS]
constructor: *)
Check ev_SS.
(* ===> ev_SS : forall n,
even n ->
even (S (S n)) *)
(** This can be read "[ev_SS] is a constructor that takes two
arguments -- a number [n] and evidence for the proposition [even
n] -- and yields evidence for the proposition [even (S (S n))]." *)
(** Now let's look again at a previous proof involving [even]. *)
Theorem ev_4 : even 4.
Proof.
apply ev_SS. apply ev_SS. apply ev_0. Qed.
(** As with ordinary data values and functions, we can use the [Print]
command to see the _proof object_ that results from this proof
script. *)
Print ev_4.
(* ===> ev_4 = ev_SS 2 (ev_SS 0 ev_0)
: even 4 *)
(** Indeed, we can also write down this proof object _directly_,
without the need for a separate proof script: *)
Check (ev_SS 2 (ev_SS 0 ev_0)).
(* ===> even 4 *)
(** The expression [ev_SS 2 (ev_SS 0 ev_0)] can be thought of as
instantiating the parameterized constructor [ev_SS] with the
specific arguments [2] and [0] plus the corresponding proof
objects for its premises [even 2] and [even 0]. Alternatively, we can
think of [ev_SS] as a primitive "evidence constructor" that, when
applied to a particular number, wants to be further applied to
evidence that that number is even; its type,
forall n, even n -> even (S (S n)),
expresses this functionality, in the same way that the polymorphic
type [forall X, list X] expresses the fact that the constructor
[nil] can be thought of as a function from types to empty lists
with elements of that type. *)
(** We saw in the [Logic] chapter that we can use function
application syntax to instantiate universally quantified variables
in lemmas, as well as to supply evidence for assumptions that
these lemmas impose. For instance: *)
Theorem ev_4': even 4.
Proof.
apply (ev_SS 2 (ev_SS 0 ev_0)).
Qed.
(* ################################################################# *)
(** * Proof Scripts *)
(** The _proof objects_ we've been discussing lie at the core of how
Coq operates. When Coq is following a proof script, what is
happening internally is that it is gradually constructing a proof
object -- a term whose type is the proposition being proved. The
tactics between [Proof] and [Qed] tell it how to build up a term
of the required type. To see this process in action, let's use
the [Show Proof] command to display the current state of the proof
tree at various points in the following tactic proof. *)
Theorem ev_4'' : even 4.
Proof.
Show Proof.
apply ev_SS.
Show Proof.
apply ev_SS.
Show Proof.
apply ev_0.
Show Proof.
Qed.
(** At any given moment, Coq has constructed a term with a
"hole" (indicated by [?Goal] here, and so on), and it knows what
type of evidence is needed to fill this hole.
Each hole corresponds to a subgoal, and the proof is
finished when there are no more subgoals. At this point, the
evidence we've built stored in the global context under the name
given in the [Theorem] command. *)
(** Tactic proofs are useful and convenient, but they are not
essential: in principle, we can always construct the required
evidence by hand, as shown above. Then we can use [Definition]
(rather than [Theorem]) to give a global name directly to this
evidence. *)
Definition ev_4''' : even 4 :=
ev_SS 2 (ev_SS 0 ev_0).
(** All these different ways of building the proof lead to exactly the
same evidence being saved in the global environment. *)
Print ev_4.
(* ===> ev_4 = ev_SS 2 (ev_SS 0 ev_0) : even 4 *)
Print ev_4'.
(* ===> ev_4' = ev_SS 2 (ev_SS 0 ev_0) : even 4 *)
Print ev_4''.
(* ===> ev_4'' = ev_SS 2 (ev_SS 0 ev_0) : even 4 *)
Print ev_4'''.
(* ===> ev_4''' = ev_SS 2 (ev_SS 0 ev_0) : even 4 *)
(** **** Exercise: 2 stars, standard (eight_is_even)
Give a tactic proof and a proof object showing that [even 8]. *)
Theorem ev_8 : even 8.
Proof.
(* FILL IN HERE *) Admitted.
Definition ev_8' : even 8
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
(** [] *)
(* ################################################################# *)
(** * Quantifiers, Implications, Functions *)
(** In Coq's computational universe (where data structures and
programs live), there are two sorts of values with arrows in their
types: _constructors_ introduced by [Inductive]ly defined data
types, and _functions_.
Similarly, in Coq's logical universe (where we carry out proofs),
there are two ways of giving evidence for an implication:
constructors introduced by [Inductive]ly defined propositions,
and... functions! *)
(** For example, consider this statement: *)
Theorem ev_plus4 : forall n, even n -> even (4 + n).
Proof.
intros n H. simpl.
apply ev_SS.
apply ev_SS.
apply H.
Qed.
(** What is the proof object corresponding to [ev_plus4]?
We're looking for an expression whose _type_ is [forall n, even n ->
even (4 + n)] -- that is, a _function_ that takes two arguments (one
number and a piece of evidence) and returns a piece of evidence!
Here it is: *)
Definition ev_plus4' : forall n, even n -> even (4 + n) :=
fun (n : nat) => fun (H : even n) =>
ev_SS (S (S n)) (ev_SS n H).
(** Recall that [fun n => blah] means "the function that, given [n],
yields [blah]," and that Coq treats [4 + n] and [S (S (S (S n)))]
as synonyms. Another equivalent way to write this definition is: *)
Definition ev_plus4'' (n : nat) (H : even n)
: even (4 + n) :=
ev_SS (S (S n)) (ev_SS n H).
Check ev_plus4''.
(* ===>
: forall n : nat, even n -> even (4 + n) *)
(** When we view the proposition being proved by [ev_plus4] as a
function type, one interesting point becomes apparent: The second
argument's type, [even n], mentions the _value_ of the first
argument, [n].
While such _dependent types_ are not found in conventional
programming languages, they can be useful in programming too, as
the recent flurry of activity in the functional programming
community demonstrates. *)
(** Notice that both implication ([->]) and quantification ([forall])
correspond to functions on evidence. In fact, they are really the
same thing: [->] is just a shorthand for a degenerate use of
[forall] where there is no dependency, i.e., no need to give a
name to the type on the left-hand side of the arrow:
forall (x:nat), nat
= forall (_:nat), nat
= nat -> nat
*)
(** For example, consider this proposition: *)
Definition ev_plus2 : Prop :=
forall n, forall (E : even n), even (n + 2).
(** A proof term inhabiting this proposition would be a function
with two arguments: a number [n] and some evidence [E] that [n] is
even. But the name [E] for this evidence is not used in the rest
of the statement of [ev_plus2], so it's a bit silly to bother
making up a name for it. We could write it like this instead,
using the dummy identifier [_] in place of a real name: *)
Definition ev_plus2' : Prop :=
forall n, forall (_ : even n), even (n + 2).
(** Or, equivalently, we can write it in more familiar notation: *)
Definition ev_plus2'' : Prop :=
forall n, even n -> even (n + 2).
(** In general, "[P -> Q]" is just syntactic sugar for
"[forall (_:P), Q]". *)
(* ################################################################# *)
(** * Programming with Tactics *)
(** If we can build proofs by giving explicit terms rather than
executing tactic scripts, you may be wondering whether we can
build _programs_ using _tactics_ rather than explicit terms.
Naturally, the answer is yes! *)
Definition add1 : nat -> nat.
intro n.
Show Proof.
apply S.
Show Proof.
apply n. Defined.
Print add1.
(* ==>
add1 = fun n : nat => S n
: nat -> nat
*)
Compute add1 2.
(* ==> 3 : nat *)
(** Notice that we terminate the [Definition] with a [.] rather than
with [:=] followed by a term. This tells Coq to enter _proof
scripting mode_ to build an object of type [nat -> nat]. Also, we
terminate the proof with [Defined] rather than [Qed]; this makes
the definition _transparent_ so that it can be used in computation
like a normally-defined function. ([Qed]-defined objects are
opaque during computation.)
This feature is mainly useful for writing functions with dependent
types, which we won't explore much further in this book. But it
does illustrate the uniformity and orthogonality of the basic
ideas in Coq. *)
(* ################################################################# *)
(** * Logical Connectives as Inductive Types *)
(** Inductive definitions are powerful enough to express most of the
connectives we have seen so far. Indeed, only universal
quantification (with implication as a special case) is built into
Coq; all the others are defined inductively. We'll see these
definitions in this section. *)
Module Props.
(* ================================================================= *)
(** ** Conjunction *)
(** To prove that [P /\ Q] holds, we must present evidence for both
[P] and [Q]. Thus, it makes sense to define a proof object for [P
/\ Q] as consisting of a pair of two proofs: one for [P] and
another one for [Q]. This leads to the following definition. *)
Module And.
Inductive and (P Q : Prop) : Prop :=
| conj : P -> Q -> and P Q.
End And.
(** Notice the similarity with the definition of the [prod] type,
given in chapter [Poly]; the only difference is that [prod] takes
[Type] arguments, whereas [and] takes [Prop] arguments. *)
Print prod.
(* ===>
Inductive prod (X Y : Type) : Type :=
| pair : X -> Y -> X * Y. *)
(** This similarity should clarify why [destruct] and [intros]
patterns can be used on a conjunctive hypothesis. Case analysis
allows us to consider all possible ways in which [P /\ Q] was
proved -- here just one (the [conj] constructor).
Similarly, the [split] tactic actually works for any inductively
defined proposition with exactly one constructor. In particular,
it works for [and]: *)
Lemma and_comm : forall P Q : Prop, P /\ Q <-> Q /\ P.
Proof.
intros P Q. split.
- intros [HP HQ]. split.
+ apply HQ.
+ apply HP.
- intros [HP HQ]. split.
+ apply HQ.
+ apply HP.
Qed.
(** This shows why the inductive definition of [and] can be
manipulated by tactics as we've been doing. We can also use it to
build proofs directly, using pattern-matching. For instance: *)
Definition and_comm'_aux P Q (H : P /\ Q) : Q /\ P :=
match H with
| conj HP HQ => conj HQ HP
end.
Definition and_comm' P Q : P /\ Q <-> Q /\ P :=
conj (and_comm'_aux P Q) (and_comm'_aux Q P).
(** **** Exercise: 2 stars, standard, optional (conj_fact)
Construct a proof object demonstrating the following proposition. *)
Definition conj_fact : forall P Q R, P /\ Q -> Q /\ R -> P /\ R
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
(** [] *)
(* ================================================================= *)
(** ** Disjunction *)
(** The inductive definition of disjunction uses two constructors, one
for each side of the disjunct: *)
Module Or.
Inductive or (P Q : Prop) : Prop :=
| or_introl : P -> or P Q
| or_intror : Q -> or P Q.
End Or.
(** This declaration explains the behavior of the [destruct] tactic on
a disjunctive hypothesis, since the generated subgoals match the
shape of the [or_introl] and [or_intror] constructors.
Once again, we can also directly write proof objects for theorems
involving [or], without resorting to tactics. *)
(** **** Exercise: 2 stars, standard, optional (or_commut'')
Try to write down an explicit proof object for [or_commut] (without
using [Print] to peek at the ones we already defined!). *)
Definition or_comm : forall P Q, P \/ Q -> Q \/ P
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
(** [] *)
(* ================================================================= *)
(** ** Existential Quantification *)
(** To give evidence for an existential quantifier, we package a
witness [x] together with a proof that [x] satisfies the property
[P]: *)
Module Ex.
Inductive ex {A : Type} (P : A -> Prop) : Prop :=
| ex_intro : forall x : A, P x -> ex P.
End Ex.
(** This may benefit from a little unpacking. The core definition is
for a type former [ex] that can be used to build propositions of
the form [ex P], where [P] itself is a _function_ from witness
values in the type [A] to propositions. The [ex_intro]
constructor then offers a way of constructing evidence for [ex P],
given a witness [x] and a proof of [P x]. *)
(** The more familiar form [exists x, P x] desugars to an expression
involving [ex]: *)
Check ex (fun n => even n).
(* ===> exists n : nat, even n
: Prop *)
(** Here's how to define an explicit proof object involving [ex]: *)
Definition some_nat_is_even : exists n, even n :=
ex_intro even 4 (ev_SS 2 (ev_SS 0 ev_0)).
(** **** Exercise: 2 stars, standard, optional (ex_ev_Sn)
Complete the definition of the following proof object: *)
Definition ex_ev_Sn : ex (fun n => even (S n))
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
(** [] *)
(* ================================================================= *)
(** ** [True] and [False] *)
(** The inductive definition of the [True] proposition is simple: *)
Inductive True : Prop :=
| I : True.
(** It has one constructor (so every proof of [True] is the same, so
being given a proof of [True] is not informative.) *)
(** [False] is equally simple -- indeed, so simple it may look
syntactically wrong at first glance! *)
Inductive False : Prop := .
(** That is, [False] is an inductive type with _no_ constructors --
i.e., no way to build evidence for it. *)
End Props.
(* ################################################################# *)
(** * Equality *)
(** Even Coq's equality relation is not built in. It has the
following inductive definition. (Actually, the definition in the
standard library is a slight variant of this, which gives an
induction principle that is slightly easier to use.) *)
Module MyEquality.
Inductive eq {X:Type} : X -> X -> Prop :=
| eq_refl : forall x, eq x x.
Notation "x == y" := (eq x y)
(at level 70, no associativity)
: type_scope.
(** The way to think about this definition is that, given a set [X],
it defines a _family_ of propositions "[x] is equal to [y],"
indexed by pairs of values ([x] and [y]) from [X]. There is just
one way of constructing evidence for members of this family:
applying the constructor [eq_refl] to a type [X] and a single
value [x : X], which yields evidence that [x] is equal to [x].
Other types of the form [eq x y] where [x] and [y] are not the
same are thus uninhabited. *)
(** We can use [eq_refl] to construct evidence that, for example, [2 =
2]. Can we also use it to construct evidence that [1 + 1 = 2]?
Yes, we can. Indeed, it is the very same piece of evidence!
The reason is that Coq treats as "the same" any two terms that are
_convertible_ according to a simple set of computation rules.
These rules, which are similar to those used by [Compute], include
evaluation of function application, inlining of definitions, and
simplification of [match]es. *)
Lemma four: 2 + 2 == 1 + 3.
Proof.
apply eq_refl.
Qed.
(** The [reflexivity] tactic that we have used to prove equalities up
to now is essentially just shorthand for [apply eq_refl].
In tactic-based proofs of equality, the conversion rules are
normally hidden in uses of [simpl] (either explicit or implicit in
other tactics such as [reflexivity]).
But you can see them directly at work in the following explicit
proof objects: *)
Definition four' : 2 + 2 == 1 + 3 :=
eq_refl 4.
Definition singleton : forall (X:Type) (x:X), []++[x] == x::[] :=
fun (X:Type) (x:X) => eq_refl [x].
(** **** Exercise: 2 stars, standard (equality__leibniz_equality)
The inductive definition of equality implies _Leibniz equality_:
what we mean when we say "[x] and [y] are equal" is that every
property on [P] that is true of [x] is also true of [y]. *)
Lemma equality__leibniz_equality : forall (X : Type) (x y: X),
x == y -> forall P:X->Prop, P x -> P y.
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** **** Exercise: 5 stars, standard, optional (leibniz_equality__equality)
Show that, in fact, the inductive definition of equality is
_equivalent_ to Leibniz equality: *)
Lemma leibniz_equality__equality : forall (X : Type) (x y: X),
(forall P:X->Prop, P x -> P y) -> x == y.
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
End MyEquality.
(* ================================================================= *)
(** ** Inversion, Again *)
(** We've seen [inversion] used with both equality hypotheses and
hypotheses about inductively defined propositions. Now that we've
seen that these are actually the same thing, we're in a position
to take a closer look at how [inversion] behaves.
In general, the [inversion] tactic...
- takes a hypothesis [H] whose type [P] is inductively defined,
and
- for each constructor [C] in [P]'s definition,
- generates a new subgoal in which we assume [H] was
built with [C],
- adds the arguments (premises) of [C] to the context of
the subgoal as extra hypotheses,
- matches the conclusion (result type) of [C] against the
current goal and calculates a set of equalities that must
hold in order for [C] to be applicable,
- adds these equalities to the context (and, for convenience,
rewrites them in the goal), and
- if the equalities are not satisfiable (e.g., they involve
things like [S n = O]), immediately solves the subgoal. *)
(** _Example_: If we invert a hypothesis built with [or], there are
two constructors, so two subgoals get generated. The
conclusion (result type) of the constructor ([P \/ Q]) doesn't
place any restrictions on the form of [P] or [Q], so we don't get
any extra equalities in the context of the subgoal. *)
(** _Example_: If we invert a hypothesis built with [and], there is
only one constructor, so only one subgoal gets generated. Again,
the conclusion (result type) of the constructor ([P /\ Q]) doesn't
place any restrictions on the form of [P] or [Q], so we don't get
any extra equalities in the context of the subgoal. The
constructor does have two arguments, though, and these can be seen
in the context in the subgoal. *)
(** _Example_: If we invert a hypothesis built with [eq], there is
again only one constructor, so only one subgoal gets generated.
Now, though, the form of the [eq_refl] constructor does give us
some extra information: it tells us that the two arguments to [eq]
must be the same! The [inversion] tactic adds this fact to the
context. *)
(* Wed Jan 9 12:02:45 EST 2019 *)

88
ProofObjectsTest.v Normal file
View file

@ -0,0 +1,88 @@
Set Warnings "-notation-overridden,-parsing".
From Coq Require Export String.
From LF Require Import ProofObjects.
Parameter MISSING: Type.
Module Check.
Ltac check_type A B :=
match type of A with
| context[MISSING] => idtac "Missing:" A
| ?T => first [unify T B; idtac "Type: ok" | idtac "Type: wrong - should be (" B ")"]
end.
Ltac print_manual_grade A :=
match eval compute in A with
| Some (_ ?S ?C) =>
idtac "Score:" S;
match eval compute in C with
| ""%string => idtac "Comment: None"
| _ => idtac "Comment:" C
end
| None =>
idtac "Score: Ungraded";
idtac "Comment: None"
end.
End Check.
From LF Require Import ProofObjects.
Import Check.
Goal True.
idtac "------------------- eight_is_even --------------------".
idtac " ".
idtac "#> ev_8".
idtac "Possible points: 1".
check_type @ev_8 ((even 8)).
idtac "Assumptions:".
Abort.
Print Assumptions ev_8.
Goal True.
idtac " ".
idtac "#> ev_8'".
idtac "Possible points: 1".
check_type @ev_8' ((even 8)).
idtac "Assumptions:".
Abort.
Print Assumptions ev_8'.
Goal True.
idtac " ".
idtac "------------------- equality__leibniz_equality --------------------".
idtac " ".
idtac "#> MyEquality.equality__leibniz_equality".
idtac "Possible points: 2".
check_type @MyEquality.equality__leibniz_equality (
(forall (X : Type) (x y : X),
@MyEquality.eq X x y -> forall P : X -> Prop, P x -> P y)).
idtac "Assumptions:".
Abort.
Print Assumptions MyEquality.equality__leibniz_equality.
Goal True.
idtac " ".
idtac " ".
idtac "Max points - standard: 4".
idtac "Max points - advanced: 4".
idtac "".
idtac "********** Summary **********".
idtac "".
idtac "********** Standard **********".
idtac "---------- ev_8 ---------".
Print Assumptions ev_8.
idtac "---------- ev_8' ---------".
Print Assumptions ev_8'.
idtac "---------- MyEquality.equality__leibniz_equality ---------".
Print Assumptions MyEquality.equality__leibniz_equality.
idtac "".
idtac "********** Advanced **********".
Abort.
(* Wed Jan 9 12:02:18 EST 2019 *)

397
Rel.v Normal file
View file

@ -0,0 +1,397 @@
(** * Rel: Properties of Relations *)
(** This short (and optional) chapter develops some basic definitions
and a few theorems about binary relations in Coq. The key
definitions are repeated where they are actually used (in the
[Smallstep] chapter of _Programming Language Foundations_),
so readers who are already comfortable with these ideas can safely
skim or skip this chapter. However, relations are also a good
source of exercises for developing facility with Coq's basic
reasoning facilities, so it may be useful to look at this material
just after the [IndProp] chapter. *)
Set Warnings "-notation-overridden,-parsing".
From LF Require Export IndProp.
(* ################################################################# *)
(** * Relations *)
(** A binary _relation_ on a set [X] is a family of propositions
parameterized by two elements of [X] -- i.e., a proposition about
pairs of elements of [X]. *)
Definition relation (X: Type) := X -> X -> Prop.
(** Confusingly, the Coq standard library hijacks the generic term
"relation" for this specific instance of the idea. To maintain
consistency with the library, we will do the same. So, henceforth
the Coq identifier [relation] will always refer to a binary
relation between some set and itself, whereas the English word
"relation" can refer either to the specific Coq concept or the
more general concept of a relation between any number of possibly
different sets. The context of the discussion should always make
clear which is meant. *)
(** An example relation on [nat] is [le], the less-than-or-equal-to
relation, which we usually write [n1 <= n2]. *)
Print le.
(* ====> Inductive le (n : nat) : nat -> Prop :=
le_n : n <= n
| le_S : forall m : nat, n <= m -> n <= S m *)
Check le : nat -> nat -> Prop.
Check le : relation nat.
(** (Why did we write it this way instead of starting with [Inductive
le : relation nat...]? Because we wanted to put the first [nat]
to the left of the [:], which makes Coq generate a somewhat nicer
induction principle for reasoning about [<=].) *)
(* ################################################################# *)
(** * Basic Properties *)
(** As anyone knows who has taken an undergraduate discrete math
course, there is a lot to be said about relations in general,
including ways of classifying relations (as reflexive, transitive,
etc.), theorems that can be proved generically about certain sorts
of relations, constructions that build one relation from another,
etc. For example... *)
(* ----------------------------------------------------------------- *)
(** *** Partial Functions *)
(** A relation [R] on a set [X] is a _partial function_ if, for every
[x], there is at most one [y] such that [R x y] -- i.e., [R x y1]
and [R x y2] together imply [y1 = y2]. *)
Definition partial_function {X: Type} (R: relation X) :=
forall x y1 y2 : X, R x y1 -> R x y2 -> y1 = y2.
(** For example, the [next_nat] relation defined earlier is a partial
function. *)
Print next_nat.
(* ====> Inductive next_nat (n : nat) : nat -> Prop :=
nn : next_nat n (S n) *)
Check next_nat : relation nat.
Theorem next_nat_partial_function :
partial_function next_nat.
Proof.
unfold partial_function.
intros x y1 y2 H1 H2.
inversion H1. inversion H2.
reflexivity. Qed.
(** However, the [<=] relation on numbers is not a partial
function. (Assume, for a contradiction, that [<=] is a partial
function. But then, since [0 <= 0] and [0 <= 1], it follows that
[0 = 1]. This is nonsense, so our assumption was
contradictory.) *)
Theorem le_not_a_partial_function :
~ (partial_function le).
Proof.
unfold not. unfold partial_function. intros Hc.
assert (0 = 1) as Nonsense. {
apply Hc with (x := 0).
- apply le_n.
- apply le_S. apply le_n. }
discriminate Nonsense. Qed.
(** **** Exercise: 2 stars, standard, optional (total_relation_not_partial)
Show that the [total_relation] defined in (an exercise in)
[IndProp] is not a partial function. *)
(* FILL IN HERE
[] *)
(** **** Exercise: 2 stars, standard, optional (empty_relation_partial)
Show that the [empty_relation] defined in (an exercise in)
[IndProp] is a partial function. *)
(* FILL IN HERE
[] *)
(* ----------------------------------------------------------------- *)
(** *** Reflexive Relations *)
(** A _reflexive_ relation on a set [X] is one for which every element
of [X] is related to itself. *)
Definition reflexive {X: Type} (R: relation X) :=
forall a : X, R a a.
Theorem le_reflexive :
reflexive le.
Proof.
unfold reflexive. intros n. apply le_n. Qed.
(* ----------------------------------------------------------------- *)
(** *** Transitive Relations *)
(** A relation [R] is _transitive_ if [R a c] holds whenever [R a b]
and [R b c] do. *)
Definition transitive {X: Type} (R: relation X) :=
forall a b c : X, (R a b) -> (R b c) -> (R a c).
Theorem le_trans :
transitive le.
Proof.
intros n m o Hnm Hmo.
induction Hmo.
- (* le_n *) apply Hnm.
- (* le_S *) apply le_S. apply IHHmo. Qed.
Theorem lt_trans:
transitive lt.
Proof.
unfold lt. unfold transitive.
intros n m o Hnm Hmo.
apply le_S in Hnm.
apply le_trans with (a := (S n)) (b := (S m)) (c := o).
apply Hnm.
apply Hmo. Qed.
(** **** Exercise: 2 stars, standard, optional (le_trans_hard_way)
We can also prove [lt_trans] more laboriously by induction,
without using [le_trans]. Do this. *)
Theorem lt_trans' :
transitive lt.
Proof.
(* Prove this by induction on evidence that [m] is less than [o]. *)
unfold lt. unfold transitive.
intros n m o Hnm Hmo.
induction Hmo as [| m' Hm'o].
(* FILL IN HERE *) Admitted.
(** [] *)
(** **** Exercise: 2 stars, standard, optional (lt_trans'')
Prove the same thing again by induction on [o]. *)
Theorem lt_trans'' :
transitive lt.
Proof.
unfold lt. unfold transitive.
intros n m o Hnm Hmo.
induction o as [| o'].
(* FILL IN HERE *) Admitted.
(** [] *)
(** The transitivity of [le], in turn, can be used to prove some facts
that will be useful later (e.g., for the proof of antisymmetry
below)... *)
Theorem le_Sn_le : forall n m, S n <= m -> n <= m.
Proof.
intros n m H. apply le_trans with (S n).
- apply le_S. apply le_n.
- apply H.
Qed.
(** **** Exercise: 1 star, standard, optional (le_S_n) *)
Theorem le_S_n : forall n m,
(S n <= S m) -> (n <= m).
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** **** Exercise: 2 stars, standard, optional (le_Sn_n_inf)
Provide an informal proof of the following theorem:
Theorem: For every [n], [~ (S n <= n)]
A formal proof of this is an optional exercise below, but try
writing an informal proof without doing the formal proof first.
Proof: *)
(* FILL IN HERE
[] *)
(** **** Exercise: 1 star, standard, optional (le_Sn_n) *)
Theorem le_Sn_n : forall n,
~ (S n <= n).
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** Reflexivity and transitivity are the main concepts we'll need for
later chapters, but, for a bit of additional practice working with
relations in Coq, let's look at a few other common ones... *)
(* ----------------------------------------------------------------- *)
(** *** Symmetric and Antisymmetric Relations *)
(** A relation [R] is _symmetric_ if [R a b] implies [R b a]. *)
Definition symmetric {X: Type} (R: relation X) :=
forall a b : X, (R a b) -> (R b a).
(** **** Exercise: 2 stars, standard, optional (le_not_symmetric) *)
Theorem le_not_symmetric :
~ (symmetric le).
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** A relation [R] is _antisymmetric_ if [R a b] and [R b a] together
imply [a = b] -- that is, if the only "cycles" in [R] are trivial
ones. *)
Definition antisymmetric {X: Type} (R: relation X) :=
forall a b : X, (R a b) -> (R b a) -> a = b.
(** **** Exercise: 2 stars, standard, optional (le_antisymmetric) *)
Theorem le_antisymmetric :
antisymmetric le.
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** **** Exercise: 2 stars, standard, optional (le_step) *)
Theorem le_step : forall n m p,
n < m ->
m <= S p ->
n <= p.
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(* ----------------------------------------------------------------- *)
(** *** Equivalence Relations *)
(** A relation is an _equivalence_ if it's reflexive, symmetric, and
transitive. *)
Definition equivalence {X:Type} (R: relation X) :=
(reflexive R) /\ (symmetric R) /\ (transitive R).
(* ----------------------------------------------------------------- *)
(** *** Partial Orders and Preorders *)
(** A relation is a _partial order_ when it's reflexive,
_anti_-symmetric, and transitive. In the Coq standard library
it's called just "order" for short. *)
Definition order {X:Type} (R: relation X) :=
(reflexive R) /\ (antisymmetric R) /\ (transitive R).
(** A preorder is almost like a partial order, but doesn't have to be
antisymmetric. *)
Definition preorder {X:Type} (R: relation X) :=
(reflexive R) /\ (transitive R).
Theorem le_order :
order le.
Proof.
unfold order. split.
- (* refl *) apply le_reflexive.
- split.
+ (* antisym *) apply le_antisymmetric.
+ (* transitive. *) apply le_trans. Qed.
(* ################################################################# *)
(** * Reflexive, Transitive Closure *)
(** The _reflexive, transitive closure_ of a relation [R] is the
smallest relation that contains [R] and that is both reflexive and
transitive. Formally, it is defined like this in the Relations
module of the Coq standard library: *)
Inductive clos_refl_trans {A: Type} (R: relation A) : relation A :=
| rt_step x y (H : R x y) : clos_refl_trans R x y
| rt_refl x : clos_refl_trans R x x
| rt_trans x y z
(Hxy : clos_refl_trans R x y)
(Hyz : clos_refl_trans R y z) :
clos_refl_trans R x z.
(** For example, the reflexive and transitive closure of the
[next_nat] relation coincides with the [le] relation. *)
Theorem next_nat_closure_is_le : forall n m,
(n <= m) <-> ((clos_refl_trans next_nat) n m).
Proof.
intros n m. split.
- (* -> *)
intro H. induction H.
+ (* le_n *) apply rt_refl.
+ (* le_S *)
apply rt_trans with m. apply IHle. apply rt_step.
apply nn.
- (* <- *)
intro H. induction H.
+ (* rt_step *) inversion H. apply le_S. apply le_n.
+ (* rt_refl *) apply le_n.
+ (* rt_trans *)
apply le_trans with y.
apply IHclos_refl_trans1.
apply IHclos_refl_trans2. Qed.
(** The above definition of reflexive, transitive closure is natural:
it says, explicitly, that the reflexive and transitive closure of
[R] is the least relation that includes [R] and that is closed
under rules of reflexivity and transitivity. But it turns out
that this definition is not very convenient for doing proofs,
since the "nondeterminism" of the [rt_trans] rule can sometimes
lead to tricky inductions. Here is a more useful definition: *)
Inductive clos_refl_trans_1n {A : Type}
(R : relation A) (x : A)
: A -> Prop :=
| rt1n_refl : clos_refl_trans_1n R x x
| rt1n_trans (y z : A)
(Hxy : R x y) (Hrest : clos_refl_trans_1n R y z) :
clos_refl_trans_1n R x z.
(** Our new definition of reflexive, transitive closure "bundles"
the [rt_step] and [rt_trans] rules into the single rule step.
The left-hand premise of this step is a single use of [R],
leading to a much simpler induction principle.
Before we go on, we should check that the two definitions do
indeed define the same relation...
First, we prove two lemmas showing that [clos_refl_trans_1n] mimics
the behavior of the two "missing" [clos_refl_trans]
constructors. *)
Lemma rsc_R : forall (X:Type) (R:relation X) (x y : X),
R x y -> clos_refl_trans_1n R x y.
Proof.
intros X R x y H.
apply rt1n_trans with y. apply H. apply rt1n_refl. Qed.
(** **** Exercise: 2 stars, standard, optional (rsc_trans) *)
Lemma rsc_trans :
forall (X:Type) (R: relation X) (x y z : X),
clos_refl_trans_1n R x y ->
clos_refl_trans_1n R y z ->
clos_refl_trans_1n R x z.
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** Then we use these facts to prove that the two definitions of
reflexive, transitive closure do indeed define the same
relation. *)
(** **** Exercise: 3 stars, standard, optional (rtc_rsc_coincide) *)
Theorem rtc_rsc_coincide :
forall (X:Type) (R: relation X) (x y : X),
clos_refl_trans R x y <-> clos_refl_trans_1n R x y.
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(* Wed Jan 9 12:02:46 EST 2019 *)

47
RelTest.v Normal file
View file

@ -0,0 +1,47 @@
Set Warnings "-notation-overridden,-parsing".
From Coq Require Export String.
From LF Require Import Rel.
Parameter MISSING: Type.
Module Check.
Ltac check_type A B :=
match type of A with
| context[MISSING] => idtac "Missing:" A
| ?T => first [unify T B; idtac "Type: ok" | idtac "Type: wrong - should be (" B ")"]
end.
Ltac print_manual_grade A :=
match eval compute in A with
| Some (_ ?S ?C) =>
idtac "Score:" S;
match eval compute in C with
| ""%string => idtac "Comment: None"
| _ => idtac "Comment:" C
end
| None =>
idtac "Score: Ungraded";
idtac "Comment: None"
end.
End Check.
From LF Require Import Rel.
Import Check.
Goal True.
idtac " ".
idtac "Max points - standard: 0".
idtac "Max points - advanced: 0".
idtac "".
idtac "********** Summary **********".
idtac "".
idtac "********** Standard **********".
idtac "".
idtac "********** Advanced **********".
Abort.
(* Wed Jan 9 12:02:21 EST 2019 *)

1183
Tactics.v Normal file

File diff suppressed because it is too large Load diff

222
TacticsTest.v Normal file
View file

@ -0,0 +1,222 @@
Set Warnings "-notation-overridden,-parsing".
From Coq Require Export String.
From LF Require Import Tactics.
Parameter MISSING: Type.
Module Check.
Ltac check_type A B :=
match type of A with
| context[MISSING] => idtac "Missing:" A
| ?T => first [unify T B; idtac "Type: ok" | idtac "Type: wrong - should be (" B ")"]
end.
Ltac print_manual_grade A :=
match eval compute in A with
| Some (_ ?S ?C) =>
idtac "Score:" S;
match eval compute in C with
| ""%string => idtac "Comment: None"
| _ => idtac "Comment:" C
end
| None =>
idtac "Score: Ungraded";
idtac "Comment: None"
end.
End Check.
From LF Require Import Tactics.
Import Check.
Goal True.
idtac "------------------- apply_exercise1 --------------------".
idtac " ".
idtac "#> rev_exercise1".
idtac "Possible points: 3".
check_type @rev_exercise1 ((forall l l' : list nat, l = @rev nat l' -> l' = @rev nat l)).
idtac "Assumptions:".
Abort.
Print Assumptions rev_exercise1.
Goal True.
idtac " ".
idtac "------------------- injection_ex3 --------------------".
idtac " ".
idtac "#> injection_ex3".
idtac "Possible points: 1".
check_type @injection_ex3 (
(forall (X : Type) (x y z : X) (l j : list X),
x :: y :: l = z :: j -> y :: l = x :: j -> x = y)).
idtac "Assumptions:".
Abort.
Print Assumptions injection_ex3.
Goal True.
idtac " ".
idtac "------------------- discriminate_ex3 --------------------".
idtac " ".
idtac "#> discriminate_ex3".
idtac "Possible points: 1".
check_type @discriminate_ex3 (
(forall (X : Type) (x y z : X) (l : list X),
list X -> x :: y :: l = [ ] -> x = z)).
idtac "Assumptions:".
Abort.
Print Assumptions discriminate_ex3.
Goal True.
idtac " ".
idtac "------------------- plus_n_n_injective --------------------".
idtac " ".
idtac "#> plus_n_n_injective".
idtac "Possible points: 3".
check_type @plus_n_n_injective ((forall n m : nat, n + n = m + m -> n = m)).
idtac "Assumptions:".
Abort.
Print Assumptions plus_n_n_injective.
Goal True.
idtac " ".
idtac "------------------- eqb_true --------------------".
idtac " ".
idtac "#> eqb_true".
idtac "Possible points: 2".
check_type @eqb_true ((forall n m : nat, (n =? m) = true -> n = m)).
idtac "Assumptions:".
Abort.
Print Assumptions eqb_true.
Goal True.
idtac " ".
idtac "------------------- eqb_true_informal --------------------".
idtac " ".
idtac "#> Manually graded: informal_proof".
idtac "Advanced".
idtac "Possible points: 2".
print_manual_grade manual_grade_for_informal_proof.
idtac " ".
idtac "------------------- gen_dep_practice --------------------".
idtac " ".
idtac "#> nth_error_after_last".
idtac "Possible points: 3".
check_type @nth_error_after_last (
(forall (n : nat) (X : Type) (l : list X),
@length X l = n -> @nth_error X l n = @None X)).
idtac "Assumptions:".
Abort.
Print Assumptions nth_error_after_last.
Goal True.
idtac " ".
idtac "------------------- destruct_eqn_practice --------------------".
idtac " ".
idtac "#> bool_fn_applied_thrice".
idtac "Possible points: 2".
check_type @bool_fn_applied_thrice (
(forall (f : bool -> bool) (b : bool), f (f (f b)) = f b)).
idtac "Assumptions:".
Abort.
Print Assumptions bool_fn_applied_thrice.
Goal True.
idtac " ".
idtac "------------------- eqb_sym --------------------".
idtac " ".
idtac "#> eqb_sym".
idtac "Possible points: 3".
check_type @eqb_sym ((forall n m : nat, (n =? m) = (m =? n))).
idtac "Assumptions:".
Abort.
Print Assumptions eqb_sym.
Goal True.
idtac " ".
idtac "------------------- split_combine --------------------".
idtac " ".
idtac "#> Manually graded: split_combine".
idtac "Advanced".
idtac "Possible points: 3".
print_manual_grade manual_grade_for_split_combine.
idtac " ".
idtac "------------------- filter_exercise --------------------".
idtac " ".
idtac "#> filter_exercise".
idtac "Advanced".
idtac "Possible points: 3".
check_type @filter_exercise (
(forall (X : Type) (test : X -> bool) (x : X) (l lf : list X),
@filter X test l = x :: lf -> test x = true)).
idtac "Assumptions:".
Abort.
Print Assumptions filter_exercise.
Goal True.
idtac " ".
idtac "------------------- forall_exists_challenge --------------------".
idtac " ".
idtac "#> existsb_existsb'".
idtac "Advanced".
idtac "Possible points: 4".
check_type @existsb_existsb' (
(forall (X : Type) (test : X -> bool) (l : list X),
@existsb X test l = @existsb' X test l)).
idtac "Assumptions:".
Abort.
Print Assumptions existsb_existsb'.
Goal True.
idtac " ".
idtac " ".
idtac "Max points - standard: 18".
idtac "Max points - advanced: 30".
idtac "".
idtac "********** Summary **********".
idtac "".
idtac "********** Standard **********".
idtac "---------- rev_exercise1 ---------".
Print Assumptions rev_exercise1.
idtac "---------- injection_ex3 ---------".
Print Assumptions injection_ex3.
idtac "---------- discriminate_ex3 ---------".
Print Assumptions discriminate_ex3.
idtac "---------- plus_n_n_injective ---------".
Print Assumptions plus_n_n_injective.
idtac "---------- eqb_true ---------".
Print Assumptions eqb_true.
idtac "---------- nth_error_after_last ---------".
Print Assumptions nth_error_after_last.
idtac "---------- bool_fn_applied_thrice ---------".
Print Assumptions bool_fn_applied_thrice.
idtac "---------- eqb_sym ---------".
Print Assumptions eqb_sym.
idtac "".
idtac "********** Advanced **********".
idtac "---------- informal_proof ---------".
idtac "MANUAL".
idtac "---------- split_combine ---------".
idtac "MANUAL".
idtac "---------- filter_exercise ---------".
Print Assumptions filter_exercise.
idtac "---------- existsb_existsb' ---------".
Print Assumptions existsb_existsb'.
Abort.
(* Wed Jan 9 12:02:11 EST 2019 *)

2042
imp.ml Normal file

File diff suppressed because it is too large Load diff

182
imp.mli Normal file
View file

@ -0,0 +1,182 @@
val negb : bool -> bool
val app : 'a1 list -> 'a1 list -> 'a1 list
val add : int -> int -> int
val mul : int -> int -> int
val sub : int -> int -> int
val eqb : int -> int -> bool
val leb : int -> int -> bool
module Nat :
sig
val eqb : int -> int -> bool
val leb : int -> int -> bool
end
type positive =
| XI of positive
| XO of positive
| XH
type n =
| N0
| Npos of positive
module Pos :
sig
val succ : positive -> positive
val add : positive -> positive -> positive
val add_carry : positive -> positive -> positive
val mul : positive -> positive -> positive
val iter_op : ('a1 -> 'a1 -> 'a1) -> positive -> 'a1 -> 'a1
val to_nat : positive -> int
end
module N :
sig
val add : n -> n -> n
val mul : n -> n -> n
val to_nat : n -> int
end
val rev : 'a1 list -> 'a1 list
val map : ('a1 -> 'a2) -> 'a1 list -> 'a2 list
val fold_left : ('a1 -> 'a2 -> 'a1) -> 'a2 list -> 'a1 -> 'a1
val fold_right : ('a2 -> 'a1 -> 'a1) -> 'a1 -> 'a2 list -> 'a1
val forallb : ('a1 -> bool) -> 'a1 list -> bool
val n_of_digits : bool list -> n
val n_of_ascii : char -> n
val nat_of_ascii : char -> int
val string_dec : char list -> char list -> bool
val append : char list -> char list -> char list
val eqb_string : char list -> char list -> bool
type 'a total_map = char list -> 'a
val t_empty : 'a1 -> 'a1 total_map
val t_update : 'a1 total_map -> char list -> 'a1 -> char list -> 'a1
type state = int total_map
type aexp =
| ANum of int
| AId of char list
| APlus of aexp * aexp
| AMinus of aexp * aexp
| AMult of aexp * aexp
type bexp =
| BTrue
| BFalse
| BEq of aexp * aexp
| BLe of aexp * aexp
| BNot of bexp
| BAnd of bexp * bexp
val aeval : state -> aexp -> int
val beval : state -> bexp -> bool
val empty_st : int total_map
type com =
| CSkip
| CAss of char list * aexp
| CSeq of com * com
| CIf of bexp * com * com
| CWhile of bexp * com
val ceval_step : state -> com -> int -> state option
val isWhite : char -> bool
val isLowerAlpha : char -> bool
val isAlpha : char -> bool
val isDigit : char -> bool
type chartype =
| White
| Alpha
| Digit
| Other
val classifyChar : char -> chartype
val list_of_string : char list -> char list
val string_of_list : char list -> char list
type token = char list
val tokenize_helper : chartype -> char list -> char list -> char list list
val tokenize : char list -> char list list
type 'x optionE =
| SomeE of 'x
| NoneE of char list
type 't parser0 = token list -> ('t * token list) optionE
val many_helper :
'a1 parser0 -> 'a1 list -> int -> token list -> ('a1 list * token list)
optionE
val many : 'a1 parser0 -> int -> 'a1 list parser0
val firstExpect : token -> 'a1 parser0 -> 'a1 parser0
val expect : token -> unit parser0
val parseIdentifier : token list -> (char list * token list) optionE
val parseNumber : token list -> (int * token list) optionE
val parsePrimaryExp : int -> token list -> (aexp * token list) optionE
val parseProductExp : int -> token list -> (aexp * token list) optionE
val parseSumExp : int -> token list -> (aexp * token list) optionE
val parseAExp : int -> token list -> (aexp * token list) optionE
val parseAtomicExp : int -> token list -> (bexp * token list) optionE
val parseConjunctionExp : int -> token list -> (bexp * token list) optionE
val parseBExp : int -> token list -> (bexp * token list) optionE
val parseSimpleCommand : int -> token list -> (com * token list) optionE
val parseSequencedCommand : int -> token list -> (com * token list) optionE
val bignumber : int
val parse : char list -> com optionE

211
imp1.ml Normal file
View file

@ -0,0 +1,211 @@
type bool =
| True
| False
(** val negb : bool -> bool **)
let negb = function
| True -> False
| False -> True
type nat =
| O
| S of nat
type 'a option =
| Some of 'a
| None
type sumbool =
| Left
| Right
(** val add : nat -> nat -> nat **)
let rec add n m =
match n with
| O -> m
| S p -> S (add p m)
(** val mul : nat -> nat -> nat **)
let rec mul n m =
match n with
| O -> O
| S p -> add m (mul p m)
(** val sub : nat -> nat -> nat **)
let rec sub n m =
match n with
| O -> n
| S k -> (match m with
| O -> n
| S l -> sub k l)
(** val eqb : nat -> nat -> bool **)
let rec eqb n m =
match n with
| O -> (match m with
| O -> True
| S _ -> False)
| S n' -> (match m with
| O -> False
| S m' -> eqb n' m')
(** val leb : nat -> nat -> bool **)
let rec leb n m =
match n with
| O -> True
| S n' -> (match m with
| O -> False
| S m' -> leb n' m')
(** val bool_dec : bool -> bool -> sumbool **)
let bool_dec b1 b2 =
match b1 with
| True -> (match b2 with
| True -> Left
| False -> Right)
| False -> (match b2 with
| True -> Right
| False -> Left)
type ascii =
| Ascii of bool * bool * bool * bool * bool * bool * bool * bool
(** val ascii_dec : ascii -> ascii -> sumbool **)
let ascii_dec a b =
let Ascii (x, x0, x1, x2, x3, x4, x5, x6) = a in
let Ascii (b8, b9, b10, b11, b12, b13, b14, b15) = b in
(match bool_dec x b8 with
| Left ->
(match bool_dec x0 b9 with
| Left ->
(match bool_dec x1 b10 with
| Left ->
(match bool_dec x2 b11 with
| Left ->
(match bool_dec x3 b12 with
| Left ->
(match bool_dec x4 b13 with
| Left ->
(match bool_dec x5 b14 with
| Left -> bool_dec x6 b15
| Right -> Right)
| Right -> Right)
| Right -> Right)
| Right -> Right)
| Right -> Right)
| Right -> Right)
| Right -> Right)
type string =
| EmptyString
| String of ascii * string
(** val string_dec : string -> string -> sumbool **)
let rec string_dec s x =
match s with
| EmptyString -> (match x with
| EmptyString -> Left
| String (_, _) -> Right)
| String (a, s0) ->
(match x with
| EmptyString -> Right
| String (a0, s1) ->
(match ascii_dec a a0 with
| Left -> string_dec s0 s1
| Right -> Right))
(** val eqb_string : string -> string -> bool **)
let eqb_string x y =
match string_dec x y with
| Left -> True
| Right -> False
type 'a total_map = string -> 'a
(** val t_update : 'a1 total_map -> string -> 'a1 -> string -> 'a1 **)
let t_update m x v x' =
match eqb_string x x' with
| True -> v
| False -> m x'
type state = nat total_map
type aexp =
| ANum of nat
| AId of string
| APlus of aexp * aexp
| AMinus of aexp * aexp
| AMult of aexp * aexp
type bexp =
| BTrue
| BFalse
| BEq of aexp * aexp
| BLe of aexp * aexp
| BNot of bexp
| BAnd of bexp * bexp
(** val aeval : state -> aexp -> nat **)
let rec aeval st = function
| ANum n -> n
| AId x -> st x
| APlus (a1, a2) -> add (aeval st a1) (aeval st a2)
| AMinus (a1, a2) -> sub (aeval st a1) (aeval st a2)
| AMult (a1, a2) -> mul (aeval st a1) (aeval st a2)
(** val beval : state -> bexp -> bool **)
let rec beval st = function
| BTrue -> True
| BFalse -> False
| BEq (a1, a2) -> eqb (aeval st a1) (aeval st a2)
| BLe (a1, a2) -> leb (aeval st a1) (aeval st a2)
| BNot b1 -> negb (beval st b1)
| BAnd (b1, b2) ->
(match beval st b1 with
| True -> beval st b2
| False -> False)
type com =
| CSkip
| CAss of string * aexp
| CSeq of com * com
| CIf of bexp * com * com
| CWhile of bexp * com
(** val ceval_step : state -> com -> nat -> state option **)
let rec ceval_step st c = function
| O -> None
| S i' ->
(match c with
| CSkip -> Some st
| CAss (l, a1) -> Some (t_update st l (aeval st a1))
| CSeq (c1, c2) ->
(match ceval_step st c1 i' with
| Some st' -> ceval_step st' c2 i'
| None -> None)
| CIf (b, c1, c2) ->
(match beval st b with
| True -> ceval_step st c1 i'
| False -> ceval_step st c2 i')
| CWhile (b1, c1) ->
(match beval st b1 with
| True ->
(match ceval_step st c1 i' with
| Some st' -> ceval_step st' c i'
| None -> None)
| False -> Some st))

77
imp1.mli Normal file
View file

@ -0,0 +1,77 @@
type bool =
| True
| False
val negb : bool -> bool
type nat =
| O
| S of nat
type 'a option =
| Some of 'a
| None
type sumbool =
| Left
| Right
val add : nat -> nat -> nat
val mul : nat -> nat -> nat
val sub : nat -> nat -> nat
val eqb : nat -> nat -> bool
val leb : nat -> nat -> bool
val bool_dec : bool -> bool -> sumbool
type ascii =
| Ascii of bool * bool * bool * bool * bool * bool * bool * bool
val ascii_dec : ascii -> ascii -> sumbool
type string =
| EmptyString
| String of ascii * string
val string_dec : string -> string -> sumbool
val eqb_string : string -> string -> bool
type 'a total_map = string -> 'a
val t_update : 'a1 total_map -> string -> 'a1 -> string -> 'a1
type state = nat total_map
type aexp =
| ANum of nat
| AId of string
| APlus of aexp * aexp
| AMinus of aexp * aexp
| AMult of aexp * aexp
type bexp =
| BTrue
| BFalse
| BEq of aexp * aexp
| BLe of aexp * aexp
| BNot of bexp
| BAnd of bexp * bexp
val aeval : state -> aexp -> nat
val beval : state -> bexp -> bool
type com =
| CSkip
| CAss of string * aexp
| CSeq of com * com
| CIf of bexp * com * com
| CWhile of bexp * com
val ceval_step : state -> com -> nat -> state option

189
imp2.ml Normal file
View file

@ -0,0 +1,189 @@
(** val negb : bool -> bool **)
let negb = function
| true -> false
| false -> true
type 'a option =
| Some of 'a
| None
type sumbool =
| Left
| Right
(** val add : int -> int -> int **)
let rec add = ( + )
(** val mul : int -> int -> int **)
let rec mul = ( * )
(** val sub : int -> int -> int **)
let rec sub n m =
(fun zero succ n ->
if n=0 then zero () else succ (n-1))
(fun _ -> n)
(fun k ->
(fun zero succ n ->
if n=0 then zero () else succ (n-1))
(fun _ -> n)
(fun l -> sub k l)
m)
n
(** val eqb : int -> int -> bool **)
let rec eqb = ( = )
(** val leb : int -> int -> bool **)
let rec leb n m =
(fun zero succ n ->
if n=0 then zero () else succ (n-1))
(fun _ -> true)
(fun n' ->
(fun zero succ n ->
if n=0 then zero () else succ (n-1))
(fun _ -> false)
(fun m' -> leb n' m')
m)
n
(** val bool_dec : bool -> bool -> sumbool **)
let bool_dec b1 b2 =
if b1 then if b2 then Left else Right else if b2 then Right else Left
type ascii =
| Ascii of bool * bool * bool * bool * bool * bool * bool * bool
(** val ascii_dec : ascii -> ascii -> sumbool **)
let ascii_dec a b =
let Ascii (x, x0, x1, x2, x3, x4, x5, x6) = a in
let Ascii (b8, b9, b10, b11, b12, b13, b14, b15) = b in
(match bool_dec x b8 with
| Left ->
(match bool_dec x0 b9 with
| Left ->
(match bool_dec x1 b10 with
| Left ->
(match bool_dec x2 b11 with
| Left ->
(match bool_dec x3 b12 with
| Left ->
(match bool_dec x4 b13 with
| Left ->
(match bool_dec x5 b14 with
| Left -> bool_dec x6 b15
| Right -> Right)
| Right -> Right)
| Right -> Right)
| Right -> Right)
| Right -> Right)
| Right -> Right)
| Right -> Right)
type string =
| EmptyString
| String of ascii * string
(** val string_dec : string -> string -> sumbool **)
let rec string_dec s x =
match s with
| EmptyString -> (match x with
| EmptyString -> Left
| String (_, _) -> Right)
| String (a, s0) ->
(match x with
| EmptyString -> Right
| String (a0, s1) ->
(match ascii_dec a a0 with
| Left -> string_dec s0 s1
| Right -> Right))
(** val eqb_string : string -> string -> bool **)
let eqb_string x y =
match string_dec x y with
| Left -> true
| Right -> false
type 'a total_map = string -> 'a
(** val t_update : 'a1 total_map -> string -> 'a1 -> string -> 'a1 **)
let t_update m x v x' =
if eqb_string x x' then v else m x'
type state = int total_map
type aexp =
| ANum of int
| AId of string
| APlus of aexp * aexp
| AMinus of aexp * aexp
| AMult of aexp * aexp
type bexp =
| BTrue
| BFalse
| BEq of aexp * aexp
| BLe of aexp * aexp
| BNot of bexp
| BAnd of bexp * bexp
(** val aeval : state -> aexp -> int **)
let rec aeval st = function
| ANum n -> n
| AId x -> st x
| APlus (a1, a2) -> add (aeval st a1) (aeval st a2)
| AMinus (a1, a2) -> sub (aeval st a1) (aeval st a2)
| AMult (a1, a2) -> mul (aeval st a1) (aeval st a2)
(** val beval : state -> bexp -> bool **)
let rec beval st = function
| BTrue -> true
| BFalse -> false
| BEq (a1, a2) -> eqb (aeval st a1) (aeval st a2)
| BLe (a1, a2) -> leb (aeval st a1) (aeval st a2)
| BNot b1 -> negb (beval st b1)
| BAnd (b1, b2) -> if beval st b1 then beval st b2 else false
type com =
| CSkip
| CAss of string * aexp
| CSeq of com * com
| CIf of bexp * com * com
| CWhile of bexp * com
(** val ceval_step : state -> com -> int -> state option **)
let rec ceval_step st c i =
(fun zero succ n ->
if n=0 then zero () else succ (n-1))
(fun _ -> None)
(fun i' ->
match c with
| CSkip -> Some st
| CAss (l, a1) -> Some (t_update st l (aeval st a1))
| CSeq (c1, c2) ->
(match ceval_step st c1 i' with
| Some st' -> ceval_step st' c2 i'
| None -> None)
| CIf (b, c1, c2) ->
if beval st b then ceval_step st c1 i' else ceval_step st c2 i'
| CWhile (b1, c1) ->
if beval st b1
then (match ceval_step st c1 i' with
| Some st' -> ceval_step st' c i'
| None -> None)
else Some st)
i

69
imp2.mli Normal file
View file

@ -0,0 +1,69 @@
val negb : bool -> bool
type 'a option =
| Some of 'a
| None
type sumbool =
| Left
| Right
val add : int -> int -> int
val mul : int -> int -> int
val sub : int -> int -> int
val eqb : int -> int -> bool
val leb : int -> int -> bool
val bool_dec : bool -> bool -> sumbool
type ascii =
| Ascii of bool * bool * bool * bool * bool * bool * bool * bool
val ascii_dec : ascii -> ascii -> sumbool
type string =
| EmptyString
| String of ascii * string
val string_dec : string -> string -> sumbool
val eqb_string : string -> string -> bool
type 'a total_map = string -> 'a
val t_update : 'a1 total_map -> string -> 'a1 -> string -> 'a1
type state = int total_map
type aexp =
| ANum of int
| AId of string
| APlus of aexp * aexp
| AMinus of aexp * aexp
| AMult of aexp * aexp
type bexp =
| BTrue
| BFalse
| BEq of aexp * aexp
| BLe of aexp * aexp
| BNot of bexp
| BAnd of bexp * bexp
val aeval : state -> aexp -> int
val beval : state -> bexp -> bool
type com =
| CSkip
| CAss of string * aexp
| CSeq of com * com
| CIf of bexp * com * com
| CWhile of bexp * com
val ceval_step : state -> com -> int -> state option