mirror of
https://github.com/achlipala/frap.git
synced 2024-11-10 00:07:51 +00:00
Set simplification for ModelChecking
This commit is contained in:
parent
c88ae6d484
commit
4371d08696
3 changed files with 97 additions and 6 deletions
2
Frap.v
2
Frap.v
|
@ -43,7 +43,7 @@ Ltac invert0 e := invert e; fail.
|
|||
Ltac invert1 e := invert0 e || (invert e; []).
|
||||
Ltac invert2 e := invert1 e || (invert e; [|]).
|
||||
|
||||
Ltac simplify := repeat progress (simpl in *; intros; try autorewrite with core in *).
|
||||
Ltac simplify := repeat progress (simpl in *; intros; try autorewrite with core in *); repeat removeDups.
|
||||
Ltac propositional := intuition idtac.
|
||||
|
||||
Ltac linear_arithmetic := intros;
|
||||
|
|
|
@ -190,9 +190,6 @@ Proof.
|
|||
apply oneStepClosure_split; simplify.
|
||||
invert H; simplify.
|
||||
apply singleton_in.
|
||||
apply oneStepClosure_split; simplify.
|
||||
invert H; simplify.
|
||||
apply singleton_in.
|
||||
apply oneStepClosure_empty.
|
||||
simplify.
|
||||
|
||||
|
@ -202,10 +199,55 @@ Proof.
|
|||
propositional; invert H0; try equality.
|
||||
invert H; equality.
|
||||
invert H1; equality.
|
||||
invert H; equality.
|
||||
invert H; try equality.
|
||||
|
||||
simplify.
|
||||
propositional; subst; simplify; propositional.
|
||||
(* ^-- *)
|
||||
Qed.
|
||||
|
||||
Hint Rewrite fact_init_is.
|
||||
|
||||
Ltac model_check_done :=
|
||||
apply MscDone; apply prove_oneStepClosure; simplify; propositional; subst;
|
||||
repeat match goal with
|
||||
| [ H : _ |- _ ] => invert H
|
||||
end; simplify; equality.
|
||||
|
||||
Ltac model_check_step :=
|
||||
eapply MscStep; [
|
||||
repeat ((apply oneStepClosure_empty; simplify)
|
||||
|| (apply oneStepClosure_split; [ simplify;
|
||||
repeat match goal with
|
||||
| [ H : _ |- _ ] => invert H
|
||||
end; apply singleton_in | ]))
|
||||
| simplify ].
|
||||
|
||||
Ltac model_check_steps1 := model_check_done || model_check_step.
|
||||
Ltac model_check_steps := repeat model_check_steps1.
|
||||
|
||||
Ltac model_check_finish := simplify; propositional; subst; simplify; equality.
|
||||
|
||||
Ltac model_check_find_invariant :=
|
||||
simplify; eapply invariantFor_weaken; [
|
||||
apply multiStepClosure_ok; simplify; model_check_steps
|
||||
| ].
|
||||
|
||||
Ltac model_check := model_check_find_invariant; model_check_finish.
|
||||
|
||||
Theorem factorial_ok_2_snazzy :
|
||||
invariantFor (factorial_sys 2) (fact_correct 2).
|
||||
Proof.
|
||||
model_check.
|
||||
Qed.
|
||||
|
||||
Theorem factorial_ok_3 :
|
||||
invariantFor (factorial_sys 3) (fact_correct 3).
|
||||
Proof.
|
||||
model_check.
|
||||
Qed.
|
||||
|
||||
Theorem factorial_ok_4 :
|
||||
invariantFor (factorial_sys 4) (fact_correct 4).
|
||||
Proof.
|
||||
model_check.
|
||||
Qed.
|
||||
|
|
49
Sets.v
49
Sets.v
|
@ -126,3 +126,52 @@ End properties.
|
|||
Hint Resolve subseteq_refl subseteq_In.
|
||||
|
||||
Hint Rewrite union_constant.
|
||||
|
||||
|
||||
(** * Removing duplicates from constant sets *)
|
||||
|
||||
Inductive removeDups A : list A -> list A -> Prop :=
|
||||
| RdNil : removeDups nil nil
|
||||
| RdNew : forall x ls ls',
|
||||
~List.In x ls
|
||||
-> removeDups ls ls'
|
||||
-> removeDups (x :: ls) (x :: ls')
|
||||
| RdDup : forall x ls ls',
|
||||
List.In x ls
|
||||
-> removeDups ls ls'
|
||||
-> removeDups (x :: ls) ls'.
|
||||
|
||||
Theorem removeDups_fwd : forall A x (ls ls' : list A),
|
||||
removeDups ls ls'
|
||||
-> List.In x ls
|
||||
-> List.In x ls'.
|
||||
Proof.
|
||||
induction 1; simpl; intuition.
|
||||
subst; eauto.
|
||||
Qed.
|
||||
|
||||
Theorem removeDups_bwd : forall A x (ls ls' : list A),
|
||||
removeDups ls ls'
|
||||
-> List.In x ls'
|
||||
-> List.In x ls.
|
||||
Proof.
|
||||
induction 1; simpl; intuition.
|
||||
Qed.
|
||||
|
||||
Theorem removeDups_ok : forall A (ls ls' : list A),
|
||||
removeDups ls ls'
|
||||
-> constant ls = constant ls'.
|
||||
Proof.
|
||||
intros.
|
||||
apply sets_equal.
|
||||
unfold constant; intuition eauto using removeDups_fwd, removeDups_bwd.
|
||||
Qed.
|
||||
|
||||
Ltac removeDups :=
|
||||
match goal with
|
||||
| [ |- context[constant ?ls] ] =>
|
||||
erewrite (@removeDups_ok _ ls)
|
||||
by repeat (apply RdNil
|
||||
|| (apply RdNew; [ simpl; intuition congruence | ])
|
||||
|| (apply RdDup; [ simpl; intuition congruence | ]))
|
||||
end.
|
||||
|
|
Loading…
Reference in a new issue