frap/Frap.v

135 lines
4.3 KiB
Coq
Raw Normal View History

2016-02-21 17:16:31 +00:00
Require Import String Arith Omega Program Sets Relations Map Var Invariant Bool ModelCheck.
Export String Arith Sets Relations Map Var Invariant Bool ModelCheck.
2015-12-31 20:44:34 +00:00
Require Import List.
2016-02-02 17:38:00 +00:00
Export List ListNotations.
2015-12-31 20:44:34 +00:00
Open Scope string_scope.
2016-02-06 23:24:06 +00:00
Open Scope list_scope.
2015-12-31 20:44:34 +00:00
Ltac inductN n :=
match goal with
| [ |- forall x : ?E, _ ] =>
match type of E with
| Prop =>
let H := fresh in intro H;
match n with
| 1 => dependent induction H
| S ?n' => inductN n'
end
| _ => intro; inductN n
end
end.
Ltac induct e := inductN e || dependent induction e.
Ltac invert' H := inversion H; clear H; subst.
Ltac invertN n :=
match goal with
| [ |- forall x : ?E, _ ] =>
match type of E with
| Prop =>
let H := fresh in intro H;
match n with
| 1 => invert' H
| S ?n' => invertN n'
end
| _ => intro; invertN n
end
end.
Ltac invert e := invertN e || invert' e.
Ltac invert0 e := invert e; fail.
Ltac invert1 e := invert0 e || (invert e; []).
Ltac invert2 e := invert1 e || (invert e; [|]).
2016-02-18 22:51:58 +00:00
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).
2016-02-09 18:11:58 +00:00
Ltac propositional := intuition idtac.
2015-12-31 20:44:34 +00:00
2016-01-17 01:32:12 +00:00
Ltac linear_arithmetic := intros;
2015-12-31 20:44:34 +00:00
repeat match goal with
| [ |- context[max ?a ?b] ] =>
let Heq := fresh "Heq" in destruct (Max.max_spec a b) as [[? Heq] | [? Heq]];
rewrite Heq in *; clear Heq
| [ _ : context[max ?a ?b] |- _ ] =>
let Heq := fresh "Heq" in destruct (Max.max_spec a b) as [[? Heq] | [? Heq]];
rewrite Heq in *; clear Heq
| [ |- context[min ?a ?b] ] =>
let Heq := fresh "Heq" in destruct (Min.min_spec a b) as [[? Heq] | [? Heq]];
rewrite Heq in *; clear Heq
| [ _ : context[min ?a ?b] |- _ ] =>
let Heq := fresh "Heq" in destruct (Min.min_spec a b) as [[? Heq] | [? Heq]];
rewrite Heq in *; clear Heq
end; omega.
2016-02-09 21:28:48 +00:00
Ltac equality := intuition congruence.
2015-12-31 20:44:34 +00:00
Ltac cases E :=
2016-02-06 23:24:06 +00:00
((is_var E; destruct E)
|| match type of E with
| {_} + {_} => destruct E
| _ => let Heq := fresh "Heq" in destruct E eqn:Heq
end);
repeat match goal with
| [ H : _ = left _ |- _ ] => clear H
| [ H : _ = right _ |- _ ] => clear H
end.
2016-01-17 01:32:12 +00:00
Global Opaque max min.
2016-02-07 03:09:37 +00:00
Infix "==n" := eq_nat_dec (no associativity, at level 50).
Export Frap.Map.
Ltac maps_equal := Frap.Map.M.maps_equal; simplify.
2016-02-16 00:39:36 +00:00
Ltac first_order := firstorder idtac.
2016-02-21 17:16:31 +00:00
(** * Model checking *)
Ltac model_check_done :=
apply MscDone; apply prove_oneStepClosure; simplify; propositional; subst;
repeat match goal with
2016-02-21 22:00:01 +00:00
| [ H : ?P |- _ ] =>
match type of P with
| Prop => invert H; simplify
end
end; equality.
2016-02-21 17:16:31 +00:00
Ltac singletoner :=
repeat match goal with
| _ => apply singleton_in
| [ |- (_ \cup _) _ ] => apply singleton_in_other
end.
Ltac model_check_step :=
eapply MscStep; [
repeat (apply oneStepClosure_empty
2016-02-21 17:16:31 +00:00
|| (apply oneStepClosure_split; [ simplify;
repeat match goal with
2016-02-21 22:00:01 +00:00
| [ H : ?P |- _ ] =>
match type of P with
| Prop => invert H; simplify; try congruence
end
2016-02-21 17:16:31 +00:00
end; solve [ singletoner ] | ]))
| 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_infer :=
apply multiStepClosure_ok; simplify; model_check_steps.
Ltac model_check_find_invariant :=
simplify; eapply invariant_weaken; [ model_check_infer | ]; cbv beta in *.
Ltac model_check := model_check_find_invariant; model_check_finish.