src/Doc/Corec/Corec.thy
changeset 62745 257a022f7e7b
parent 62742 bfb5a70e4319
child 62747 f65ef4723aca
equal deleted inserted replaced
62744:b4f139bf02e3 62745:257a022f7e7b
    24 \cite{isabelle-datatypes}
    24 \cite{isabelle-datatypes}
    25 
    25 
    26 * friend
    26 * friend
    27 * up to
    27 * up to
    28 
    28 
       
    29 * versioning
       
    30 
    29 BNF
    31 BNF
    30 
    32 
    31 link to papers
    33 link to papers
    32 \<close>
    34 \<close>
    33 
    35 
    58 corecursive call occurs under two or more constructors:
    60 corecursive call occurs under two or more constructors:
    59 \<close>
    61 \<close>
    60 
    62 
    61     corec oneTwos :: "nat stream" where
    63     corec oneTwos :: "nat stream" where
    62       "oneTwos = SCons 1 (SCons 2 oneTwos)"
    64       "oneTwos = SCons 1 (SCons 2 oneTwos)"
       
    65 
       
    66 thm oneTwos.cong_intros
    63 
    67 
    64 text \<open>
    68 text \<open>
    65 \noindent
    69 \noindent
    66 This is already beyond the syntactic fragment supported by \keyw{primcorec}.
    70 This is already beyond the syntactic fragment supported by \keyw{primcorec}.
    67 
    71 
   304 %
   308 %
   305 This rule is almost identical to structural coinduction, except that the
   309 This rule is almost identical to structural coinduction, except that the
   306 corecursive application of @{term R} is replaced by
   310 corecursive application of @{term R} is replaced by
   307 @{term "stream.v5.congclp R"}. The @{const stream.v5.congclp} predicate is
   311 @{term "stream.v5.congclp R"}. The @{const stream.v5.congclp} predicate is
   308 equipped with the following introduction rules:
   312 equipped with the following introduction rules:
   309 \<close>
   313 
   310 
   314 \begin{indentblock}
   311 (* FIXME:
   315 \begin{description}
   312 thm
   316 
   313   sskew.cong_base
   317 \item[@{thm [source] sskew.cong_base}\rm:] ~ \\
   314   sskew.cong_refl
   318 @{thm sskew.cong_base[no_vars]}
   315   sskew.cong_intros
   319 
   316 
   320 \item[@{thm [source] sskew.cong_refl}\rm:] ~ \\
   317 thm stream.coinduct_upto(2)
   321 @{thm sskew.cong_refl[no_vars]}
   318   sskew.coinduct
   322 
   319 
   323 \item[@{thm [source] sskew.cong_sym}\rm:] ~ \\
   320 thm stream.cong_intros
   324 @{thm sskew.cong_sym[no_vars]}
   321 thm sfsup.cong_intros
   325 
   322 *)
   326 \item[@{thm [source] sskew.cong_trans}\rm:] ~ \\
       
   327 @{thm sskew.cong_trans[no_vars]}
       
   328 
       
   329 \item[@{thm [source] sskew.cong_SCons}\rm:] ~ \\
       
   330 @{thm sskew.cong_SCons[no_vars]}
       
   331 
       
   332 \item[@{thm [source] sskew.cong_ssum}\rm:] ~ \\
       
   333 @{thm sskew.cong_ssum[no_vars]}
       
   334 
       
   335 \item[@{thm [source] sskew.cong_sprod}\rm:] ~ \\
       
   336 @{thm sskew.cong_sprod[no_vars]}
       
   337 
       
   338 \item[@{thm [source] sskew.cong_sskew}\rm:] ~ \\
       
   339 @{thm sskew.cong_sskew[no_vars]}
       
   340 
       
   341 \end{description}
       
   342 \end{indentblock}
       
   343 %
       
   344 The introduction rules are also available as
       
   345 @{thm [source] sskew.cong_intros}.
       
   346 
       
   347 Notice that there is no introduction rule corresponding to @{const sexp},
       
   348 because @{const sexp} has a more restrictive result type than @{const sskew}
       
   349 (@{typ "nat stream"} vs. @{typ "('a :: {plus,times}) stream"}.
       
   350 
       
   351 Since the package maintains a set of incomparable corecursors, there is also a
       
   352 set of associated coinduction principles and a set of sets of introduction
       
   353 rules. A technically subtle point is to make Isabelle choose the right rules in
       
   354 most situations. For this purpose, the package maintains the collection
       
   355 @{thm [source] stream.coinduct_upto} of coinduction principles ordered by
       
   356 increasing generality, which works well with Isabelle's philosophy of applying
       
   357 the first rule that matches. For example, after registering @{const ssum} as a
       
   358 friend, proving the equality @{term "l = r"} on @{typ "nat stream"} might
       
   359 require coinduction principle for @{term "nat stream"}, which is up to
       
   360 @{const ssum}.
       
   361 
       
   362 The collection @{thm [source] stream.coinduct_upto} is guaranteed to be complete
       
   363 and up to date with respect to the type instances of definitions considered so
       
   364 far, but occasionally it may be necessary to take the union of two incomparable
       
   365 coinduction principles. This can be done using the @{command coinduction_upto}
       
   366 command. Consider the following definitions:
       
   367 \<close>
       
   368 
       
   369     codatatype (tset: 'a, 'b) tllist =
       
   370       TNil (terminal : 'b)
       
   371     | TCons (thd : 'a) (ttl : "('a, 'b) tllist")
       
   372     for
       
   373       map: tmap
       
   374       rel: tllist_all2
       
   375 
       
   376     corec (friend) square_elems :: "(nat, 'b) tllist \<Rightarrow> (nat, 'b) tllist" where
       
   377       "square_elems xs =
       
   378        (case xs of
       
   379          TNil z \<Rightarrow> TNil z
       
   380        | TCons y ys \<Rightarrow> TCons (y ^^ 2) (square_elems ys))"
       
   381 
       
   382     corec (friend) square_terminal :: "('a, int) tllist \<Rightarrow> ('a, int) tllist" where
       
   383       "square_terminal xs =
       
   384        (case xs of
       
   385          TNil z \<Rightarrow> TNil (z ^^ 2)
       
   386        | TCons y ys \<Rightarrow> TCons y (square_terminal ys))"
       
   387 
       
   388 text \<open>
       
   389 At this point, @{thm [source] tllist.coinduct_upto} contains three variants of the
       
   390 coinduction principles:
       
   391 %
       
   392 \begin{itemize}
       
   393 \item @{typ "('a, int) tllist"} up to @{const TNil}, @{const TCons}, and
       
   394   @{const square_terminal};
       
   395 \item @{typ "(nat, 'b) tllist"} up to @{const TNil}, @{const TCons}, and
       
   396   @{const square_elems};
       
   397 \item @{typ "('a, 'b) tllist"} up to @{const TNil} and @{const TCons}.
       
   398 \end{itemize}
       
   399 %
       
   400 The following variant is missing:
       
   401 %
       
   402 \begin{itemize}
       
   403 \item @{typ "(nat, int) tllist"} up to @{const TNil}, @{const TCons},
       
   404   @{const square_elems}, and @{const square_terminal}.
       
   405 \end{itemize}
       
   406 %
       
   407 To generate it, without having to define a new function with @{command corec},
       
   408 we can use the following command:
       
   409 \<close>
       
   410 
       
   411     coinduction_upto nat_int_tllist: "(nat, int) tllist"
       
   412 
       
   413 text \<open>
       
   414 This produces the theorems @{thm [source] nat_int_tllist.coinduct_upto} and
       
   415 @{thm [source] nat_int_tllist.cong_intros} (as well as the individually named
       
   416 introduction rules), and extends @{thm [source] tllist.coinduct_upto}.
       
   417 \<close>
   323 
   418 
   324 
   419 
   325 subsection \<open>Uniqueness Reasoning
   420 subsection \<open>Uniqueness Reasoning
   326   \label{ssec:uniqueness-reasoning}\<close>
   421   \label{ssec:uniqueness-reasoning}\<close>
   327 
   422 
   328 text \<open>
   423 text \<open>
   329 ...
   424 t is sometimes possible to achieve better automation by using a more specialized
       
   425 proof method than coinduction. Uniqueness principles maintain a good balance
       
   426 between expressiveness and automation. They exploit the property that a
       
   427 corecursive specification is the unique solution to a fixpoint equation.
   330 \<close>
   428 \<close>
   331 
   429 
   332 
   430 
   333 section \<open>Command Syntax
   431 section \<open>Command Syntax
   334   \label{sec:command-syntax}\<close>
   432   \label{sec:command-syntax}\<close>