src/Doc/Datatypes/Datatypes.thy
changeset 54287 7f096d8eb3d0
parent 54278 c830ead80c91
child 54386 3514fdfdf59b
equal deleted inserted replaced
54286:22616f65d4ea 54287:7f096d8eb3d0
   922 non-functions) is handled in a more modular fashion. The old-style recursor can
   922 non-functions) is handled in a more modular fashion. The old-style recursor can
   923 be generated on demand using @{command primrec_new}, as explained in
   923 be generated on demand using @{command primrec_new}, as explained in
   924 Section~\ref{sssec:primrec-nested-as-mutual-recursion}, if the recursion is via
   924 Section~\ref{sssec:primrec-nested-as-mutual-recursion}, if the recursion is via
   925 new-style datatypes.
   925 new-style datatypes.
   926 
   926 
   927 \item \emph{Accordingly, the induction principle is different for nested
   927 \item \emph{Accordingly, the induction rule is different for nested recursive
   928 recursive datatypes.} Again, the old-style induction principle can be generated
   928 datatypes.} Again, the old-style induction rule can be generated on demand using
   929 on demand using @{command primrec_new}, as explained in
   929 @{command primrec_new}, as explained in
   930 Section~\ref{sssec:primrec-nested-as-mutual-recursion}, if the recursion is via
   930 Section~\ref{sssec:primrec-nested-as-mutual-recursion}, if the recursion is via
   931 new-style datatypes.
   931 new-style datatypes.
   932 
   932 
   933 \item \emph{The internal constructions are completely different.} Proof texts
   933 \item \emph{The internal constructions are completely different.} Proof texts
   934 that unfold the definition of constants introduced by \keyw{datatype} will be
   934 that unfold the definition of constants introduced by \keyw{datatype} will be
  1225             Zero \<Rightarrow> at\<^sub>f\<^sub>f t
  1225             Zero \<Rightarrow> at\<^sub>f\<^sub>f t
  1226           | Suc j' \<Rightarrow> ats\<^sub>f\<^sub>f ts j')"
  1226           | Suc j' \<Rightarrow> ats\<^sub>f\<^sub>f ts j')"
  1227 
  1227 
  1228 text {*
  1228 text {*
  1229 \noindent
  1229 \noindent
  1230 Appropriate induction principles are generated under the names
  1230 Appropriate induction rules are generated as
  1231 @{thm [source] at\<^sub>f\<^sub>f.induct},
  1231 @{thm [source] at\<^sub>f\<^sub>f.induct},
  1232 @{thm [source] ats\<^sub>f\<^sub>f.induct}, and
  1232 @{thm [source] ats\<^sub>f\<^sub>f.induct}, and
  1233 @{thm [source] at\<^sub>f\<^sub>f_ats\<^sub>f\<^sub>f.induct}.
  1233 @{thm [source] at\<^sub>f\<^sub>f_ats\<^sub>f\<^sub>f.induct}. The
  1234 
  1234 induction rules and the underlying recursors are generated on a per-need basis
  1235 %%% TODO: Add recursors.
  1235 and are kept in a cache to speed up subsequent definitions.
  1236 
  1236 
  1237 Here is a second example:
  1237 Here is a second example:
  1238 *}
  1238 *}
  1239 
  1239 
  1240     primrec_new
  1240     primrec_new
  1934 were mutually recursive
  1934 were mutually recursive
  1935 (Section~\ref{sssec:primrec-nested-as-mutual-recursion}), it is possible to
  1935 (Section~\ref{sssec:primrec-nested-as-mutual-recursion}), it is possible to
  1936 pretend that nested codatatypes are mutually corecursive. For example:
  1936 pretend that nested codatatypes are mutually corecursive. For example:
  1937 *}
  1937 *}
  1938 
  1938 
       
  1939 (*<*)
       
  1940     context late
       
  1941     begin
       
  1942 (*>*)
  1939     primcorec
  1943     primcorec
  1940       (*<*)(in late) (*>*)iterate\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>i" and
  1944       iterate\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>i" and
  1941       iterates\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a llist \<Rightarrow> 'a tree\<^sub>i\<^sub>i llist"
  1945       iterates\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a llist \<Rightarrow> 'a tree\<^sub>i\<^sub>i llist"
  1942     where
  1946     where
  1943       "iterate\<^sub>i\<^sub>i g x = Node\<^sub>i\<^sub>i x (iterates\<^sub>i\<^sub>i g (g x))" |
  1947       "iterate\<^sub>i\<^sub>i g x = Node\<^sub>i\<^sub>i x (iterates\<^sub>i\<^sub>i g (g x))" |
  1944       "iterates\<^sub>i\<^sub>i g xs =
  1948       "iterates\<^sub>i\<^sub>i g xs =
  1945          (case xs of
  1949          (case xs of
  1946             LNil \<Rightarrow> LNil
  1950             LNil \<Rightarrow> LNil
  1947           | LCons x xs' \<Rightarrow> LCons (iterate\<^sub>i\<^sub>i g x) (iterates\<^sub>i\<^sub>i g xs'))"
  1951           | LCons x xs' \<Rightarrow> LCons (iterate\<^sub>i\<^sub>i g x) (iterates\<^sub>i\<^sub>i g xs'))"
       
  1952 
       
  1953 text {*
       
  1954 \noindent
       
  1955 Coinduction rules are generated as
       
  1956 @{thm [source] iterate\<^sub>i\<^sub>i.coinduct},
       
  1957 @{thm [source] iterates\<^sub>i\<^sub>i.coinduct}, and
       
  1958 @{thm [source] iterate\<^sub>i\<^sub>i_iterates\<^sub>i\<^sub>i.coinduct}
       
  1959 and analogously for @{text strong_coinduct}. These rules and the
       
  1960 underlying corecursors are generated on a per-need basis and are kept in a cache
       
  1961 to speed up subsequent definitions.
       
  1962 *}
       
  1963 
       
  1964 (*<*)
       
  1965     end
       
  1966 (*>*)
  1948 
  1967 
  1949 
  1968 
  1950 subsubsection {* Constructor View
  1969 subsubsection {* Constructor View
  1951   \label{ssec:primrec-constructor-view} *}
  1970   \label{ssec:primrec-constructor-view} *}
  1952 
  1971