annotate WhileTest.agda @ 13:724766af8b12

add impl
author soto
date Thu, 11 Feb 2021 19:33:35 +0900
parents 77f0530f2eff
children
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
rev   line source
11
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
1 open import Level renaming (suc to Suc ; zero to Zero )
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
2 module WhileTest where
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
3
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
4 open import Relation.Binary.PropositionalEquality
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
5 open import Relation.Binary.Core
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
6 open import Data.Nat hiding (compare)
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
7 open import Data.Maybe
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
8 open import Data.List
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
9 open import Function
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
10 open import logic
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
11
13
724766af8b12 add impl
soto
parents: 12
diff changeset
12 open import Data.Bool hiding ( _∧_ ; _≟_ ) -- ; _∧_ ; _≤_ ; _<_)
724766af8b12 add impl
soto
parents: 12
diff changeset
13 open import Data.Product
11
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
14 open import Agda.Builtin.Unit
13
724766af8b12 add impl
soto
parents: 12
diff changeset
15 open import Relation.Nullary using (¬_; Dec; yes; no)
724766af8b12 add impl
soto
parents: 12
diff changeset
16
724766af8b12 add impl
soto
parents: 12
diff changeset
17 open import Data.Empty
724766af8b12 add impl
soto
parents: 12
diff changeset
18 open import Data.Nat.Properties
724766af8b12 add impl
soto
parents: 12
diff changeset
19
724766af8b12 add impl
soto
parents: 12
diff changeset
20 -- logicへ
724766af8b12 add impl
soto
parents: 12
diff changeset
21 +zero : { y : ℕ } → y + zero ≡ y
724766af8b12 add impl
soto
parents: 12
diff changeset
22 +zero {zero} = refl
724766af8b12 add impl
soto
parents: 12
diff changeset
23 +zero {suc y} = cong ( λ x → suc x ) ( +zero {y} )
724766af8b12 add impl
soto
parents: 12
diff changeset
24
724766af8b12 add impl
soto
parents: 12
diff changeset
25 -- utilへ
724766af8b12 add impl
soto
parents: 12
diff changeset
26 _-_ : ℕ → ℕ → ℕ
724766af8b12 add impl
soto
parents: 12
diff changeset
27 x - zero = x
724766af8b12 add impl
soto
parents: 12
diff changeset
28 zero - _ = zero
724766af8b12 add impl
soto
parents: 12
diff changeset
29 (suc x) - (suc y) = x - y
11
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
30
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
31 record Env : Set (Suc Zero) where
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
32 field
13
724766af8b12 add impl
soto
parents: 12
diff changeset
33 c10 : ℕ
11
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
34 varn : ℕ
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
35 vari : ℕ
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
36 open Env
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
37
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
38 data _implies_ (A B : Set ) : Set (Suc Zero) where
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
39 proof : ( A → B ) → A implies B
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
40
13
724766af8b12 add impl
soto
parents: 12
diff changeset
41 implies2p : {A B : Set } → A implies B → A → B
724766af8b12 add impl
soto
parents: 12
diff changeset
42 implies2p (proof x) = x
724766af8b12 add impl
soto
parents: 12
diff changeset
43
11
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
44 data whileTestState : Set where
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
45 s1 : whileTestState
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
46 s2 : whileTestState
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
47 sf : whileTestState
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
48
13
724766af8b12 add impl
soto
parents: 12
diff changeset
49 whileTestStateP : whileTestState → Env → Set
724766af8b12 add impl
soto
parents: 12
diff changeset
50 whileTestStateP s1 env = (vari env ≡ 0) ∧ (varn env ≡ c10 env)
724766af8b12 add impl
soto
parents: 12
diff changeset
51 whileTestStateP s2 env = (varn env + vari env ≡ c10 env)
724766af8b12 add impl
soto
parents: 12
diff changeset
52 whileTestStateP sf env = (vari env ≡ c10 env)
11
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
53
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
54 record WhileTest {m : Level } {t : Set m } : Set (Suc m) where
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
55 field
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
56 env : Env
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
57 whileInit : {m : Level } {t : Set m } → (c10 : ℕ) → (Env → t) → t
13
724766af8b12 add impl
soto
parents: 12
diff changeset
58 whileInit c10 next = next (record {c10 = c10 ; varn = c10 ; vari = 0 } )
11
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
59 whileLoop : Env → (Code : Env → t) → t
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
60 whileLoop env next = whileLoop1 (varn env) env where
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
61 whileLoop1 : ℕ → Env → t
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
62 whileLoop1 zero env = next env
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
63 whileLoop1 (suc t ) env = whileLoop1 t (record env {varn = t ; vari = (vari env) + 1})
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
64 whileTest : (c10 : ℕ) → (Env → t) → t
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
65 whileTest c10 next = whileInit c10 $ λ env → whileLoop env next
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
66
13
724766af8b12 add impl
soto
parents: 12
diff changeset
67 loopPP : (n : ℕ) → (input : Env ) → (n ≡ varn input) → Env
724766af8b12 add impl
soto
parents: 12
diff changeset
68 loopPP zero input refl = input
724766af8b12 add impl
soto
parents: 12
diff changeset
69 loopPP (suc n) input refl = loopPP n (record input { varn = pred (varn input) ; vari = suc (vari input)}) refl
724766af8b12 add impl
soto
parents: 12
diff changeset
70
724766af8b12 add impl
soto
parents: 12
diff changeset
71 -- init
724766af8b12 add impl
soto
parents: 12
diff changeset
72 whileInit-impl : (c10 : ℕ) → whileInit c10 (λ env → ⊤ implies (whileTestStateP s1 env) )
724766af8b12 add impl
soto
parents: 12
diff changeset
73 whileInit-impl c = proof ( λ _ → record { proj1 = refl ; proj2 = refl } )
724766af8b12 add impl
soto
parents: 12
diff changeset
74 whileTestPSemSound : (c : ℕ ) (output : Env ) → output ≡ whileInit c (λ e → e) → ⊤ implies ((vari output ≡ 0) ∧ (varn output ≡ c))
724766af8b12 add impl
soto
parents: 12
diff changeset
75 whileTestPSemSound c output refl = whileInit-impl c
724766af8b12 add impl
soto
parents: 12
diff changeset
76 -- init → loop
724766af8b12 add impl
soto
parents: 12
diff changeset
77 whileConvPSemSound : {l : Level} → (input : Env) → (whileTestStateP s1 input ) implies (whileTestStateP s2 input)
724766af8b12 add impl
soto
parents: 12
diff changeset
78 whileConvPSemSound input = proof λ x → (conv input x) where
724766af8b12 add impl
soto
parents: 12
diff changeset
79 conv : (env : Env ) → (vari env ≡ 0) ∧ (varn env ≡ c10 env ) → varn env + vari env ≡ c10 env
724766af8b12 add impl
soto
parents: 12
diff changeset
80 conv e record { proj1 = refl ; proj2 = refl } = +zero
724766af8b12 add impl
soto
parents: 12
diff changeset
81 -- loop → loop
724766af8b12 add impl
soto
parents: 12
diff changeset
82 whileLoopPSem : {l : Level} {t : Set l} → (input : Env ) → whileTestStateP s2 input
724766af8b12 add impl
soto
parents: 12
diff changeset
83 → (next : (output : Env ) → (whileTestStateP s2 input ) implies (whileTestStateP s2 output) → t)
724766af8b12 add impl
soto
parents: 12
diff changeset
84 → (exit : (output : Env ) → (whileTestStateP s2 input ) implies (whileTestStateP sf output) → t) → t
724766af8b12 add impl
soto
parents: 12
diff changeset
85 whileLoopPSem env s next exit with varn env | s
724766af8b12 add impl
soto
parents: 12
diff changeset
86 ... | zero | _ = exit env (proof (λ z → z))
724766af8b12 add impl
soto
parents: 12
diff changeset
87 ... | (suc varn ) | refl = next ( record env { varn = varn ; vari = suc (vari env) } ) (proof λ x → +-suc varn (vari env) )
724766af8b12 add impl
soto
parents: 12
diff changeset
88 -- loop → fin
724766af8b12 add impl
soto
parents: 12
diff changeset
89
724766af8b12 add impl
soto
parents: 12
diff changeset
90 loopPPSem : (input output : Env ) → output ≡ loopPP (varn input) input refl
724766af8b12 add impl
soto
parents: 12
diff changeset
91 → (whileTestStateP s2 input ) → (whileTestStateP s2 input ) implies (whileTestStateP sf output)
724766af8b12 add impl
soto
parents: 12
diff changeset
92 loopPPSem input output refl s2p = loopPPSemInduct (varn input) input refl refl s2p
724766af8b12 add impl
soto
parents: 12
diff changeset
93 where
724766af8b12 add impl
soto
parents: 12
diff changeset
94 lem : (n : ℕ) → (env : Env) → n + suc (vari env) ≡ suc (n + vari env)
724766af8b12 add impl
soto
parents: 12
diff changeset
95 lem n env = +-suc (n) (vari env)
724766af8b12 add impl
soto
parents: 12
diff changeset
96 loopPPSemInduct : (n : ℕ) → (current : Env) → (eq : n ≡ varn current) → (loopeq : output ≡ loopPP n current eq)
724766af8b12 add impl
soto
parents: 12
diff changeset
97 → (whileTestStateP s2 current ) → (whileTestStateP s2 current ) implies (whileTestStateP sf output)
724766af8b12 add impl
soto
parents: 12
diff changeset
98 loopPPSemInduct zero current refl loopeq refl rewrite loopeq = proof (λ x → refl)
724766af8b12 add impl
soto
parents: 12
diff changeset
99 loopPPSemInduct (suc n) current refl loopeq refl rewrite (sym (lem n current)) =
724766af8b12 add impl
soto
parents: 12
diff changeset
100 whileLoopPSem current refl
724766af8b12 add impl
soto
parents: 12
diff changeset
101 (λ output x → loopPPSemInduct n (record { c10 = n + suc (vari current) ; varn = n ; vari = suc (vari current) }) refl loopeq refl)
724766af8b12 add impl
soto
parents: 12
diff changeset
102 (λ output x → loopPPSemInduct n (record { c10 = n + suc (vari current) ; varn = n ; vari = suc (vari current) }) refl loopeq refl)
724766af8b12 add impl
soto
parents: 12
diff changeset
103
724766af8b12 add impl
soto
parents: 12
diff changeset
104
724766af8b12 add impl
soto
parents: 12
diff changeset
105 whileLoopPSemSound : {l : Level} → (input output : Env )
724766af8b12 add impl
soto
parents: 12
diff changeset
106 → whileTestStateP s2 input
724766af8b12 add impl
soto
parents: 12
diff changeset
107 → output ≡ loopPP (varn input) input refl
724766af8b12 add impl
soto
parents: 12
diff changeset
108 → (whileTestStateP s2 input ) implies ( whileTestStateP sf output )
724766af8b12 add impl
soto
parents: 12
diff changeset
109 whileLoopPSemSound {l} input output pre eq = loopPPSem input output eq pre
724766af8b12 add impl
soto
parents: 12
diff changeset
110
724766af8b12 add impl
soto
parents: 12
diff changeset
111 whileTestSound : {l : Level} (c : ℕ) → (output : Env) → ⊤ → whileTestStateP sf output
724766af8b12 add impl
soto
parents: 12
diff changeset
112 whileTestSound {l} c record { c10 = c10 ; varn = varn ; vari = vari } top =
724766af8b12 add impl
soto
parents: 12
diff changeset
113 implies2p (whileLoopPSemSound {l} (record { c10 = c ; varn = c ; vari = zero }) (record { c10 = c10 ; varn = c ; vari = vari}) (+zero) {!!}) (implies2p (whileConvPSemSound {l} (record { c10 = c ; varn = c ; vari = zero })) (implies2p (whileTestPSemSound c (whileInit c (λ e → e)) refl) top))
724766af8b12 add impl
soto
parents: 12
diff changeset
114
11
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
115 open WhileTest
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
116
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
117 createWhileTest : {m : Level} {t : Set m } → WhileTest {m} {t}
13
724766af8b12 add impl
soto
parents: 12
diff changeset
118 createWhileTest = record { env = record { c10 = 0; varn = 0; vari = 0 } }
11
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
119
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
120 test2 : ℕ
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
121 test2 = whileTest createWhileTest 10 $ λ e → vari e
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
122
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
123 ---
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
124
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
125 {-
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
126 whileTestStateP : whileTestState → Envc → Set
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
127 whileTestStateP s1 env = (vari env ≡ 0) /\ (varn env ≡ c10 env)
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
128 whileTestStateP s2 env = (varn env + vari env ≡ c10 env)
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
129 whileTestStateP sf env = (vari env ≡ c10 env)
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
130
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
131 whileTestP : {l : Level} {t : Set l} → (c10 : ℕ) → (Code : Envc → t) → t
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
132 whileTestP c10 next = next (record {varn = c10 ; vari = 0 ; c10 = c10 } )
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
133
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
134 whileTestPSem : (c : ℕ) → whileTestP c ( λ env → ⊤ implies (whileTestStateP s1 env) )
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
135 whileTestPSem c = proof ( λ _ → record { pi1 = refl ; pi2 = refl } )
e6a7215fb00c add while_init_imple
soto
parents:
diff changeset
136 -}