diff paper/src/agda-hoare-soundness.agda @ 19:046b2b20d6c7 default tip

fix
author ryokka
date Mon, 09 Mar 2020 11:25:49 +0900
parents
children
line wrap: on
line diff
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/paper/src/agda-hoare-soundness.agda	Mon Mar 09 11:25:49 2020 +0900
@@ -0,0 +1,66 @@
+Soundness : {bPre : Cond} -> {cm : Comm} -> {bPost : Cond} ->
+            HTProof bPre cm bPost -> Satisfies bPre cm bPost
+Soundness (PrimRule {bPre} {cm} {bPost} pr) s1 s2 q1 q2
+  = axiomValid bPre cm bPost pr s1 s2 q1 q2
+Soundness {.bPost} {.Skip} {bPost} (SkipRule .bPost) s1 s2 q1 q2
+  = substId1 State {Level.zero} {State} {s1} {s2} (proj₂ q2) (SemCond bPost) q1
+Soundness {bPre} {.Abort} {bPost} (AbortRule .bPre .bPost) s1 s2 q1 ()
+Soundness (WeakeningRule {bPre} {bPre'} {cm} {bPost'} {bPost} tautPre pr tautPost)
+          s1 s2 q1 q2
+  = let hyp : Satisfies bPre' cm bPost'
+        hyp = Soundness pr
+    in tautValid bPost' bPost tautPost s2 (hyp s1 s2 (tautValid bPre bPre' tautPre s1 q1) q2)
+Soundness (SeqRule {bPre} {cm1} {bMid} {cm2} {bPost} pr1 pr2)
+           s1 s2 q1 q2
+  = let hyp1 : Satisfies bPre cm1 bMid
+        hyp1 = Soundness pr1
+        hyp2 : Satisfies bMid cm2 bPost
+        hyp2 = Soundness pr2
+    in hyp2 (proj₁ q2) s2 (hyp1 s1 (proj₁ q2) q1 (proj₁ (proj₂ q2))) (proj₂ (proj₂ q2))
+Soundness (IfRule {cmThen} {cmElse} {bPre} {bPost} {b} pThen pElse)
+          s1 s2 q1 q2
+  = let hypThen : Satisfies (bPre /\ b) cmThen bPost
+        hypThen = Soundness pThen
+        hypElse : Satisfies (bPre /\ neg b) cmElse bPost
+        hypElse = Soundness pElse
+        rThen : RelOpState.comp (RelOpState.delta (SemCond b))
+                  (SemComm cmThen) s1 s2 -> SemCond bPost s2
+        rThen = λ h -> hypThen s1 s2 ((proj₂ (respAnd bPre b s1)) (q1 , proj₁ t1))
+          (proj₂ ((proj₂ (RelOpState.deltaRestPre (SemCond b) (SemComm cmThen) s1 s2)) h))
+        rElse : RelOpState.comp (RelOpState.delta (NotP (SemCond b)))
+                  (SemComm cmElse) s1 s2 -> SemCond bPost s2
+        rElse = λ h ->
+                  let t10 : (NotP (SemCond b) s1) × (SemComm cmElse s1 s2)
+                      t10 = proj₂ (RelOpState.deltaRestPre
+                                  (NotP (SemCond b)) (SemComm cmElse) s1 s2) h
+                  in hypElse s1 s2 (proj₂ (respAnd bPre (neg b) s1)
+                             (q1 , (proj₂ (respNeg b s1) (proj₁ t10)))) (proj₂ t10)
+    in when rThen rElse q2
+Soundness (WhileRule {cm'} {bInv} {b} pr) s1 s2 q1 q2
+  = proj₂ (respAnd bInv (neg b) s2)
+          (lem1 (proj₁ q2) s2 (proj₁ t15) , proj₂ (respNeg b s2) (proj₂ t15))
+    where
+      hyp : Satisfies (bInv /\ b) cm' bInv
+      hyp = Soundness pr
+      Rel1 : ℕ -> Rel State (Level.zero)
+      Rel1 = λ m ->
+               RelOpState.repeat
+                 m
+                 (RelOpState.comp (RelOpState.delta (SemCond b))
+                                  (SemComm cm'))
+      t15 : (Rel1 (proj₁ q2) s1 s2) × (NotP (SemCond b) s2)
+      t15 = proj₂ (RelOpState.deltaRestPost
+        (NotP (SemCond b)) (Rel1 (proj₁ q2)) s1 s2) (proj₂ q2)
+      lem1 : (m : ℕ) -> (ss2 : State) -> Rel1 m s1 ss2 -> SemCond bInv ss2
+      lem1 zero ss2 h = substId1 State (proj₂ h) (SemCond bInv) q1
+      lem1 (suc n) ss2 h
+    = let hyp2 : (z : State) -> Rel1 (proj₁ q2) s1 z ->
+                     SemCond bInv z
+              hyp2 = lem1 n
+              t22 : (SemCond b (proj₁ h)) × (SemComm cm' (proj₁ h) ss2)
+              t22 = proj₂ (RelOpState.deltaRestPre (SemCond b) (SemComm cm') (proj₁ h) ss2)
+                    (proj₂ (proj₂ h))
+              t23 : SemCond (bInv /\ b) (proj₁ h)
+              t23 = proj₂ (respAnd bInv b (proj₁ h))
+                (hyp2 (proj₁ h) (proj₁ (proj₂ h)) , proj₁ t22)
+      in hyp (proj₁ h) ss2 t23 (proj₂ t22)