src/Doc/Datatypes/Datatypes.thy
changeset 53753 ae7f50e70c09
parent 53752 1a883094fbe0
child 53798 6a4e3299dfd1
equal deleted inserted replaced
53752:1a883094fbe0 53753:ae7f50e70c09
     8 *)
     8 *)
     9 
     9 
    10 theory Datatypes
    10 theory Datatypes
    11 imports Setup
    11 imports Setup
    12 keywords
    12 keywords
    13   "primcorec_notyet" :: thy_decl
    13   "primcorec" :: thy_goal and
       
    14   "primcorecursive_notyet" :: thy_decl
    14 begin
    15 begin
    15 
    16 
    16 (*<*)
    17 (*<*)
    17 (* FIXME: Temporary setup until "primcorec" is fully implemented. *)
    18 (* FIXME: Temporary setup until "primcorec" and "primcorecursive" are fully
       
    19    implemented. *)
    18 ML_command {*
    20 ML_command {*
    19 fun add_dummy_cmd _ _ lthy = lthy;
    21 fun add_dummy_cmd _ _ lthy = lthy;
    20 
    22 
    21 val _ = Outer_Syntax.local_theory @{command_spec "primcorec_notyet"} ""
    23 val _ = Outer_Syntax.local_theory @{command_spec "primcorec"} ""
       
    24   (Parse.fixes -- Parse_Spec.where_alt_specs >> uncurry add_dummy_cmd);
       
    25 
       
    26 val _ = Outer_Syntax.local_theory @{command_spec "primcorecursive_notyet"} ""
    22   (Parse.fixes -- Parse_Spec.where_alt_specs >> uncurry add_dummy_cmd);
    27   (Parse.fixes -- Parse_Spec.where_alt_specs >> uncurry add_dummy_cmd);
    23 *}
    28 *}
    24 (*>*)
    29 (*>*)
    25 
    30 
    26 
    31 
   124 \item Section \ref{sec:defining-codatatypes}, ``Defining Codatatypes,''
   129 \item Section \ref{sec:defining-codatatypes}, ``Defining Codatatypes,''
   125 describes how to specify codatatypes using the \keyw{codatatype} command.
   130 describes how to specify codatatypes using the \keyw{codatatype} command.
   126 
   131 
   127 \item Section \ref{sec:defining-corecursive-functions}, ``Defining Corecursive
   132 \item Section \ref{sec:defining-corecursive-functions}, ``Defining Corecursive
   128 Functions,'' describes how to specify corecursive functions using the
   133 Functions,'' describes how to specify corecursive functions using the
   129 @{command primcorec} command.
   134 @{command primcorecursive} command.
   130 
   135 
   131 \item Section \ref{sec:registering-bounded-natural-functors}, ``Registering
   136 \item Section \ref{sec:registering-bounded-natural-functors}, ``Registering
   132 Bounded Natural Functors,'' explains how to use the @{command bnf} command
   137 Bounded Natural Functors,'' explains how to use the @{command bnf} command
   133 to register arbitrary type constructors as BNFs.
   138 to register arbitrary type constructors as BNFs.
   134 
   139 
  1563 
  1568 
  1564 section {* Defining Corecursive Functions
  1569 section {* Defining Corecursive Functions
  1565   \label{sec:defining-corecursive-functions} *}
  1570   \label{sec:defining-corecursive-functions} *}
  1566 
  1571 
  1567 text {*
  1572 text {*
  1568 Corecursive functions can be specified using @{command primcorec}, which
  1573 Corecursive functions can be specified using @{command primcorec} and
  1569 supports primitive corecursion, or using the more general
  1574 @{command primcorecursive}, which support primitive corecursion, or using the
  1570 \keyw{partial\_function} command. Here, the focus is on @{command primcorec}.
  1575 more general \keyw{partial\_function} command. Here, the focus is on
  1571 More examples can be found in the directory \verb|~~/src/HOL/BNF/Examples|.
  1576 the former two. More examples can be found in the directory
       
  1577 \verb|~~/src/HOL/BNF/Examples|.
  1572 
  1578 
  1573 Whereas recursive functions consume datatypes one constructor at a time,
  1579 Whereas recursive functions consume datatypes one constructor at a time,
  1574 corecursive functions construct codatatypes one constructor at a time.
  1580 corecursive functions construct codatatypes one constructor at a time.
  1575 Partly reflecting a lack of agreement among proponents of coalgebraic methods,
  1581 Partly reflecting a lack of agreement among proponents of coalgebraic methods,
  1576 Isabelle supports three competing syntaxes for specifying a function $f$:
  1582 Isabelle supports three competing syntaxes for specifying a function $f$:
  1596 with restrictions on the format of the right-hand side. Lazy functional
  1602 with restrictions on the format of the right-hand side. Lazy functional
  1597 programming languages such as Haskell support a generalized version of this
  1603 programming languages such as Haskell support a generalized version of this
  1598 style.
  1604 style.
  1599 \end{itemize}
  1605 \end{itemize}
  1600 
  1606 
  1601 All three styles are available as input syntax to @{command primcorec}.
  1607 All three styles are available as input syntax. Whichever syntax is chosen,
  1602 Whichever syntax is chosen, characteristic theorems for all three styles are
  1608 characteristic theorems for all three styles are generated.
  1603 generated.
       
  1604 
  1609 
  1605 \begin{framed}
  1610 \begin{framed}
  1606 \noindent
  1611 \noindent
  1607 \textbf{Warning:}\enskip The @{command primcorec} command is under development.
  1612 \textbf{Warning:}\enskip The @{command primcorec} and @{command primcorecursive}
  1608 Some of the functionality described here is vaporware. An alternative is to
  1613 commands are under development. Some of the functionality described here is
  1609 define corecursive functions directly using the generated @{text t_unfold} or
  1614 vaporware. An alternative is to define corecursive functions directly using the
  1610 @{text t_corec} combinators.
  1615 generated @{text t_unfold} or @{text t_corec} combinators.
  1611 \end{framed}
  1616 \end{framed}
  1612 
  1617 
  1613 %%% TODO: partial_function? E.g. for defining tail recursive function on lazy
  1618 %%% TODO: partial_function? E.g. for defining tail recursive function on lazy
  1614 %%% lists (cf. terminal0 in TLList.thy)
  1619 %%% lists (cf. terminal0 in TLList.thy)
  1615 *}
  1620 *}
  1639 Following the code view, corecursive calls are allowed on the right-hand side as
  1644 Following the code view, corecursive calls are allowed on the right-hand side as
  1640 long as they occur under a constructor, which itself appears either directly to
  1645 long as they occur under a constructor, which itself appears either directly to
  1641 the right of the equal sign or in a conditional expression:
  1646 the right of the equal sign or in a conditional expression:
  1642 *}
  1647 *}
  1643 
  1648 
  1644     primcorec literate :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a llist" where
  1649     primcorecursive literate :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a llist" where
  1645       "literate f x = LCons x (literate f (f x))"
  1650       "literate f x = LCons x (literate f (f x))"
  1646     .
  1651     .
  1647 
  1652 
  1648 text {* \blankline *}
  1653 text {* \blankline *}
  1649 
  1654 
  1650     primcorec siterate :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a stream" where
  1655     primcorecursive siterate :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a stream" where
  1651       "siterate f x = SCons x (siterate f (f x))"
  1656       "siterate f x = SCons x (siterate f (f x))"
  1652     .
  1657     .
  1653 
  1658 
  1654 text {*
  1659 text {*
  1655 \noindent
  1660 \noindent
  1664 Corecursive functions construct codatatype values, but nothing prevents them
  1669 Corecursive functions construct codatatype values, but nothing prevents them
  1665 from also consuming such values. The following function drops ever second
  1670 from also consuming such values. The following function drops ever second
  1666 element in a stream:
  1671 element in a stream:
  1667 *}
  1672 *}
  1668 
  1673 
  1669     primcorec every_snd :: "'a stream \<Rightarrow> 'a stream" where
  1674     primcorecursive every_snd :: "'a stream \<Rightarrow> 'a stream" where
  1670       "every_snd s = SCons (shd s) (stl (stl s))"
  1675       "every_snd s = SCons (shd s) (stl (stl s))"
  1671     .
  1676     .
  1672 
  1677 
  1673 text {*
  1678 text {*
  1674 \noindent
  1679 \noindent
  1675 Constructs such as @{text "let"}---@{text "in"}, @{text
  1680 Constructs such as @{text "let"}---@{text "in"}, @{text
  1676 "if"}---@{text "then"}---@{text "else"}, and @{text "case"}---@{text "of"} may
  1681 "if"}---@{text "then"}---@{text "else"}, and @{text "case"}---@{text "of"} may
  1677 appear around constructors that guard corecursive calls:
  1682 appear around constructors that guard corecursive calls:
  1678 *}
  1683 *}
  1679 
  1684 
  1680     primcorec_notyet lappend :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist" where
  1685     primcorecursive_notyet lappend :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist" where
  1681       "lappend xs ys =
  1686       "lappend xs ys =
  1682          (case xs of
  1687          (case xs of
  1683             LNil \<Rightarrow> ys
  1688             LNil \<Rightarrow> ys
  1684           | LCons x xs' \<Rightarrow> LCons x (lappend xs' ys))"
  1689           | LCons x xs' \<Rightarrow> LCons x (lappend xs' ys))"
  1685 
  1690 
  1686 text {*
  1691 text {*
  1687 \noindent
  1692 \noindent
  1688 Corecursion is useful to specify not only functions but also infinite objects:
  1693 Corecursion is useful to specify not only functions but also infinite objects:
  1689 *}
  1694 *}
  1690 
  1695 
  1691     primcorec infty :: enat where
  1696     primcorecursive infty :: enat where
  1692       "infty = ESuc infty"
  1697       "infty = ESuc infty"
  1693     .
  1698     .
  1694 
  1699 
  1695 text {*
  1700 text {*
  1696 \noindent
  1701 \noindent
  1697 The example below constructs a pseudorandom process value. It takes a stream of
  1702 The example below constructs a pseudorandom process value. It takes a stream of
  1698 actions (@{text s}), a pseudorandom function generator (@{text f}), and a
  1703 actions (@{text s}), a pseudorandom function generator (@{text f}), and a
  1699 pseudorandom seed (@{text n}):
  1704 pseudorandom seed (@{text n}):
  1700 *}
  1705 *}
  1701 
  1706 
  1702     primcorec_notyet
  1707     primcorecursive_notyet
  1703       random_process :: "'a stream \<Rightarrow> (int \<Rightarrow> int) \<Rightarrow> int \<Rightarrow> 'a process"
  1708       random_process :: "'a stream \<Rightarrow> (int \<Rightarrow> int) \<Rightarrow> int \<Rightarrow> 'a process"
  1704     where
  1709     where
  1705       "random_process s f n =
  1710       "random_process s f n =
  1706          (if n mod 4 = 0 then
  1711          (if n mod 4 = 0 then
  1707             Fail
  1712             Fail
  1727 text {*
  1732 text {*
  1728 The syntax for mutually corecursive functions over mutually corecursive
  1733 The syntax for mutually corecursive functions over mutually corecursive
  1729 datatypes is unsurprising:
  1734 datatypes is unsurprising:
  1730 *}
  1735 *}
  1731 
  1736 
  1732     primcorec
  1737     primcorecursive
  1733       even_infty :: even_enat and
  1738       even_infty :: even_enat and
  1734       odd_infty :: odd_enat
  1739       odd_infty :: odd_enat
  1735     where
  1740     where
  1736       "even_infty = Even_ESuc odd_infty" |
  1741       "even_infty = Even_ESuc odd_infty" |
  1737       "odd_infty = Odd_ESuc even_infty"
  1742       "odd_infty = Odd_ESuc even_infty"
  1746 functions (Section~\ref{sssec:primcorec-nested-corecursion}) to possibly
  1751 functions (Section~\ref{sssec:primcorec-nested-corecursion}) to possibly
  1747 infinite trees in which subnodes are organized either as a lazy list (@{text
  1752 infinite trees in which subnodes are organized either as a lazy list (@{text
  1748 tree\<^sub>i\<^sub>i}) or as a finite set (@{text tree\<^sub>i\<^sub>s}):
  1753 tree\<^sub>i\<^sub>i}) or as a finite set (@{text tree\<^sub>i\<^sub>s}):
  1749 *}
  1754 *}
  1750 
  1755 
  1751     primcorec iterate\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>i" where
  1756     primcorecursive iterate\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>i" where
  1752       "iterate\<^sub>i\<^sub>i f x = Node\<^sub>i\<^sub>i x (lmap (iterate\<^sub>i\<^sub>i f) (f x))"
  1757       "iterate\<^sub>i\<^sub>i f x = Node\<^sub>i\<^sub>i x (lmap (iterate\<^sub>i\<^sub>i f) (f x))"
  1753     .
  1758     .
  1754 
  1759 
  1755 text {* \blankline *}
  1760 text {* \blankline *}
  1756 
  1761 
  1757     primcorec iterate\<^sub>i\<^sub>s :: "('a \<Rightarrow> 'a fset) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>s" where
  1762     primcorecursive iterate\<^sub>i\<^sub>s :: "('a \<Rightarrow> 'a fset) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>s" where
  1758       "iterate\<^sub>i\<^sub>s f x = Node\<^sub>i\<^sub>s x (fmap (iterate\<^sub>i\<^sub>s f) (f x))"
  1763       "iterate\<^sub>i\<^sub>s f x = Node\<^sub>i\<^sub>s x (fmap (iterate\<^sub>i\<^sub>s f) (f x))"
  1759     .
  1764     .
  1760 
  1765 
  1761 text {*
  1766 text {*
  1762 \noindent
  1767 \noindent
  1765 @{text \<Sigma>} is a finite alphabet, @{text \<delta>} is a transition function, @{text q\<^sub>0}
  1770 @{text \<Sigma>} is a finite alphabet, @{text \<delta>} is a transition function, @{text q\<^sub>0}
  1766 is an initial state, and @{text F} is a set of final states. The following
  1771 is an initial state, and @{text F} is a set of final states. The following
  1767 function translates a DFA into a @{type state_machine}:
  1772 function translates a DFA into a @{type state_machine}:
  1768 *}
  1773 *}
  1769 
  1774 
  1770     primcorec (*<*)(in early) (*>*)
  1775     primcorecursive (*<*)(in early) (*>*)
  1771       sm_of_dfa :: "('q \<Rightarrow> 'a \<Rightarrow> 'q) \<Rightarrow> 'q set \<Rightarrow> 'q \<Rightarrow> 'a state_machine"
  1776       sm_of_dfa :: "('q \<Rightarrow> 'a \<Rightarrow> 'q) \<Rightarrow> 'q set \<Rightarrow> 'q \<Rightarrow> 'a state_machine"
  1772     where
  1777     where
  1773       "sm_of_dfa \<delta> F q = State_Machine (q \<in> F) (sm_of_dfa \<delta> F o \<delta> q)"
  1778       "sm_of_dfa \<delta> F q = State_Machine (q \<in> F) (sm_of_dfa \<delta> F o \<delta> q)"
  1774     .
  1779     .
  1775 
  1780 
  1779 (@{text "op \<circ>"}). For convenience, corecursion through functions can be
  1784 (@{text "op \<circ>"}). For convenience, corecursion through functions can be
  1780 expressed using @{text \<lambda>}-expressions and function application rather
  1785 expressed using @{text \<lambda>}-expressions and function application rather
  1781 than composition. For example:
  1786 than composition. For example:
  1782 *}
  1787 *}
  1783 
  1788 
  1784     primcorec
  1789     primcorecursive
  1785       sm_of_dfa :: "('q \<Rightarrow> 'a \<Rightarrow> 'q) \<Rightarrow> 'q set \<Rightarrow> 'q \<Rightarrow> 'a state_machine"
  1790       sm_of_dfa :: "('q \<Rightarrow> 'a \<Rightarrow> 'q) \<Rightarrow> 'q set \<Rightarrow> 'q \<Rightarrow> 'a state_machine"
  1786     where
  1791     where
  1787       "sm_of_dfa \<delta> F q = State_Machine (q \<in> F) (sm_of_dfa \<delta> F o \<delta> q)"
  1792       "sm_of_dfa \<delta> F q = State_Machine (q \<in> F) (sm_of_dfa \<delta> F o \<delta> q)"
  1788     .
  1793     .
  1789 
  1794 
  1790 text {* \blankline *}
  1795 text {* \blankline *}
  1791 
  1796 
  1792     primcorec empty_sm :: "'a state_machine" where
  1797     primcorecursive empty_sm :: "'a state_machine" where
  1793       "empty_sm = State_Machine False (\<lambda>_. empty_sm)"
  1798       "empty_sm = State_Machine False (\<lambda>_. empty_sm)"
  1794     .
  1799     .
  1795 
  1800 
  1796 text {* \blankline *}
  1801 text {* \blankline *}
  1797 
  1802 
  1798     primcorec not_sm :: "'a state_machine \<Rightarrow> 'a state_machine" where
  1803     primcorecursive not_sm :: "'a state_machine \<Rightarrow> 'a state_machine" where
  1799       "not_sm M = State_Machine (\<not> accept M) (\<lambda>a. not_sm (trans M a))"
  1804       "not_sm M = State_Machine (\<not> accept M) (\<lambda>a. not_sm (trans M a))"
  1800     .
  1805     .
  1801 
  1806 
  1802 text {* \blankline *}
  1807 text {* \blankline *}
  1803 
  1808 
  1804     primcorec
  1809     primcorecursive
  1805       or_sm :: "'a state_machine \<Rightarrow> 'a state_machine \<Rightarrow> 'a state_machine"
  1810       or_sm :: "'a state_machine \<Rightarrow> 'a state_machine \<Rightarrow> 'a state_machine"
  1806     where
  1811     where
  1807       "or_sm M N =
  1812       "or_sm M N =
  1808          State_Machine (accept M \<or> accept N)
  1813          State_Machine (accept M \<or> accept N)
  1809            (\<lambda>a. or_sm (trans M a) (trans N a))"
  1814            (\<lambda>a. or_sm (trans M a) (trans N a))"
  1818 were mutually recursive
  1823 were mutually recursive
  1819 (Section~\ref{sssec:primrec-nested-as-mutual-recursion}), it is possible to
  1824 (Section~\ref{sssec:primrec-nested-as-mutual-recursion}), it is possible to
  1820 pretend that nested codatatypes are mutually corecursive. For example:
  1825 pretend that nested codatatypes are mutually corecursive. For example:
  1821 *}
  1826 *}
  1822 
  1827 
  1823     primcorec_notyet
  1828     primcorecursive_notyet
  1824       iterate\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>i" and
  1829       iterate\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>i" and
  1825       iterates\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a llist \<Rightarrow> 'a tree\<^sub>i\<^sub>i llist"
  1830       iterates\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a llist \<Rightarrow> 'a tree\<^sub>i\<^sub>i llist"
  1826     where
  1831     where
  1827       "iterate\<^sub>i\<^sub>i f x = Node\<^sub>i\<^sub>i x (iterates\<^sub>i\<^sub>i f (f x))" |
  1832       "iterate\<^sub>i\<^sub>i f x = Node\<^sub>i\<^sub>i x (iterates\<^sub>i\<^sub>i f (f x))" |
  1828       "iterates\<^sub>i\<^sub>i f xs =
  1833       "iterates\<^sub>i\<^sub>i f xs =
  1850 and @{const siterate}, are identical in both styles.
  1855 and @{const siterate}, are identical in both styles.
  1851 
  1856 
  1852 Here is an example where there is a difference:
  1857 Here is an example where there is a difference:
  1853 *}
  1858 *}
  1854 
  1859 
  1855     primcorec lappend :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist" where
  1860     primcorecursive lappend :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist" where
  1856       "lnull xs \<Longrightarrow> lnull ys \<Longrightarrow> lappend xs ys = LNil" |
  1861       "lnull xs \<Longrightarrow> lnull ys \<Longrightarrow> lappend xs ys = LNil" |
  1857       "_ \<Longrightarrow> lappend xs ys = LCons (lhd (if lnull xs then ys else xs))
  1862       "_ \<Longrightarrow> lappend xs ys = LCons (lhd (if lnull xs then ys else xs))
  1858          (if xs = LNil then ltl ys else lappend (ltl xs) ys)"
  1863          (if xs = LNil then ltl ys else lappend (ltl xs) ys)"
  1859     .
  1864     .
  1860 
  1865 
  1874 
  1879 
  1875 In contrast, the next example is arguably more naturally expressed in the
  1880 In contrast, the next example is arguably more naturally expressed in the
  1876 constructor view:
  1881 constructor view:
  1877 *}
  1882 *}
  1878 
  1883 
  1879     primcorec_notyet
  1884     primcorecursive_notyet
  1880       random_process :: "'a stream \<Rightarrow> (int \<Rightarrow> int) \<Rightarrow> int \<Rightarrow> 'a process"
  1885       random_process :: "'a stream \<Rightarrow> (int \<Rightarrow> int) \<Rightarrow> int \<Rightarrow> 'a process"
  1881     where
  1886     where
  1882       "n mod 4 = 0 \<Longrightarrow> random_process s f n = Fail" |
  1887       "n mod 4 = 0 \<Longrightarrow> random_process s f n = Fail" |
  1883       "n mod 4 = 1 \<Longrightarrow>
  1888       "n mod 4 = 1 \<Longrightarrow>
  1884          random_process s f n = Skip (random_process s f (f n))" |
  1889          random_process s f n = Skip (random_process s f (f n))" |
  1920 determine which constructor to choose, and these conditions are interpreted
  1925 determine which constructor to choose, and these conditions are interpreted
  1921 sequentially or not depending on the @{text "sequential"} option.
  1926 sequentially or not depending on the @{text "sequential"} option.
  1922 Consider the following examples:
  1927 Consider the following examples:
  1923 *}
  1928 *}
  1924 
  1929 
  1925     primcorec literate :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a llist" where
  1930     primcorecursive literate :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a llist" where
  1926       "\<not> lnull (literate _ x)" |
  1931       "\<not> lnull (literate _ x)" |
  1927       "lhd (literate _ x) = x" |
  1932       "lhd (literate _ x) = x" |
  1928       "ltl (literate f x) = literate f (f x)"
  1933       "ltl (literate f x) = literate f (f x)"
  1929     .
  1934     .
  1930 
  1935 
  1931 text {* \blankline *}
  1936 text {* \blankline *}
  1932 
  1937 
  1933     primcorec siterate :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a stream" where
  1938     primcorecursive siterate :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a stream" where
  1934       "shd (siterate _ x) = x" |
  1939       "shd (siterate _ x) = x" |
  1935       "stl (siterate f x) = siterate f (f x)"
  1940       "stl (siterate f x) = siterate f (f x)"
  1936     .
  1941     .
  1937 
  1942 
  1938 text {* \blankline *}
  1943 text {* \blankline *}
  1939 
  1944 
  1940     primcorec every_snd :: "'a stream \<Rightarrow> 'a stream" where
  1945     primcorecursive every_snd :: "'a stream \<Rightarrow> 'a stream" where
  1941       "shd (every_snd s) = shd s" |
  1946       "shd (every_snd s) = shd s" |
  1942       "stl (every_snd s) = stl (stl s)"
  1947       "stl (every_snd s) = stl (stl s)"
  1943     .
  1948     .
  1944 
  1949 
  1945 text {*
  1950 text {*
  1953 
  1958 
  1954 The next example shows how to specify functions that rely on more than one
  1959 The next example shows how to specify functions that rely on more than one
  1955 constructor:
  1960 constructor:
  1956 *}
  1961 *}
  1957 
  1962 
  1958     primcorec lappend :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist" where
  1963     primcorecursive lappend :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist" where
  1959       "lnull xs \<Longrightarrow> lnull ys \<Longrightarrow> lnull (lappend xs ys)" |
  1964       "lnull xs \<Longrightarrow> lnull ys \<Longrightarrow> lnull (lappend xs ys)" |
  1960       "lhd (lappend xs ys) = lhd (if lnull xs then ys else xs)" |
  1965       "lhd (lappend xs ys) = lhd (if lnull xs then ys else xs)" |
  1961       "ltl (lappend xs ys) = (if xs = LNil then ltl ys else lappend (ltl xs) ys)"
  1966       "ltl (lappend xs ys) = (if xs = LNil then ltl ys else lappend (ltl xs) ys)"
  1962     .
  1967     .
  1963 
  1968 
  1969 *}
  1974 *}
  1970 
  1975 
  1971 (*<*)
  1976 (*<*)
  1972     end
  1977     end
  1973 
  1978 
  1974     primcorec lappend :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist" where
  1979     primcorecursive lappend :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist" where
  1975       "lnull xs \<Longrightarrow> lnull ys \<Longrightarrow> lnull (lappend xs ys)" |
  1980       "lnull xs \<Longrightarrow> lnull ys \<Longrightarrow> lnull (lappend xs ys)" |
  1976 (*>*)
  1981 (*>*)
  1977       "_ \<Longrightarrow> \<not> lnull (lappend xs ys)"
  1982       "_ \<Longrightarrow> \<not> lnull (lappend xs ys)"
  1978 (*<*) |
  1983 (*<*) |
  1979       "lhd (lappend xs ys) = lhd (if lnull xs then ys else xs)" |
  1984       "lhd (lappend xs ys) = lhd (if lnull xs then ys else xs)" |
  1989 
  1994 
  1990 The next example illustrates how to cope with selectors defined for several
  1995 The next example illustrates how to cope with selectors defined for several
  1991 constructors:
  1996 constructors:
  1992 *}
  1997 *}
  1993 
  1998 
  1994     primcorec_notyet
  1999     primcorecursive_notyet
  1995       random_process :: "'a stream \<Rightarrow> (int \<Rightarrow> int) \<Rightarrow> int \<Rightarrow> 'a process"
  2000       random_process :: "'a stream \<Rightarrow> (int \<Rightarrow> int) \<Rightarrow> int \<Rightarrow> 'a process"
  1996     where
  2001     where
  1997       "n mod 4 = 0 \<Longrightarrow> is_Fail (random_process s f n)" |
  2002       "n mod 4 = 0 \<Longrightarrow> is_Fail (random_process s f n)" |
  1998       "n mod 4 = 1 \<Longrightarrow> is_Skip (random_process s f n)" |
  2003       "n mod 4 = 1 \<Longrightarrow> is_Skip (random_process s f n)" |
  1999       "n mod 4 = 2 \<Longrightarrow> is_Action (random_process s f n)" |
  2004       "n mod 4 = 2 \<Longrightarrow> is_Action (random_process s f n)" |
  2012 cont} depending on which constructor is selected.
  2017 cont} depending on which constructor is selected.
  2013 
  2018 
  2014 Here are more examples to conclude:
  2019 Here are more examples to conclude:
  2015 *}
  2020 *}
  2016 
  2021 
  2017     primcorec
  2022     primcorecursive
  2018       even_infty :: even_enat and
  2023       even_infty :: even_enat and
  2019       odd_infty :: odd_enat
  2024       odd_infty :: odd_enat
  2020     where
  2025     where
  2021       "\<not> is_Even_EZero even_infty" |
  2026       "\<not> is_Even_EZero even_infty" |
  2022       "un_Even_ESuc even_infty = odd_infty" |
  2027       "un_Even_ESuc even_infty = odd_infty" |
  2023       "un_Odd_ESuc odd_infty = even_infty"
  2028       "un_Odd_ESuc odd_infty = even_infty"
  2024     .
  2029     .
  2025 
  2030 
  2026 text {* \blankline *}
  2031 text {* \blankline *}
  2027 
  2032 
  2028     primcorec iterate\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>i" where
  2033     primcorecursive iterate\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>i" where
  2029       "lbl\<^sub>i\<^sub>i (iterate\<^sub>i\<^sub>i f x) = x" |
  2034       "lbl\<^sub>i\<^sub>i (iterate\<^sub>i\<^sub>i f x) = x" |
  2030       "sub\<^sub>i\<^sub>i (iterate\<^sub>i\<^sub>i f x) = lmap (iterate\<^sub>i\<^sub>i f) (f x)"
  2035       "sub\<^sub>i\<^sub>i (iterate\<^sub>i\<^sub>i f x) = lmap (iterate\<^sub>i\<^sub>i f) (f x)"
  2031     .
  2036     .
  2032 
  2037 
  2033 (*<*)
  2038 (*<*)
  2037 
  2042 
  2038 subsection {* Command Syntax
  2043 subsection {* Command Syntax
  2039   \label{ssec:primcorec-command-syntax} *}
  2044   \label{ssec:primcorec-command-syntax} *}
  2040 
  2045 
  2041 
  2046 
  2042 subsubsection {* \keyw{primcorec}
  2047 subsubsection {* \keyw{primcorecursive} and \keyw{primcorec}
  2043   \label{sssec:primcorec} *}
  2048   \label{sssec:primcorecursive-and-primcorec} *}
  2044 
  2049 
  2045 text {*
  2050 text {*
  2046 Primitive corecursive definitions have the following general syntax:
  2051 Primitive corecursive definitions have the following general syntax:
  2047 
  2052 
  2048 @{rail "
  2053 @{rail "
  2049   @@{command_def primcorec} target? @{syntax pcr_options}? fixes \\ @'where'
  2054   @@{command_def primcorecursive} target? @{syntax pcr_option}? fixes \\ @'where'
  2050     (@{syntax pcr_formula} + '|')
  2055     (@{syntax pcr_formula} + '|')
  2051   ;
  2056   ;
  2052   @{syntax_def pcr_options}: '(' 'sequential' ')'
  2057   @@{command_def primcorec} target? fixes \\ @'where'
       
  2058     (@{syntax pcr_formula} + '|')
       
  2059   ;
       
  2060   @{syntax_def pcr_option}: '(' 'sequential' ')'
  2053   ;
  2061   ;
  2054   @{syntax_def pcr_formula}: thmdecl? prop (@'of' (term * ))?
  2062   @{syntax_def pcr_formula}: thmdecl? prop (@'of' (term * ))?
  2055 "}
  2063 "}
  2056 
  2064 
  2057 The optional target is optionally followed by a corecursion-specific option:
  2065 The optional target is optionally followed by a corecursion-specific option:
  2063 The @{text "sequential"} option indicates that the conditions in specifications
  2071 The @{text "sequential"} option indicates that the conditions in specifications
  2064 expressed using the constructor or destructor view are to be interpreted
  2072 expressed using the constructor or destructor view are to be interpreted
  2065 sequentially.
  2073 sequentially.
  2066 \end{itemize}
  2074 \end{itemize}
  2067 
  2075 
       
  2076 The @{command primcorec} command is an abbreviation for
       
  2077 @{command primcorecursive} with @{text "sequential"} enabled. It has no proof
       
  2078 obligations.
  2068 *}
  2079 *}
  2069 
  2080 
  2070 
  2081 
  2071 (*
  2082 (*
  2072 subsection {* Generated Theorems
  2083 subsection {* Generated Theorems
  2237 text {*
  2248 text {*
  2238 Known open issues of the package.
  2249 Known open issues of the package.
  2239 *}
  2250 *}
  2240 
  2251 
  2241 text {*
  2252 text {*
  2242 %* primcorec is unfinished
  2253 %* primcorecursive and primcorec is unfinished
  2243 %
  2254 %
  2244 %* slow n-ary mutual (co)datatype, avoid as much as possible (e.g. using nesting)
  2255 %* slow n-ary mutual (co)datatype, avoid as much as possible (e.g. using nesting)
  2245 %
  2256 %
  2246 %* issues with HOL-Proofs?
  2257 %* issues with HOL-Proofs?
  2247 %
  2258 %