aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
-rw-r--r--backend/Duplicate.v6
-rw-r--r--backend/Duplicateproof.v4
2 files changed, 9 insertions, 1 deletions
diff --git a/backend/Duplicate.v b/backend/Duplicate.v
index c313e3fa..ec03009d 100644
--- a/backend/Duplicate.v
+++ b/backend/Duplicate.v
@@ -66,7 +66,6 @@ Remark builtin_res_eq_pos: forall (a b: builtin_res positive), {a=b} + {a<>b}.
Proof. intros. apply (builtin_res_eq Pos.eq_dec). Qed.
Global Opaque builtin_res_eq_pos.
-
Definition verify_match_inst revmap inst tinst :=
match inst with
| Inop n => match tinst with Inop n' => do u <- verify_is_copy revmap n n'; OK tt | _ => Error(msg "verify_match_inst Inop") end
@@ -127,6 +126,11 @@ Definition verify_match_inst revmap inst tinst :=
else Error (msg "Different lbar in Ibuiltin")
else Error (msg "Different ef in Ibuiltin")
| _ => Error (msg "verify_match_inst Ibuiltin") end
+ | Ireturn or => match tinst with
+ | Ireturn or' =>
+ if (option_eq Pos.eq_dec or or') then OK tt
+ else Error (msg "Different or in Ireturn")
+ | _ => Error (msg "verify_match_inst Ireturn") end
| _ => Error(msg "not implemented")
end.
diff --git a/backend/Duplicateproof.v b/backend/Duplicateproof.v
index 7369c3ea..37c583df 100644
--- a/backend/Duplicateproof.v
+++ b/backend/Duplicateproof.v
@@ -121,6 +121,10 @@ Proof.
destruct (builtin_res_eq_pos _ _); try discriminate.
eapply verify_is_copy_correct_one. destruct x. eassumption. subst.
constructor.
+(* Ireturn *)
+ - destruct i'; try (inversion H; fail).
+ destruct (option_eq _ _ _); try discriminate. subst. clear H.
+ constructor.
Qed.
Lemma verify_mapping_mn_correct: