Skip to content

Commit

Permalink
moving tactics to separate .v
Browse files Browse the repository at this point in the history
  • Loading branch information
vzaliva committed Nov 1, 2023
1 parent 0c45da2 commit 3c6f213
Show file tree
Hide file tree
Showing 3 changed files with 46 additions and 46 deletions.
47 changes: 2 additions & 45 deletions coq/Proofs/Revocation.v
Original file line number Diff line number Diff line change
Expand Up @@ -29,6 +29,8 @@ From Morello Require Import CapabilitiesGS MorelloCapsGS.

From CheriMemory Require Import CheriMorelloMemory Memory_model CoqMem_common ErrorWithState CoqUndefined ErrorWithState CoqLocation CoqSymbol CoqImplementation CoqTags CoqSwitches CerbSwitches CoqAilTypesAux.

Require Import Tactics.

Local Open Scope string_scope.
Local Open Scope type_scope.
Local Open Scope list_scope.
Expand All @@ -54,18 +56,6 @@ Module AbstTagDefs: TagDefs.
Definition tagDefs := abst_tagDefs.
End AbstTagDefs.

(* This (seemingly) prevents slow `Qed` problem.
Usage: replace `cbn in H` with `cbn_hyp H`.
See aslo:
- discussion: https://coq.zulipchat.com/#narrow/stream/237977-Coq-users/topic/Coq.20hangs.20on.20Qed
- FAQ: https://github.com/coq/coq/wiki/Troubleshooting#what-can-i-do-when-qed-is-slow
TODO: move to separate file with other tactics we define
*)
Ltac cbn_hyp H :=
let T := type of H in let T' := eval cbn in T in replace T with T' in H by (cbn;reflexivity).

Module RevocationProofs.

(* --- Memory models instantiated with and without PNVI --- *)
Expand Down Expand Up @@ -1140,39 +1130,6 @@ Module RevocationProofs.

End allocate_region_proofs.

(* TODO : move *)
Ltac inl_inr :=
match goal with
| [H1: inl _ = inr _ |- _] => inversion H1
| [H1: inr _ = inl _ |- _] => inversion H1

| [H1: inl _ = Monad.ret _ |- _] => inversion H1
| [H1: Monad.ret _ = inl _ |- _] => inversion H1

| [H1: raise _ = inr _ |- _] => inversion H1
| [H1: inr _ = raise _ |- _] => inversion H1

| [ |- inl ?x = inl ?x ] => reflexivity
| [ |- inr ?x = inr ?x ] => reflexivity

| [H1: ?a = inr _,
H2: ?a = inl _ |- _] => rewrite H1 in H2;
inversion H2

end.

(* TODO : move *)
Ltac inl_inr_inv :=
match goal with
| [H1: inl _ = inl _ |- _] => inversion H1; clear H1
| [H1: inr _ = inr _ |- _] => inversion H1; clear H1

| [H1: inr _ = Monad.ret _ |- _] => inversion H1; clear H1
| [H1: Monad.ret _ = inr _ |- _] => inversion H1; clear H1
| [H1: raise _ = inl _ |- _] => inversion H1; clear H1
| [H1: inl _ = raise _ |- _] => inversion H1; clear H1
end.

Lemma bind_Same_eq {T1 T2 T:Type}
{R: T1 -> T2 -> Prop} (* relation between values *)
{M1: CheriMemoryWithPNVI.memM T}
Expand Down
43 changes: 43 additions & 0 deletions coq/Proofs/Tactics.v
Original file line number Diff line number Diff line change
@@ -0,0 +1,43 @@
From ExtLib.Structures Require Import Monad Monads MonadExc MonadState Traversable.
From ExtLib.Data.Monads Require Import EitherMonad OptionMonad.

(* This (seemingly) prevents slow `Qed` problem.
Usage: replace `cbn in H` with `cbn_hyp H`.
See aslo:
- discussion: https://coq.zulipchat.com/#narrow/stream/237977-Coq-users/topic/Coq.20hangs.20on.20Qed
- FAQ: https://github.com/coq/coq/wiki/Troubleshooting#what-can-i-do-when-qed-is-slow
*)
Ltac cbn_hyp H :=
let T := type of H in let T' := eval cbn in T in replace T with T' in H by (cbn;reflexivity).

Ltac inl_inr :=
match goal with
| [H1: inl _ = inr _ |- _] => inversion H1
| [H1: inr _ = inl _ |- _] => inversion H1

| [H1: inl _ = Monad.ret _ |- _] => inversion H1
| [H1: Monad.ret _ = inl _ |- _] => inversion H1

| [H1: raise _ = inr _ |- _] => inversion H1
| [H1: inr _ = raise _ |- _] => inversion H1

| [ |- inl ?x = inl ?x ] => reflexivity
| [ |- inr ?x = inr ?x ] => reflexivity

| [H1: ?a = inr _,
H2: ?a = inl _ |- _] => rewrite H1 in H2;
inversion H2

end.

Ltac inl_inr_inv :=
match goal with
| [H1: inl _ = inl _ |- _] => inversion H1; clear H1
| [H1: inr _ = inr _ |- _] => inversion H1; clear H1

| [H1: inr _ = Monad.ret _ |- _] => inversion H1; clear H1
| [H1: Monad.ret _ = inr _ |- _] => inversion H1; clear H1
| [H1: raise _ = inl _ |- _] => inversion H1; clear H1
| [H1: inl _ = raise _ |- _] => inversion H1; clear H1
end.
2 changes: 1 addition & 1 deletion coq/Proofs/dune
Original file line number Diff line number Diff line change
Expand Up @@ -2,6 +2,6 @@
(mode vo)
(name Proofs)
(package cerberus-cheri)
(modules Revocation)
(modules Tactics Revocation)
(theories CheriMemory Common Morello)
)

0 comments on commit 3c6f213

Please sign in to comment.