mirror of
https://github.com/achlipala/frap.git
synced 2024-11-10 00:07:51 +00:00
Fixed to work in Coq 8.4, too
This commit is contained in:
parent
cf65c18ebf
commit
2c601b04a0
3 changed files with 20 additions and 3 deletions
8
Frap.v
8
Frap.v
|
@ -43,8 +43,12 @@ 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 *; repeat unifyTails);
|
||||
repeat (removeDups || doSubtract).
|
||||
Ltac simplify := repeat (unifyTails; pose proof I);
|
||||
repeat match goal with
|
||||
| [ H : True |- _ ] => clear H
|
||||
end;
|
||||
repeat progress (simpl in *; intros; try autorewrite with core in *);
|
||||
repeat (removeDups || doSubtract).
|
||||
Ltac propositional := intuition idtac.
|
||||
|
||||
Ltac linear_arithmetic := intros;
|
||||
|
|
|
@ -248,7 +248,7 @@ Ltac model_check_infer :=
|
|||
apply multiStepClosure_ok; simplify; model_check_steps.
|
||||
|
||||
Ltac model_check_find_invariant :=
|
||||
simplify; eapply invariant_weaken; [ model_check_infer | ].
|
||||
simplify; eapply invariant_weaken; [ model_check_infer | ]; cbv beta in *.
|
||||
|
||||
Ltac model_check := model_check_find_invariant; model_check_finish.
|
||||
|
||||
|
|
13
Sets.v
13
Sets.v
|
@ -169,9 +169,22 @@ Proof.
|
|||
unfold constant; intuition eauto using removeDups_fwd, removeDups_bwd.
|
||||
Qed.
|
||||
|
||||
Ltac someMatch ls :=
|
||||
match ls with
|
||||
| ?x :: ?ls' =>
|
||||
let rec someMatch' ls :=
|
||||
match ls with
|
||||
| x :: _ => idtac
|
||||
| _ :: ?ls' => someMatch' ls'
|
||||
end
|
||||
in someMatch' ls'
|
||||
| _ :: ?ls' => someMatch ls'
|
||||
end.
|
||||
|
||||
Ltac removeDups :=
|
||||
match goal with
|
||||
| [ |- context[constant ?ls] ] =>
|
||||
someMatch ls;
|
||||
erewrite (@removeDups_ok _ ls)
|
||||
by repeat (apply RdNil
|
||||
|| (apply RdNew; [ simpl; intuition congruence | ])
|
||||
|
|
Loading…
Reference in a new issue