doc-src/TutorialI/Recdef/Nested2.thy
changeset 9721 7e51c9f3d5a0
parent 9690 50f22b1b136a
child 9754 a123a64cadeb
     1.1 --- a/doc-src/TutorialI/Recdef/Nested2.thy	Tue Aug 29 12:28:48 2000 +0200
     1.2 +++ b/doc-src/TutorialI/Recdef/Nested2.thy	Tue Aug 29 15:13:10 2000 +0200
     1.3 @@ -34,10 +34,21 @@
     1.4  both of which are solved by simplification:
     1.5  *};
     1.6  
     1.7 -by(simp_all del:map_compose add:sym[OF map_compose] rev_map);
     1.8 +by(simp_all add:rev_map sym[OF map_compose]);
     1.9  
    1.10  text{*\noindent
    1.11 -If this surprises you, see Datatype/Nested2......
    1.12 +If the proof of the induction step mystifies you, we recommend to go through
    1.13 +the chain of simplification steps in detail, probably with the help of
    1.14 +\isa{trace\_simp}.
    1.15 +%\begin{quote}
    1.16 +%{term[display]"trev(trev(App f ts))"}\\
    1.17 +%{term[display]"App f (rev(map trev (rev(map trev ts))))"}\\
    1.18 +%{term[display]"App f (map trev (rev(rev(map trev ts))))"}\\
    1.19 +%{term[display]"App f (map trev (map trev ts))"}\\
    1.20 +%{term[display]"App f (map (trev o trev) ts)"}\\
    1.21 +%{term[display]"App f (map (%x. x) ts)"}\\
    1.22 +%{term[display]"App f ts"}
    1.23 +%\end{quote}
    1.24  
    1.25  The above definition of @{term"trev"} is superior to the one in \S\ref{sec:nested-datatype}
    1.26  because it brings @{term"rev"} into play, about which already know a lot, in particular
    1.27 @@ -48,19 +59,22 @@
    1.28  because they determine the complexity of your proofs.}
    1.29  \end{quote}
    1.30  
    1.31 -Let us now return to the question of how \isacommand{recdef} can come up with sensible termination
    1.32 -conditions in the presence of higher-order functions like @{term"map"}. For a start, if nothing
    1.33 -were known about @{term"map"}, @{term"map trev ts"} might apply @{term"trev"} to arbitrary terms,
    1.34 -and thus \isacommand{recdef} would try to prove the unprovable
    1.35 -@{term"size t < Suc (term_size ts)"}, without any assumption about \isa{t}.
    1.36 -Therefore \isacommand{recdef} has been supplied with the congruence theorem \isa{map\_cong}: 
    1.37 +Let us now return to the question of how \isacommand{recdef} can come up with
    1.38 +sensible termination conditions in the presence of higher-order functions
    1.39 +like @{term"map"}. For a start, if nothing were known about @{term"map"},
    1.40 +@{term"map trev ts"} might apply @{term"trev"} to arbitrary terms, and thus
    1.41 +\isacommand{recdef} would try to prove the unprovable @{term"size t < Suc
    1.42 +(term_size ts)"}, without any assumption about \isa{t}.  Therefore
    1.43 +\isacommand{recdef} has been supplied with the congruence theorem
    1.44 +\isa{map\_cong}:
    1.45  \begin{quote}
    1.46  @{thm[display,margin=50]"map_cong"[no_vars]}
    1.47  \end{quote}
    1.48 -Its second premise expresses (indirectly) that the second argument of @{term"map"} is only applied
    1.49 -to elements of its third argument. Congruence rules for other higher-order functions on lists would
    1.50 -look very similar but have not been proved yet because they were never needed.
    1.51 -If you get into a situation where you need to supply \isacommand{recdef} with new congruence
    1.52 +Its second premise expresses (indirectly) that the second argument of
    1.53 +@{term"map"} is only applied to elements of its third argument. Congruence
    1.54 +rules for other higher-order functions on lists would look very similar but
    1.55 +have not been proved yet because they were never needed. If you get into a
    1.56 +situation where you need to supply \isacommand{recdef} with new congruence
    1.57  rules, you can either append the line
    1.58  \begin{ttbox}
    1.59  congs <congruence rules>