\[\begin{split}\newcommand{\as}{\kw{as}}
\newcommand{\Assum}[3]{\kw{Assum}(#1)(#2:#3)}
\newcommand{\case}{\kw{case}}
\newcommand{\cons}{\textsf{cons}}
\newcommand{\consf}{\textsf{consf}}
\newcommand{\Def}[4]{\kw{Def}(#1)(#2:=#3:#4)}
\newcommand{\emptyf}{\textsf{emptyf}}
\newcommand{\End}{\kw{End}}
\newcommand{\kwend}{\kw{end}}
\newcommand{\even}{\textsf{even}}
\newcommand{\evenO}{\textsf{even}_\textsf{O}}
\newcommand{\evenS}{\textsf{even}_\textsf{S}}
\newcommand{\Fix}{\kw{Fix}}
\newcommand{\fix}{\kw{fix}}
\newcommand{\for}{\textsf{for}}
\newcommand{\forest}{\textsf{forest}}
\newcommand{\Functor}{\kw{Functor}}
\newcommand{\In}{\kw{in}}
\newcommand{\Ind}[4]{\kw{Ind}[#2](#3:=#4)}
\newcommand{\ind}[3]{\kw{Ind}~[#1]\left(#2\mathrm{~:=~}#3\right)}
\newcommand{\Indp}[5]{\kw{Ind}_{#5}(#1)[#2](#3:=#4)}
\newcommand{\Indpstr}[6]{\kw{Ind}_{#5}(#1)[#2](#3:=#4)/{#6}}
\newcommand{\injective}{\kw{injective}}
\newcommand{\kw}[1]{\textsf{#1}}
\newcommand{\length}{\textsf{length}}
\newcommand{\letin}[3]{\kw{let}~#1:=#2~\kw{in}~#3}
\newcommand{\List}{\textsf{list}}
\newcommand{\lra}{\longrightarrow}
\newcommand{\Match}{\kw{match}}
\newcommand{\Mod}[3]{{\kw{Mod}}({#1}:{#2}\,\zeroone{:={#3}})}
\newcommand{\ModImp}[3]{{\kw{Mod}}({#1}:{#2}:={#3})}
\newcommand{\ModA}[2]{{\kw{ModA}}({#1}=={#2})}
\newcommand{\ModS}[2]{{\kw{Mod}}({#1}:{#2})}
\newcommand{\ModType}[2]{{\kw{ModType}}({#1}:={#2})}
\newcommand{\mto}{.\;}
\newcommand{\nat}{\textsf{nat}}
\newcommand{\Nil}{\textsf{nil}}
\newcommand{\nilhl}{\textsf{nil\_hl}}
\newcommand{\nO}{\textsf{O}}
\newcommand{\node}{\textsf{node}}
\newcommand{\nS}{\textsf{S}}
\newcommand{\odd}{\textsf{odd}}
\newcommand{\oddS}{\textsf{odd}_\textsf{S}}
\newcommand{\ovl}[1]{\overline{#1}}
\newcommand{\Pair}{\textsf{pair}}
\newcommand{\plus}{\mathsf{plus}}
\newcommand{\SProp}{\textsf{SProp}}
\newcommand{\Prop}{\textsf{Prop}}
\newcommand{\return}{\kw{return}}
\newcommand{\Set}{\textsf{Set}}
\newcommand{\Sort}{\mathcal{S}}
\newcommand{\Str}{\textsf{Stream}}
\newcommand{\Struct}{\kw{Struct}}
\newcommand{\subst}[3]{#1\{#2/#3\}}
\newcommand{\tl}{\textsf{tl}}
\newcommand{\tree}{\textsf{tree}}
\newcommand{\trii}{\triangleright_\iota}
\newcommand{\Type}{\textsf{Type}}
\newcommand{\WEV}[3]{\mbox{$#1[] \vdash #2 \lra #3$}}
\newcommand{\WEVT}[3]{\mbox{$#1[] \vdash #2 \lra$}\\ \mbox{$ #3$}}
\newcommand{\WF}[2]{{\mathcal{W\!F}}(#1)[#2]}
\newcommand{\WFE}[1]{\WF{E}{#1}}
\newcommand{\WFT}[2]{#1[] \vdash {\mathcal{W\!F}}(#2)}
\newcommand{\WFTWOLINES}[2]{{\mathcal{W\!F}}\begin{array}{l}(#1)\\\mbox{}[{#2}]\end{array}}
\newcommand{\with}{\kw{with}}
\newcommand{\WS}[3]{#1[] \vdash #2 <: #3}
\newcommand{\WSE}[2]{\WS{E}{#1}{#2}}
\newcommand{\WT}[4]{#1[#2] \vdash #3 : #4}
\newcommand{\WTE}[3]{\WT{E}{#1}{#2}{#3}}
\newcommand{\WTEG}[2]{\WTE{\Gamma}{#1}{#2}}
\newcommand{\WTM}[3]{\WT{#1}{}{#2}{#3}}
\newcommand{\zeroone}[1]{[{#1}]}
\end{split}\]
Program derivation
Coq comes with an extension called Derive
, which supports program
derivation. Typically in the style of Bird and Meertens or derivations
of program refinements. To use the Derive extension it must first be
required with Require Coq.derive.Derive
. When the extension is loaded,
it provides the following command:
-
Command
Derive ident1 SuchThat one_term As ident2
ident1
can appear in one_term
. This command opens a new proof
presenting the user with a goal for one_term
in which the name ident1
is
bound to an existential variable ?x
(formally, there are other goals
standing for the existential variables but they are shelved, as
described in shelve
).
When the proof ends two constants are defined:
The first one is named ident1
and is defined as the proof of the
shelved goal (which is also the value of ?x
). It is always
transparent.
The second one is named ident2
. It has type type
, and its body is
the proof of the initially visible goal. It is opaque if the proof
ends with Qed
, and transparent if the proof ends with Defined
.
Example
- Require Coq.derive.Derive.
- [Loading ML file derive_plugin.cmxs ... done]
- Require Import Coq.Numbers.Natural.Peano.NPeano.
-
Section P.
-
Variables (n m k:nat).
- n is declared
m is declared
k is declared
-
Derive p SuchThat ((k*n)+(k*m) = p) As h.
- 1 focused goal (shelved: 1)
n, m, k : nat
p := ?Goal : nat
============================
k * n + k * m = p
- Proof.
- rewrite <- Nat.mul_add_distr_l.
- 1 focused goal (shelved: 1)
n, m, k : nat
p := ?Goal : nat
============================
k * (n + m) = p
- subst p.
- 1 focused goal (shelved: 1)
n, m, k : nat
============================
k * (n + m) = ?Goal
- reflexivity.
- No more goals.
- Qed.
-
End P.
-
Print p.
- p = fun n m k : nat => k * (n + m)
: nat -> nat -> nat -> nat
Arguments p (n m k)%nat_scope
- Check h.
- h
: forall n m k : nat, k * n + k * m = p n m k
Any property can be used as term
, not only an equation. In particular,
it could be an order relation specifying some form of program
refinement or a non-executable property from which deriving a program
is convenient.