AutoMore Automation
Set Warnings "-notation-overridden,-parsing,-deprecated-hint-without-locality".
From Coq Require Import Lia.
From LF Require Import Maps.
From LF Require Import Imp.
 
From Coq Require Import Lia.
From LF Require Import Maps.
From LF Require Import Imp.
Consider the proof below, showing that ceval is
    deterministic.  Notice all the repetition and near-repetition... 
Theorem ceval_deterministic: ∀ 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; inversion E2; subst.
- (* E_Skip *) reflexivity.
- (* E_Asgn *) reflexivity.
- (* E_Seq *)
rewrite (IHE1_1 st'0 H1) in ×.
apply IHE1_2. assumption.
(* E_IfTrue *)
- (* b evaluates to true *)
apply IHE1. assumption.
- (* b evaluates to false (contradiction) *)
rewrite H in H5. discriminate.
(* E_IfFalse *)
- (* b evaluates to true (contradiction) *)
rewrite H in H5. discriminate.
- (* b evaluates to false *)
apply IHE1. assumption.
(* E_WhileFalse *)
- (* b evaluates to false *)
reflexivity.
- (* b evaluates to true (contradiction) *)
rewrite H in H2. discriminate.
(* E_WhileTrue *)
- (* b evaluates to false (contradiction) *)
rewrite H in H4. discriminate.
- (* b evaluates to true *)
rewrite (IHE1_1 st'0 H3) in ×.
apply IHE1_2. assumption. Qed.
st =[ c ]=> st1 →
st =[ c ]=> st2 →
st1 = st2.
Proof.
intros c st st1 st2 E1 E2;
generalize dependent st2;
induction E1; intros st2 E2; inversion E2; subst.
- (* E_Skip *) reflexivity.
- (* E_Asgn *) reflexivity.
- (* E_Seq *)
rewrite (IHE1_1 st'0 H1) in ×.
apply IHE1_2. assumption.
(* E_IfTrue *)
- (* b evaluates to true *)
apply IHE1. assumption.
- (* b evaluates to false (contradiction) *)
rewrite H in H5. discriminate.
(* E_IfFalse *)
- (* b evaluates to true (contradiction) *)
rewrite H in H5. discriminate.
- (* b evaluates to false *)
apply IHE1. assumption.
(* E_WhileFalse *)
- (* b evaluates to false *)
reflexivity.
- (* b evaluates to true (contradiction) *)
rewrite H in H2. discriminate.
(* E_WhileTrue *)
- (* b evaluates to false (contradiction) *)
rewrite H in H4. discriminate.
- (* b evaluates to true *)
rewrite (IHE1_1 st'0 H3) in ×.
apply IHE1_2. assumption. Qed.
The auto Tactic
Example auto_example_1' : ∀ (P Q R: Prop),
(P → Q) → (Q → R) → P → R.
Proof.
auto.
Qed.
(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).
Example auto_example_2 : ∀ 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.
(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 : ∀ (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.
(* Let's see where auto gets stuck using debug auto *)
debug auto.
(* Optional argument says how deep to search (default is 5) *)
auto 6.
Qed.
(P → Q) →
(Q → R) →
(R → S) →
(S → T) →
(T → U) →
P →
U.
Proof.
(* When it cannot solve the goal, auto does nothing *)
auto.
(* Let's see where auto gets stuck using debug auto *)
debug auto.
(* Optional argument says how deep to search (default is 5) *)
auto 6.
Qed.
Example auto_example_4 : ∀ P Q R : Prop,
Q →
(Q → R) →
P ∨ (Q ∧ R).
Proof. auto. Qed.
Q →
(Q → R) →
P ∨ (Q ∧ R).
Proof. auto. Qed.
If we want to see which facts auto is using, we can use
    info_auto instead. 
Example auto_example_5: 2 = 2.
Proof.
info_auto.
Qed.
Example auto_example_5' : ∀ (P Q R S T U W: Prop),
(U → T) →
(W → U) →
(R → S) →
(S → T) →
(P → R) →
(U → T) →
P →
T.
Proof.
intros.
info_auto.
Qed.
Proof.
info_auto.
Qed.
Example auto_example_5' : ∀ (P Q R S T U W: Prop),
(U → T) →
(W → U) →
(R → S) →
(S → T) →
(P → R) →
(U → T) →
P →
T.
Proof.
intros.
info_auto.
Qed.
We can extend the hint database just for the purposes of one application of auto by writing "auto using ...".
Lemma le_antisym : ∀ n m: nat, (n ≤ m ∧ m ≤ n) → n = m.
Proof. lia. Qed.
Example auto_example_6 : ∀ n m p : nat,
(n ≤ p → (n ≤ m ∧ m ≤ n)) →
n ≤ p →
n = m.
Proof.
auto using le_antisym.
Qed.
Proof. lia. Qed.
Example auto_example_6 : ∀ n m p : nat,
(n ≤ p → (n ≤ m ∧ m ≤ n)) →
n ≤ p →
n = m.
Proof.
auto using le_antisym.
Qed.
-  Hint Resolve T : core.
Add theorem or constructor T to the global DB
-  Hint Constructors c : core.
Add all constructors of c to the global DB
-  Hint Unfold d : core.
Automatically expand defined symbol d during auto
Hint Resolve le_antisym : core.
Example auto_example_6' : ∀ n m p : nat,
(n≤ p → (n ≤ m ∧ m ≤ n)) →
n ≤ p →
n = m.
Proof.
auto. (* picks up hint from database *)
Qed.
Example auto_example_6' : ∀ n m p : nat,
(n≤ p → (n ≤ m ∧ m ≤ n)) →
n ≤ p →
n = m.
Proof.
auto. (* picks up hint from database *)
Qed.
Definition is_fortytwo x := (x = 42).
Example auto_example_7: ∀ x,
(x ≤ 42 ∧ 42 ≤ x) → is_fortytwo x.
Proof.
auto. (* does nothing *)
Abort.
Hint Unfold is_fortytwo : core.
Example auto_example_7' : ∀ x,
(x ≤ 42 ∧ 42 ≤ x) → is_fortytwo x.
Proof.
auto. (* try also: info_auto. *)
Qed.
Example auto_example_7: ∀ x,
(x ≤ 42 ∧ 42 ≤ x) → is_fortytwo x.
Proof.
auto. (* does nothing *)
Abort.
Hint Unfold is_fortytwo : core.
Example auto_example_7' : ∀ x,
(x ≤ 42 ∧ 42 ≤ x) → is_fortytwo x.
Proof.
auto. (* try also: info_auto. *)
Qed.
Theorem ceval_deterministic': ∀ 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; inversion E2; subst; auto.
- (* E_Seq *)
rewrite (IHE1_1 st'0 H1) in ×.
auto.
- (* E_IfTrue *)
+ (* b evaluates to false (contradiction) *)
rewrite H in H5. discriminate.
- (* E_IfFalse *)
+ (* b evaluates to true (contradiction) *)
rewrite H in H5. discriminate.
- (* E_WhileFalse *)
+ (* b evaluates to true (contradiction) *)
rewrite H in H2. discriminate.
(* E_WhileTrue *)
- (* b evaluates to false (contradiction) *)
rewrite H in H4. discriminate.
- (* b evaluates to true *)
rewrite (IHE1_1 st'0 H3) in ×.
auto.
Qed.
st =[ c ]=> st1 →
st =[ c ]=> st2 →
st1 = st2.
Proof.
intros c st st1 st2 E1 E2.
generalize dependent st2;
induction E1; intros st2 E2; inversion E2; subst; auto.
- (* E_Seq *)
rewrite (IHE1_1 st'0 H1) in ×.
auto.
- (* E_IfTrue *)
+ (* b evaluates to false (contradiction) *)
rewrite H in H5. discriminate.
- (* E_IfFalse *)
+ (* b evaluates to true (contradiction) *)
rewrite H in H5. discriminate.
- (* E_WhileFalse *)
+ (* b evaluates to true (contradiction) *)
rewrite H in H2. discriminate.
(* E_WhileTrue *)
- (* b evaluates to false (contradiction) *)
rewrite H in H4. discriminate.
- (* b evaluates to true *)
rewrite (IHE1_1 st'0 H3) in ×.
auto.
Qed.
Searching For Hypotheses
H1: beval st b = false as well as:
H2: beval st b = true
Ltac rwd H1 H2 := rewrite H1 in H2; discriminate.
Theorem ceval_deterministic'': ∀ 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; inversion E2; subst; auto.
- (* E_Seq *)
rewrite (IHE1_1 st'0 H1) in ×.
auto.
- (* E_IfTrue *)
+ (* b evaluates to false (contradiction) *)
rwd H H5.
- (* E_IfFalse *)
+ (* b evaluates to true (contradiction) *)
rwd H H5.
- (* E_WhileFalse *)
+ (* b evaluates to true (contradiction) *)
rwd H H2.
(* E_WhileTrue *)
- (* b evaluates to false (contradiction) *)
rwd H H4.
- (* b evaluates to true *)
rewrite (IHE1_1 st'0 H3) in ×.
auto. Qed.
st =[ c ]=> st1 →
st =[ c ]=> st2 →
st1 = st2.
Proof.
intros c st st1 st2 E1 E2.
generalize dependent st2;
induction E1; intros st2 E2; inversion E2; subst; auto.
- (* E_Seq *)
rewrite (IHE1_1 st'0 H1) in ×.
auto.
- (* E_IfTrue *)
+ (* b evaluates to false (contradiction) *)
rwd H H5.
- (* E_IfFalse *)
+ (* b evaluates to true (contradiction) *)
rwd H H5.
- (* E_WhileFalse *)
+ (* b evaluates to true (contradiction) *)
rwd H H2.
(* E_WhileTrue *)
- (* b evaluates to false (contradiction) *)
rwd H H4.
- (* b evaluates to true *)
rewrite (IHE1_1 st'0 H3) in ×.
auto. Qed.
Ltac find_rwd :=
match goal with
H1: ?E = true,
H2: ?E = false
⊢ _ ⇒ rwd H1 H2
end.
match goal with
H1: ?E = true,
H2: ?E = false
⊢ _ ⇒ rwd H1 H2
end.
The match goal tactic looks for hypotheses matching the
    pattern specified. In this case, we're looking for two equalities
    H1 and H2 equating the same expression ?E to both true and
    false. 
 
Theorem ceval_deterministic''': ∀ 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; inversion E2; subst; try find_rwd; 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.
st =[ c ]=> st1 →
st =[ c ]=> st2 →
st1 = st2.
Proof.
intros c st st1 st2 E1 E2.
generalize dependent st2;
induction E1; intros st2 E2; inversion E2; subst; try find_rwd; 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.
Ltac find_eqn :=
match goal with
H1: ∀ x, ?P x → ?L = ?R,
H2: ?P ?X
⊢ _ ⇒ rewrite (H1 X H2) in ×
end.
match goal with
H1: ∀ x, ?P x → ?L = ?R,
H2: ?P ?X
⊢ _ ⇒ rewrite (H1 X H2) in ×
end.
Now we can make use of find_eqn to repeatedly rewrite with the appropriate hypothesis, wherever it may be found.
Theorem ceval_deterministic'''': ∀ 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; inversion E2; subst; try find_rwd;
try find_eqn; auto.
Qed.
st =[ c ]=> st1 →
st =[ c ]=> st2 →
st1 = st2.
Proof.
intros c st st1 st2 E1 E2.
generalize dependent st2;
induction E1; intros st2 E2; inversion E2; subst; try find_rwd;
try find_eqn; auto.
Qed.
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).
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 "'repeat' x 'until' y 'end'" :=
(CRepeat x y)
(in custom com at level 0,
x at level 99, y at level 99).
Notation "'skip'" :=
CSkip (in custom com at level 0).
Notation "x := y" :=
(CAsgn x y)
(in custom com at level 0, x constr at level 0,
y at level 85, no associativity).
Notation "x ; y" :=
(CSeq x y)
(in custom com at level 90, right associativity).
Notation "'if' x 'then' y 'else' z 'end'" :=
(CIf x y z)
(in custom com at level 89, x at level 99,
y at level 99, z at level 99).
Notation "'while' x 'do' y 'end'" :=
(CWhile x y)
(in custom com at level 89, x at level 99, y at level 99).
(CRepeat x y)
(in custom com at level 0,
x at level 99, y at level 99).
Notation "'skip'" :=
CSkip (in custom com at level 0).
Notation "x := y" :=
(CAsgn x y)
(in custom com at level 0, x constr at level 0,
y at level 85, no associativity).
Notation "x ; y" :=
(CSeq x y)
(in custom com at level 90, right associativity).
Notation "'if' x 'then' y 'else' z 'end'" :=
(CIf x y z)
(in custom com at level 89, x at level 99,
y at level 99, z at level 99).
Notation "'while' x 'do' y 'end'" :=
(CWhile x y)
(in custom com at level 89, x at level 99, y at level 99).
Reserved Notation "st '=[' c ']=>' st'"
(at level 40, c custom com at level 99, st' constr at next level).
Inductive ceval : com → state → state → Prop :=
| E_Skip : ∀ st,
st =[ skip ]=> st
| E_Asgn : ∀ st a1 n x,
aeval st a1 = n →
st =[ x := a1 ]=> (x !-> n ; st)
| E_Seq : ∀ c1 c2 st st' st'',
st =[ c1 ]=> st' →
st' =[ c2 ]=> st'' →
st =[ c1 ; c2 ]=> st''
| E_IfTrue : ∀ st st' b c1 c2,
beval st b = true →
st =[ c1 ]=> st' →
st =[ if b then c1 else c2 end ]=> st'
| E_IfFalse : ∀ st st' b c1 c2,
beval st b = false →
st =[ c2 ]=> st' →
st =[ if b then c1 else c2 end ]=> st'
| E_WhileFalse : ∀ b st c,
beval st b = false →
st =[ while b do c end ]=> st
| E_WhileTrue : ∀ st st' st'' b c,
beval st b = true →
st =[ c ]=> st' →
st' =[ while b do c end ]=> st'' →
st =[ while b do c end ]=> st''
| E_RepeatEnd : ∀ st st' b c,
st =[ c ]=> st' →
beval st' b = true →
st =[ repeat c until b end ]=> st'
| E_RepeatLoop : ∀ st st' st'' b c,
st =[ c ]=> st' →
beval st' b = false →
st' =[ repeat c until b end ]=> st'' →
st =[ repeat c until b end ]=> st''
where "st =[ c ]=> st'" := (ceval c st st').
(at level 40, c custom com at level 99, st' constr at next level).
Inductive ceval : com → state → state → Prop :=
| E_Skip : ∀ st,
st =[ skip ]=> st
| E_Asgn : ∀ st a1 n x,
aeval st a1 = n →
st =[ x := a1 ]=> (x !-> n ; st)
| E_Seq : ∀ c1 c2 st st' st'',
st =[ c1 ]=> st' →
st' =[ c2 ]=> st'' →
st =[ c1 ; c2 ]=> st''
| E_IfTrue : ∀ st st' b c1 c2,
beval st b = true →
st =[ c1 ]=> st' →
st =[ if b then c1 else c2 end ]=> st'
| E_IfFalse : ∀ st st' b c1 c2,
beval st b = false →
st =[ c2 ]=> st' →
st =[ if b then c1 else c2 end ]=> st'
| E_WhileFalse : ∀ b st c,
beval st b = false →
st =[ while b do c end ]=> st
| E_WhileTrue : ∀ st st' st'' b c,
beval st b = true →
st =[ c ]=> st' →
st' =[ while b do c end ]=> st'' →
st =[ while b do c end ]=> st''
| E_RepeatEnd : ∀ st st' b c,
st =[ c ]=> st' →
beval st' b = true →
st =[ repeat c until b end ]=> st'
| E_RepeatLoop : ∀ st st' st'' b c,
st =[ c ]=> st' →
beval st' b = false →
st' =[ repeat c until b end ]=> st'' →
st =[ repeat c until b end ]=> st''
where "st =[ c ]=> st'" := (ceval c st st').
Theorem ceval_deterministic: ∀ 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; inversion E2; subst; try find_rwd; try find_eqn; auto.
- (* E_RepeatEnd *)
+ (* b evaluates to false (contradiction) *)
find_rwd.
(* oops: why didn't find_rwd solve this for us already?
answer: we did things in the wrong order. *)
- (* E_RepeatLoop *)
+ (* b evaluates to true (contradiction) *)
find_rwd.
Qed.
st =[ c ]=> st1 →
st =[ c ]=> st2 →
st1 = st2.
Proof.
intros c st st1 st2 E1 E2.
generalize dependent st2;
induction E1;
intros st2 E2; inversion E2; subst; try find_rwd; try find_eqn; auto.
- (* E_RepeatEnd *)
+ (* b evaluates to false (contradiction) *)
find_rwd.
(* oops: why didn't find_rwd solve this for us already?
answer: we did things in the wrong order. *)
- (* E_RepeatLoop *)
+ (* b evaluates to true (contradiction) *)
find_rwd.
Qed.
Theorem ceval_deterministic': ∀ 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; inversion E2; subst; try find_eqn; try find_rwd; auto.
Qed.
End Repeat.
st =[ c ]=> st1 →
st =[ c ]=> st2 →
st1 = st2.
Proof.
intros c st st1 st2 E1 E2.
generalize dependent st2;
induction E1;
intros st2 E2; inversion E2; subst; try find_eqn; try find_rwd; auto.
Qed.
End Repeat.
Example ceval_example1:
empty_st =[
X := 2;
if (X ≤ 1)
then Y := 3
else Z := 4
end
]=> (Z !-> 4 ; X !-> 2).
Proof.
(* We supply the intermediate state st'... *)
apply E_Seq with (X !-> 2).
- apply E_Asgn. reflexivity.
- apply E_IfFalse. reflexivity. apply E_Asgn. reflexivity.
Qed.
empty_st =[
X := 2;
if (X ≤ 1)
then Y := 3
else Z := 4
end
]=> (Z !-> 4 ; X !-> 2).
Proof.
(* We supply the intermediate state st'... *)
apply E_Seq with (X !-> 2).
- apply E_Asgn. reflexivity.
- apply E_IfFalse. reflexivity. apply E_Asgn. reflexivity.
Qed.
In the first step of the proof, we had to explicitly provide the intermediate state X !-> 2, due to the "hidden" argument st' to the E_Seq constructor:
E_Seq : ∀ c1 c2 st st' st'',
st =[ c1 ]=> st' →
st' =[ c2 ]=> st'' →
st =[ c1 ; c2 ]=> st''
With eapply, we can eliminate this silliness:
Example ceval'_example1:
empty_st =[
X := 2;
if (X ≤ 1)
then Y := 3
else Z := 4
end
]=> (Z !-> 4 ; X !-> 2).
Proof.
eapply E_Seq. (* 1 *)
- apply E_Asgn. (* 2 *)
reflexivity. (* 3 *)
- (* 4 *) apply E_IfFalse. reflexivity. apply E_Asgn. reflexivity.
Qed.
empty_st =[
X := 2;
if (X ≤ 1)
then Y := 3
else Z := 4
end
]=> (Z !-> 4 ; X !-> 2).
Proof.
eapply E_Seq. (* 1 *)
- apply E_Asgn. (* 2 *)
reflexivity. (* 3 *)
- (* 4 *) apply E_IfFalse. reflexivity. apply E_Asgn. reflexivity.
Qed.
Several of the tactics that we've seen so far, including ∃,
    constructor, and auto, have similar variants. The eauto
    tactic works like auto, except that it uses eapply instead of
    apply.  Tactic info_eauto shows us which tactics eauto uses
    in its proof search.
 
    Below is an example of eauto.  Before using it, we need to give
    some hints to auto about using the constructors of ceval
    and the definitions of state and total_map as part of its
    proof search.
Hint Constructors ceval : core.
Hint Transparent state total_map : core.
Example eauto_example : ∃ s',
(Y !-> 1 ; X !-> 2) =[
if (X ≤ Y)
then Z := Y - X
else Y := X + Z
end
]=> s'.
Proof. info_eauto. Qed.
Hint Transparent state total_map : core.
Example eauto_example : ∃ s',
(Y !-> 1 ; X !-> 2) =[
if (X ≤ Y)
then Z := Y - X
else Y := X + Z
end
]=> s'.
Proof. info_eauto. Qed.
