doc-src/TutorialI/Datatype/Nested.thy
changeset 9792 bbefb6ce5cb2
parent 9689 751fde5307e4
child 9834 109b11c4e77e
equal deleted inserted replaced
9791:a39e5d43de55 9792:bbefb6ce5cb2
     1 (*<*)
     1 (*<*)
     2 theory Nested = Main:;
     2 theory Nested = ABexpr:
     3 (*>*)
     3 (*>*)
     4 
     4 
     5 text{*
     5 text{*
     6 So far, all datatypes had the property that on the right-hand side of their
     6 So far, all datatypes had the property that on the right-hand side of their
     7 definition they occurred only at the top-level, i.e.\ directly below a
     7 definition they occurred only at the top-level, i.e.\ directly below a
    10 *}
    10 *}
    11 
    11 
    12 datatype ('a,'b)"term" = Var 'a | App 'b "('a,'b)term list";
    12 datatype ('a,'b)"term" = Var 'a | App 'b "('a,'b)term list";
    13 
    13 
    14 text{*\noindent
    14 text{*\noindent
    15 Note that we need to quote \isa{term} on the left to avoid confusion with
    15 Note that we need to quote @{text"term"} on the left to avoid confusion with
    16 the command \isacommand{term}.
    16 the command \isacommand{term}.
    17 Parameter \isa{'a} is the type of variables and \isa{'b} the type of
    17 Parameter @{typ"'a"} is the type of variables and @{typ"'b"} the type of
    18 function symbols.
    18 function symbols.
    19 A mathematical term like $f(x,g(y))$ becomes @{term"App f [Var x, App g
    19 A mathematical term like $f(x,g(y))$ becomes @{term"App f [Var x, App g
    20   [Var y]]"}, where \isa{f}, \isa{g}, \isa{x}, \isa{y} are
    20   [Var y]]"}, where @{term"f"}, @{term"g"}, @{term"x"}, @{term"y"} are
    21 suitable values, e.g.\ numbers or strings.
    21 suitable values, e.g.\ numbers or strings.
    22 
    22 
    23 What complicates the definition of \isa{term} is the nested occurrence of
    23 What complicates the definition of @{text"term"} is the nested occurrence of
    24 \isa{term} inside \isa{list} on the right-hand side. In principle,
    24 @{text"term"} inside @{text"list"} on the right-hand side. In principle,
    25 nested recursion can be eliminated in favour of mutual recursion by unfolding
    25 nested recursion can be eliminated in favour of mutual recursion by unfolding
    26 the offending datatypes, here \isa{list}. The result for \isa{term}
    26 the offending datatypes, here @{text"list"}. The result for @{text"term"}
    27 would be something like
    27 would be something like
    28 \medskip
    28 \medskip
    29 
    29 
    30 \input{Datatype/document/unfoldnested.tex}
    30 \input{Datatype/document/unfoldnested.tex}
    31 \medskip
    31 \medskip
    32 
    32 
    33 \noindent
    33 \noindent
    34 Although we do not recommend this unfolding to the user, it shows how to
    34 Although we do not recommend this unfolding to the user, it shows how to
    35 simulate nested recursion by mutual recursion.
    35 simulate nested recursion by mutual recursion.
    36 Now we return to the initial definition of \isa{term} using
    36 Now we return to the initial definition of @{text"term"} using
    37 nested recursion.
    37 nested recursion.
    38 
    38 
    39 Let us define a substitution function on terms. Because terms involve term
    39 Let us define a substitution function on terms. Because terms involve term
    40 lists, we need to define two substitution functions simultaneously:
    40 lists, we need to define two substitution functions simultaneously:
    41 *}
    41 *}
    51 
    51 
    52   "substs s [] = []"
    52   "substs s [] = []"
    53   "substs s (t # ts) = subst s t # substs s ts";
    53   "substs s (t # ts) = subst s t # substs s ts";
    54 
    54 
    55 text{*\noindent
    55 text{*\noindent
    56 You should ignore the label \isa{subst\_App} for the moment.
    56 You should ignore the label @{thm[source]subst_App} for the moment.
    57 
    57 
    58 Similarly, when proving a statement about terms inductively, we need
    58 Similarly, when proving a statement about terms inductively, we need
    59 to prove a related statement about term lists simultaneously. For example,
    59 to prove a related statement about term lists simultaneously. For example,
    60 the fact that the identity substitution does not change a term needs to be
    60 the fact that the identity substitution does not change a term needs to be
    61 strengthened and proved as follows:
    61 strengthened and proved as follows:
    64 lemma "subst  Var t  = (t ::('a,'b)term)  \\<and>
    64 lemma "subst  Var t  = (t ::('a,'b)term)  \\<and>
    65         substs Var ts = (ts::('a,'b)term list)";
    65         substs Var ts = (ts::('a,'b)term list)";
    66 by(induct_tac t and ts, simp_all);
    66 by(induct_tac t and ts, simp_all);
    67 
    67 
    68 text{*\noindent
    68 text{*\noindent
    69 Note that \isa{Var} is the identity substitution because by definition it
    69 Note that @{term"Var"} is the identity substitution because by definition it
    70 leaves variables unchanged: @{term"subst Var (Var x) = Var x"}. Note also
    70 leaves variables unchanged: @{prop"subst Var (Var x) = Var x"}. Note also
    71 that the type annotations are necessary because otherwise there is nothing in
    71 that the type annotations are necessary because otherwise there is nothing in
    72 the goal to enforce that both halves of the goal talk about the same type
    72 the goal to enforce that both halves of the goal talk about the same type
    73 parameters \isa{('a,'b)}. As a result, induction would fail
    73 parameters @{text"('a,'b)"}. As a result, induction would fail
    74 because the two halves of the goal would be unrelated.
    74 because the two halves of the goal would be unrelated.
    75 
    75 
    76 \begin{exercise}
    76 \begin{exercise}
    77 The fact that substitution distributes over composition can be expressed
    77 The fact that substitution distributes over composition can be expressed
    78 roughly as follows:
    78 roughly as follows:
    79 \begin{ttbox}
    79 \begin{quote}
    80 subst (f o g) t = subst f (subst g t)
    80 @{text[display]"subst (f o g) t = subst f (subst g t)"}
    81 \end{ttbox}
    81 \end{quote}
    82 Correct this statement (you will find that it does not type-check),
    82 Correct this statement (you will find that it does not type-check),
    83 strengthen it, and prove it. (Note: \ttindexbold{o} is function composition;
    83 strengthen it, and prove it. (Note: \isaindexbold{o} is function composition;
    84 its definition is found in theorem \isa{o_def}).
    84 its definition is found in theorem @{thm[source]o_def}).
    85 \end{exercise}
    85 \end{exercise}
    86 \begin{exercise}\label{ex:trev-trev}
    86 \begin{exercise}\label{ex:trev-trev}
    87   Define a function \isa{trev} of type @{typ"('a,'b)term => ('a,'b)term"} that
    87   Define a function @{text"trev"} of type @{typ"('a,'b)term => ('a,'b)term"}
    88   recursively reverses the order of arguments of all function symbols in a
    88 that recursively reverses the order of arguments of all function symbols in a
    89   term. Prove that \isa{trev(trev t) = t}.
    89   term. Prove that @{prop"trev(trev t) = t"}.
    90 \end{exercise}
    90 \end{exercise}
    91 
    91 
    92 The experienced functional programmer may feel that our above definition of
    92 The experienced functional programmer may feel that our above definition of
    93 \isa{subst} is unnecessarily complicated in that \isa{substs} is completely
    93 @{term"subst"} is unnecessarily complicated in that @{term"substs"} is
    94 unnecessary. The @{term"App"}-case can be defined directly as
    94 completely unnecessary. The @{term"App"}-case can be defined directly as
    95 \begin{quote}
    95 \begin{quote}
    96 @{term[display]"subst s (App f ts) = App f (map (subst s) ts)"}
    96 @{term[display]"subst s (App f ts) = App f (map (subst s) ts)"}
    97 \end{quote}
    97 \end{quote}
    98 where @{term"map"} is the standard list function such that
    98 where @{term"map"} is the standard list function such that
    99 \isa{map f [x1,...,xn] = [f x1,...,f xn]}. This is true, but Isabelle insists
    99 @{text"map f [x1,...,xn] = [f x1,...,f xn]"}. This is true, but Isabelle
   100 on the above fixed format. Fortunately, we can easily \emph{prove} that the
   100 insists on the above fixed format. Fortunately, we can easily \emph{prove}
   101 suggested equation holds:
   101 that the suggested equation holds:
   102 *}
   102 *}
   103 
   103 
   104 lemma [simp]: "subst s (App f ts) = App f (map (subst s) ts)"
   104 lemma [simp]: "subst s (App f ts) = App f (map (subst s) ts)"
   105 by(induct_tac ts, simp_all)
   105 by(induct_tac ts, simp_all)
   106 
   106 
   113 
   113 
   114 text{*\noindent
   114 text{*\noindent
   115 The advantage is that now we have replaced @{term"substs"} by
   115 The advantage is that now we have replaced @{term"substs"} by
   116 @{term"map"}, we can profit from the large number of pre-proved lemmas
   116 @{term"map"}, we can profit from the large number of pre-proved lemmas
   117 about @{term"map"}.  Unfortunately inductive proofs about type
   117 about @{term"map"}.  Unfortunately inductive proofs about type
   118 \isa{term} are still awkward because they expect a conjunction. One
   118 @{text"term"} are still awkward because they expect a conjunction. One
   119 could derive a new induction principle as well (see
   119 could derive a new induction principle as well (see
   120 \S\ref{sec:derive-ind}), but turns out to be simpler to define
   120 \S\ref{sec:derive-ind}), but turns out to be simpler to define
   121 functions by \isacommand{recdef} instead of \isacommand{primrec}.
   121 functions by \isacommand{recdef} instead of \isacommand{primrec}.
   122 The details are explained in \S\ref{sec:advanced-recdef} below.
   122 The details are explained in \S\ref{sec:advanced-recdef} below.
   123 
   123 
   124 Of course, you may also combine mutual and nested recursion. For example,
   124 Of course, you may also combine mutual and nested recursion. For example,
   125 constructor \isa{Sum} in \S\ref{sec:datatype-mut-rec} could take a list of
   125 constructor @{text"Sum"} in \S\ref{sec:datatype-mut-rec} could take a list of
   126 expressions as its argument: \isa{Sum "'a aexp list"}.
   126 expressions as its argument: @{text"Sum"}~@{typ[quotes]"'a aexp list"}.
   127 *}
   127 *}
   128 (*<*)
   128 (*<*)
   129 end
   129 end
   130 (*>*)
   130 (*>*)