annotate paper/src/whileLoopPSem.agda @ 4:bf1f62556b81

add while_test_init_imple
author soto
date Thu, 11 Feb 2021 17:03:31 +0900
parents 959f4b34d6f4
children
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
rev   line source
3
959f4b34d6f4 add final thesis
soto
parents:
diff changeset
1 whileLoopPSem : {l : Level} {t : Set l} → (input : Envc ) → (vari input) + (varn input) ≡ (c10 input)
959f4b34d6f4 add final thesis
soto
parents:
diff changeset
2 → (next : (output : Envc ) → ((vari input) + (varn input) ≡ (c10 input) ) implies ((vari output) + (varn output) ≡ (c10 output)) → t)
959f4b34d6f4 add final thesis
soto
parents:
diff changeset
3 → (exit : (output : Envc ) → ((vari input) + (varn input) ≡ (c10 input) ) implies ((vari output ≡ c10 output)) → t) → t
959f4b34d6f4 add final thesis
soto
parents:
diff changeset
4 whileLoopPSem env s next exit with varn env | s
959f4b34d6f4 add final thesis
soto
parents:
diff changeset
5 ... | zero | _ = exit env (proof (λ z → z))
959f4b34d6f4 add final thesis
soto
parents:
diff changeset
6 ... | (suc varn ) | refl = next ( record env { varn = varn ; vari = suc (vari env) } ) (proof λ x → +-suc varn (vari env) )
959f4b34d6f4 add final thesis
soto
parents:
diff changeset
7
959f4b34d6f4 add final thesis
soto
parents:
diff changeset
8
959f4b34d6f4 add final thesis
soto
parents:
diff changeset
9 loopPPSem : (input output : Envc ) → output ≡ loopPP (varn input) input refl
959f4b34d6f4 add final thesis
soto
parents:
diff changeset
10 → (vari input) + (varn input) ≡ (c10 input) → ((vari input) + (varn input) ≡ (c10 input) ) implies ((vari output ≡ c10 output))
959f4b34d6f4 add final thesis
soto
parents:
diff changeset
11 loopPPSem input output refl s2p = loopPPSemInduct (varn input) input refl refl s2p
959f4b34d6f4 add final thesis
soto
parents:
diff changeset
12 where
959f4b34d6f4 add final thesis
soto
parents:
diff changeset
13 lem : (n : ℕ) → (env : Envc) → n + suc (vari env) ≡ suc (n + vari env)
959f4b34d6f4 add final thesis
soto
parents:
diff changeset
14 lem n env = +-suc (n) (vari env)
959f4b34d6f4 add final thesis
soto
parents:
diff changeset
15 loopPPSemInduct : (n : ℕ) → (current : Envc) → (eq : n ≡ varn current) → (loopeq : output ≡ loopPP n current eq)
959f4b34d6f4 add final thesis
soto
parents:
diff changeset
16 → ((vari current) + (varn current) ≡ (c10 current) ) → ((vari current) + (varn current) ≡ (c10 current) ) implies ((vari output ≡ c10 output))
959f4b34d6f4 add final thesis
soto
parents:
diff changeset
17 loopPPSemInduct zero current refl loopeq refl rewrite loopeq = proof (λ x → refl)
959f4b34d6f4 add final thesis
soto
parents:
diff changeset
18 loopPPSemInduct (suc n) current refl loopeq refl rewrite (sym (lem n current)) =
959f4b34d6f4 add final thesis
soto
parents:
diff changeset
19 whileLoopPSem current refl
959f4b34d6f4 add final thesis
soto
parents:
diff changeset
20 (λ output x → loopPPSemInduct n (record { c10 = n + suc (vari current) ; varn = n ; vari = suc (vari current) }) refl loopeq refl)
959f4b34d6f4 add final thesis
soto
parents:
diff changeset
21 (λ output x → loopPPSemInduct n (record { c10 = n + suc (vari current) ; varn = n ; vari = suc (vari current) }) refl loopeq refl)