Library HoareTut.hoarelogicsemantics
Semantics of my Hoare Logic
Set Implicit Arguments.
Expression Language
- "Var A" is the type of variables storing values of type A
- "Expr A" is the type of expressions return values of type A
- Env is the type of environment. An environment is a mapping from variables into values.
- "upd x v e" sets v as the new value of x from environment e
- eval evaluates an expression in a given environment.
Module Type ExprLang.
Parameter Var: Type → Type.
Parameter Expr: Type → Type.
Parameter Env: Type.
Parameter upd: ∀ (A:Type), (Var A) → A → Env → Env.
Parameter eval: ∀ (A:Type), (Expr A) → Env → A.
End ExprLang.
The programming language: syntax and semantics
- Iskip is an instruction to "do nothing"
- "Iset x expr" stores in x the result of expr
- Iif represents conditional branching
- Iseq represents the sequence of two instructions
- Iwhile represents the while-loop
Inductive ImpProg: Type :=
| Iskip: ImpProg
| Iset (A:Type) (x:E.Var A) (expr:E.Expr A): ImpProg
| Iif (cond:E.Expr bool) (p1 p2:ImpProg): ImpProg
| Iseq (p1 p2:ImpProg): ImpProg
| Iwhile (cond:E.Expr bool) (p:ImpProg): ImpProg.
| Iskip: ImpProg
| Iset (A:Type) (x:E.Var A) (expr:E.Expr A): ImpProg
| Iif (cond:E.Expr bool) (p1 p2:ImpProg): ImpProg
| Iseq (p1 p2:ImpProg): ImpProg
| Iwhile (cond:E.Expr bool) (p:ImpProg): ImpProg.
The semantics of the programming language is given by the following natural semantics.
Here "exec e0 p e1" means that "in the initial environment e0 the execution of p
terminates in the final environment e1".
Inductive exec: E.Env → ImpProg → E.Env → Prop :=
| exec_Iskip:
∀ e, (exec e Iskip e)
| exec_Iset:
∀ (A:Type) e x (expr: E.Expr A),
(exec e (Iset x expr) (E.upd x (E.eval expr e) e))
| exec_Iif:
∀ e (cond: E.Expr bool) p1 p2 e',
(exec e (if (E.eval cond e) then p1 else p2) e')
→ (exec e (Iif cond p1 p2) e')
| exec_Iseq:
∀ e p1 p2 e' e'',
(exec e p1 e')
→ (exec e' p2 e'')
→ (exec e (Iseq p1 p2) e'')
| exec_Iwhile:
∀ e cond p e',
(exec e (Iif cond (Iseq p (Iwhile cond p)) Iskip) e')
→ (exec e (Iwhile cond p) e').
| exec_Iskip:
∀ e, (exec e Iskip e)
| exec_Iset:
∀ (A:Type) e x (expr: E.Expr A),
(exec e (Iset x expr) (E.upd x (E.eval expr e) e))
| exec_Iif:
∀ e (cond: E.Expr bool) p1 p2 e',
(exec e (if (E.eval cond e) then p1 else p2) e')
→ (exec e (Iif cond p1 p2) e')
| exec_Iseq:
∀ e p1 p2 e' e'',
(exec e p1 e')
→ (exec e' p2 e'')
→ (exec e (Iseq p1 p2) e'')
| exec_Iwhile:
∀ e cond p e',
(exec e (Iif cond (Iseq p (Iwhile cond p)) Iskip) e')
→ (exec e (Iwhile cond p) e').
This language is thus deterministic.
This is proved by lemma exec_deterministic in file
totalhoarelogic.
If prog is a program, pre a precondition (i.e. a required
property on the initial environment) and post a postcondition
(i.e. a property on the final environment that prog must ensure
when pre is initially satisfied), then a specification in Hoare
logic is written:
Below, we give the semantics of these notations. Here the assertion
language is shallow embedded (there is no syntax for it: we use
directly the Coq proposition language).
Pred is the type of assertions.
The specification language: syntax and semantics
- "pre |= prog {= post =}" in partial correctness (i.e. even if pre is satisfied, prog may not terminate)
- "pre |= prog [= post =]" in total correctness (i.e. when pre is satisfied, prog must terminate).
Deduction operator "|=" is here only a syntactic sugar.
In the following, p |= q can also be read "q is weaker than p".
wlp represents the Weakest Liberal Precondition (i.e.
"wlp prog post" is the weakest condition on the initial environment ensuring
that post is ensured on any final environment of prog, when such environment
exists)
Definition wlp: ImpProg → Pred → Pred
:= fun prog post e ⇒ (∀ e', (exec e prog e') → (post e')).
Notation "p {= post =}" := (wlp p post) (at level 70).
:= fun prog post e ⇒ (∀ e', (exec e prog e') → (post e')).
Notation "p {= post =}" := (wlp p post) (at level 70).
wp represents the Weakest Precondition (i.e. "wp prog post" is
the weakest condition on the initial environment ensuring that there
exists a final environment of prog satisfying post). Here as the language is
deterministic, I proved that "(wp prog post) |= (wlp prog post)"
see lemma wp_entails_wlp below.
Definition wp: ImpProg → Pred → Pred
:= fun prog post e ⇒ ∃ e', (exec e prog e') ∧ (post e').
Notation "p [= post =]" := (wp p post) (at level 70).
:= fun prog post e ⇒ ∃ e', (exec e prog e') ∧ (post e').
Notation "p [= post =]" := (wp p post) (at level 70).
Hint Resolve exec_Iskip exec_Iset exec_Iif exec_Iseq exec_Iwhile: hoare.
Parameter exec_Iif_true:
∀ e cond p1 p2 e',
(E.eval cond e)=true
→ (exec e p1 e')
→ (exec e (Iif cond p1 p2) e').
Parameter exec_Iif_false:
∀ e cond p1 p2 e',
(E.eval cond e)=false
→ (exec e p2 e')
→ (exec e (Iif cond p1 p2) e').
Hint Resolve exec_Iif_true exec_Iif_false: hoare.
End HoareLogicDefs.
Generation of proof obligations (POs)
Module Type HoareProofSystem.
Declare Module HLD: HoareLogicDefs.
Import HLD.
Parameter sem_wp: ImpProg → Pred → Pred.
Parameter synt_wp: ImpProg → Pred → Pred.
Parameter soundness: ∀ pre p post, pre |= (synt_wp p post) → pre |= (sem_wp p post).
Parameter completeness: ∀ pre p post, pre |= (sem_wp p post) → pre |= (synt_wp p post).
End HoareProofSystem.
The whole meta-theory
The whole meta-theory is described by this module signature. This signature is implemented in file hoarelogic.- module PHL contains the PO generation in partial correctness.
- module THL contains the PO generation in total correctness.
Module Type HoareLogicSem.
Declare Module E: ExprLang.
Declare Module HLD: HoareLogicDefs with Module E:=E.
Import HLD.
Declare Module PHL: HoareProofSystem with Module HLD:=HLD with Definition sem_wp:=wlp.
Declare Module THL: HoareProofSystem with Module HLD:=HLD with Definition sem_wp:=wp.
Parameter wp_entails_wlp: ∀ prog post, prog [= post =] |= prog {= post =}.
End HoareLogicSem.
"Tutorial on Hoare Logic" Library. Copyright 2007 Sylvain Boulme.
This file is distributed under the terms of the
"GNU LESSER GENERAL PUBLIC LICENSE" version 3.