mirror of
https://github.com/achlipala/frap.git
synced 2024-12-01 00:26:18 +00:00
CompilerCorrectness: a running example program to optimize
This commit is contained in:
parent
7cebd4bfba
commit
dd7ce9f869
2 changed files with 69 additions and 0 deletions
|
@ -124,6 +124,66 @@ Definition traceEquivalence (vc1 vc2 : valuation * cmd) :=
|
||||||
vc1 <| vc2 /\ vc2 <| vc1.
|
vc1 <| vc2 /\ vc2 <| vc1.
|
||||||
Infix "=|" := traceEquivalence (at level 70).
|
Infix "=|" := traceEquivalence (at level 70).
|
||||||
|
|
||||||
|
Definition daysPerWeek := 7.
|
||||||
|
Definition weeksPerMonth := 4.
|
||||||
|
Definition daysPerMonth := (daysPerWeek * weeksPerMonth)%arith.
|
||||||
|
|
||||||
|
Example month_boundaries_in_days :=
|
||||||
|
"acc" <- 0;;
|
||||||
|
while 1 loop
|
||||||
|
when daysPerMonth then
|
||||||
|
"acc" <- "acc" + daysPerMonth;;
|
||||||
|
Output "acc"
|
||||||
|
else
|
||||||
|
(* Oh no! We must have calculated it wrong, since we got zero! *)
|
||||||
|
Skip
|
||||||
|
done
|
||||||
|
done.
|
||||||
|
|
||||||
|
Hint Extern 1 (interp _ _ = _) => simplify; congruence.
|
||||||
|
Hint Extern 1 (interp _ _ <> _) => simplify; congruence.
|
||||||
|
|
||||||
|
Theorem first_few_values :
|
||||||
|
generate ($0, month_boundaries_in_days) [28; 56].
|
||||||
|
Proof.
|
||||||
|
unfold month_boundaries_in_days.
|
||||||
|
eapply GenSilent.
|
||||||
|
eapply CStep with (C := CSeq Hole _); eauto.
|
||||||
|
eapply GenSilent.
|
||||||
|
eapply CStep with (C := Hole); eauto.
|
||||||
|
eapply GenSilent.
|
||||||
|
eapply CStep with (C := Hole); eauto.
|
||||||
|
eapply GenSilent.
|
||||||
|
eapply CStep with (C := CSeq Hole _); eauto.
|
||||||
|
eapply GenSilent.
|
||||||
|
eapply CStep with (C := CSeq (CSeq Hole _) _); eauto.
|
||||||
|
eapply GenSilent.
|
||||||
|
eapply CStep with (C := CSeq Hole _); eauto.
|
||||||
|
eapply GenOutput.
|
||||||
|
eapply CStep with (C := CSeq Hole _); eauto.
|
||||||
|
replace 28 with (interp "acc"
|
||||||
|
($0 $+ ("acc", interp 0 $0)
|
||||||
|
$+ ("acc", interp ("acc" + daysPerMonth)%arith ($0 $+ ("acc", interp 0 $0))))); eauto.
|
||||||
|
eapply GenSilent.
|
||||||
|
eapply CStep with (C := Hole); eauto.
|
||||||
|
eapply GenSilent.
|
||||||
|
eapply CStep with (C := Hole); eauto.
|
||||||
|
eapply GenSilent.
|
||||||
|
eapply CStep with (C := CSeq Hole _); eauto.
|
||||||
|
eapply GenSilent.
|
||||||
|
eapply CStep with (C := CSeq (CSeq Hole _) _); eauto.
|
||||||
|
eapply GenSilent.
|
||||||
|
eapply CStep with (C := CSeq Hole _); eauto.
|
||||||
|
eapply GenOutput.
|
||||||
|
eapply CStep with (C := CSeq Hole _); eauto.
|
||||||
|
replace 56 with (interp "acc"
|
||||||
|
($0 $+ ("acc", interp 0 $0) $+ ("acc",
|
||||||
|
interp ("acc" + daysPerMonth)%arith ($0 $+ ("acc", interp 0 $0))) $+ ("acc",
|
||||||
|
interp ("acc" + daysPerMonth)%arith
|
||||||
|
($0 $+ ("acc", interp 0 $0) $+ ("acc",
|
||||||
|
interp ("acc" + daysPerMonth)%arith ($0 $+ ("acc", interp 0 $0))))))); eauto.
|
||||||
|
constructor.
|
||||||
|
Qed.
|
||||||
|
|
||||||
(** * Basic Simulation Arguments and Optimizing Expressions *)
|
(** * Basic Simulation Arguments and Optimizing Expressions *)
|
||||||
|
|
||||||
|
@ -174,6 +234,8 @@ Fixpoint cfoldExprs (c : cmd) : cmd :=
|
||||||
| Output e => Output (cfoldArith e)
|
| Output e => Output (cfoldArith e)
|
||||||
end.
|
end.
|
||||||
|
|
||||||
|
Compute cfoldExprs month_boundaries_in_days.
|
||||||
|
|
||||||
Theorem skip_or_step : forall v c,
|
Theorem skip_or_step : forall v c,
|
||||||
c = Skip
|
c = Skip
|
||||||
\/ exists v' l c', cstep (v, c) l (v', c').
|
\/ exists v' l c', cstep (v, c) l (v', c').
|
||||||
|
@ -376,6 +438,8 @@ Fixpoint cfold (c : cmd) : cmd :=
|
||||||
| Output e => Output (cfoldArith e)
|
| Output e => Output (cfoldArith e)
|
||||||
end.
|
end.
|
||||||
|
|
||||||
|
Compute cfold month_boundaries_in_days.
|
||||||
|
|
||||||
Notation silent_cstep := (fun a b => cstep a None b).
|
Notation silent_cstep := (fun a b => cstep a None b).
|
||||||
|
|
||||||
Lemma silent_generate_fwd : forall ns vc vc',
|
Lemma silent_generate_fwd : forall ns vc vc',
|
||||||
|
@ -722,6 +786,8 @@ Fixpoint noUnderscore (c : cmd) : bool :=
|
||||||
| Output e => noUnderscoreArith e
|
| Output e => noUnderscoreArith e
|
||||||
end.
|
end.
|
||||||
|
|
||||||
|
Compute noUnderscore month_boundaries_in_days.
|
||||||
|
|
||||||
Fixpoint flattenArith (tempCount : nat) (dst : var) (e : arith) : nat * cmd :=
|
Fixpoint flattenArith (tempCount : nat) (dst : var) (e : arith) : nat * cmd :=
|
||||||
match e with
|
match e with
|
||||||
| Const _
|
| Const _
|
||||||
|
@ -756,6 +822,8 @@ Fixpoint flatten (c : cmd) : cmd :=
|
||||||
| Output _ => c
|
| Output _ => c
|
||||||
end.
|
end.
|
||||||
|
|
||||||
|
Compute flatten month_boundaries_in_days.
|
||||||
|
|
||||||
Section simulation_multiple.
|
Section simulation_multiple.
|
||||||
Variable R : valuation * cmd -> valuation * cmd -> Prop.
|
Variable R : valuation * cmd -> valuation * cmd -> Prop.
|
||||||
|
|
||||||
|
|
|
@ -29,6 +29,7 @@ OperationalSemantics.v
|
||||||
LogicProgramming.v
|
LogicProgramming.v
|
||||||
LogicProgramming_template.v
|
LogicProgramming_template.v
|
||||||
AbstractInterpretation.v
|
AbstractInterpretation.v
|
||||||
|
CompilerCorrectness.v
|
||||||
LambdaCalculusAndTypeSoundness_template.v
|
LambdaCalculusAndTypeSoundness_template.v
|
||||||
LambdaCalculusAndTypeSoundness.v
|
LambdaCalculusAndTypeSoundness.v
|
||||||
TypesAndMutation.v
|
TypesAndMutation.v
|
||||||
|
|
Loading…
Reference in a new issue