src/Doc/Corec/Corec.thy
author blanchet
Tue, 29 Mar 2016 21:25:19 +0200
changeset 62747 f65ef4723aca
parent 62745 257a022f7e7b
child 62756 d4b7d128ec5a
permissions -rw-r--r--
more 'corec' docs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
     1
(*  Title:      Doc/Corec/Corec.thy
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, Inria, LORIA, MPII
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
     3
    Author:     Aymeric Bouzy, Ecole polytechnique
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
     4
    Author:     Andreas Lochbihler, ETH Zuerich
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
     5
    Author:     Andrei Popescu, Middlesex University
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
     6
    Author:     Dmitriy Traytel, ETH Zuerich
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
     7
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
     8
Tutorial for nonprimitively corecursive definitions.
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
     9
*)
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
    10
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
    11
theory Corec
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
    12
imports
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    13
  GCD
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
    14
  "../Datatypes/Setup"
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
    15
  "~~/src/HOL/Library/BNF_Corec"
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    16
  "~~/src/HOL/Library/FSet"
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
    17
begin
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
    18
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
    19
section \<open>Introduction
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
    20
  \label{sec:introduction}\<close>
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
    21
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
    22
text \<open>
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
    23
...
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
    24
@{cite "isabelle-datatypes"}
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    25
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    26
* friend
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    27
* up to
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    28
62745
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
    29
* versioning
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
    30
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    31
BNF
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    32
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    33
link to papers
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
    34
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
    35
* plugins (code)
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
    36
\<close>
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
    37
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
    38
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
    39
section \<open>Introductory Examples
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
    40
  \label{sec:introductory-examples}\<close>
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
    41
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    42
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    43
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    44
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
    45
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
    46
subsection \<open>Simple Corecursion
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
    47
  \label{ssec:simple-corecursion}\<close>
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
    48
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    49
text \<open>
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
    50
The case studies by Rutten~@{cite rutten05} and Hinze~@{cite hinze10} on stream
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    51
calculi serve as our starting point. Streams can be defined as follows
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    52
(cf. @{file "~~/src/HOL/Library/Stream.thy"}):
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    53
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    54
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
    55
    codatatype 'a stream =
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    56
      SCons (shd: 'a) (stl: "'a stream")
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    57
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    58
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    59
The @{command corec} command makes it possible to define functions where the
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    60
corecursive call occurs under two or more constructors:
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    61
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    62
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    63
    corec oneTwos :: "nat stream" where
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    64
      "oneTwos = SCons 1 (SCons 2 oneTwos)"
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    65
62745
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
    66
thm oneTwos.cong_intros
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
    67
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    68
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    69
\noindent
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    70
This is already beyond the syntactic fragment supported by \keyw{primcorec}.
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    71
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    72
The following definition of pointwise sum can be performed with either
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    73
\keyw{primcorec} or @{command corec}:
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    74
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    75
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    76
    primcorec ssum :: "('a :: plus) stream \<Rightarrow> 'a stream \<Rightarrow> 'a stream" where
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    77
      "ssum xs ys = SCons (shd xs + shd ys) (ssum (stl xs) (stl ys))"
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    78
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    79
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    80
\noindent
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    81
Pointwise sum meets the friendliness criterion. We register it as a friend using
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    82
the @{command friend_of_corec} command. The command requires us to give a
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    83
specification of @{const ssum} where a constructor (@{const SCons}) occurs at
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    84
the outermost position on the right-hand side. Here, we can simply reuse the
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    85
\keyw{primcorec} specification above:
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    86
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    87
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    88
    friend_of_corec ssum :: "('a :: plus) stream \<Rightarrow> 'a stream \<Rightarrow> 'a stream" where
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    89
      "ssum xs ys = SCons (shd xs + shd ys) (ssum (stl xs) (stl ys))"
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    90
       apply (rule ssum.code)
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    91
      by transfer_prover
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    92
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    93
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    94
\noindent
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
    95
The command emits two subgoals. The first one corresponds to the equation we
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
    96
specified and is trivial to discharge. The second one is a parametricity subgoal
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    97
and can usually be discharged using the @{text transfer_prover} proof method.
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    98
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
    99
After registering @{const ssum} as a friend, we can use it in the corecursive
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   100
call context, either inside or outside the constructor guard:
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   101
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   102
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   103
    corec fibA :: "nat stream" where
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   104
      "fibA = SCons 0 (ssum (SCons 1 fibA) fibA)"
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   105
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   106
text \<open>\blankline\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   107
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   108
    corec fibB :: "nat stream" where
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   109
      "fibB = ssum (SCons 0 (SCons 1 fibB)) (SCons 0 fibB)"
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   110
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   111
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   112
Using the @{text "friend"} option, we can simultaneously define a function and
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   113
register it as a friend:
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   114
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   115
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   116
    corec (friend)
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   117
      sprod :: "('a :: {plus,times}) stream \<Rightarrow> 'a stream \<Rightarrow> 'a stream"
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   118
    where
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   119
      "sprod xs ys =
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   120
       SCons (shd xs * shd ys) (ssum (sprod xs (stl ys)) (sprod (stl xs) ys))"
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   121
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   122
text \<open>\blankline\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   123
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   124
    corec (friend) sexp :: "nat stream \<Rightarrow> nat stream" where
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   125
      "sexp xs = SCons (2 ^^ shd xs) (sprod (stl xs) (sexp xs))"
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   126
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   127
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   128
\noindent
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   129
The parametricity subgoal is given to @{text transfer_prover}.
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   130
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   131
The @{const sprod} and @{const sexp} functions provide shuffle product and
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   132
exponentiation on streams. We can use them to define the stream of factorial
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   133
numbers in two different ways:
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   134
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   135
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   136
    corec factA :: "nat stream" where
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   137
      "factA = (let zs = SCons 1 factA in sprod zs zs)"
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   138
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   139
    corec factB :: "nat stream" where
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   140
      "factB = sexp (SCons 0 factB)"
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   141
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   142
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   143
The arguments of friendly operations can be of complex types involving the
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   144
target codatatype. The following example defines the supremum of a finite set of
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   145
streams by primitive corecursion and registers it as friendly:
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   146
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   147
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   148
    corec (friend) sfsup :: "nat stream fset \<Rightarrow> nat stream" where
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   149
      "sfsup X = SCons (Sup (fset (fimage shd X))) (sfsup (fimage stl X))"
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   150
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   151
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   152
\noindent
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   153
In general, the arguments may be any BNF, with the restriction that the target
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   154
codatatype (@{typ "nat stream"}) may occur only in a live position of the BNF.
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   155
For this reason, the following operation, on unbounded sets, cannot be
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   156
registered as a friend:
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   157
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   158
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   159
    corec ssup :: "nat stream set \<Rightarrow> nat stream" where
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   160
      "ssup X = SCons (Sup (image shd X)) (ssup (image stl X))"
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   161
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   162
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   163
subsection \<open>Nested Corecursion
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   164
  \label{ssec:nested-corecursion}\<close>
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   165
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   166
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   167
The package generally supports arbitrary codatatypes with multiple constructors
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   168
and nesting through other type constructors (BNFs). Consider the following type
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   169
of finitely branching Rose trees of potentially infinite depth:
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   170
\<close>
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   171
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   172
    codatatype 'a tree =
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   173
      Node (lab: 'a) (sub: "'a tree list")
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   174
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   175
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   176
We first define the pointwise sum of two trees analogously to @{const ssum}:
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   177
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   178
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   179
    corec (friend) tplus :: "('a :: plus) tree \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   180
      "tplus t u =
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   181
       Node (lab t + lab u) (map (\<lambda>(t', u'). tplus t' u') (zip (sub t) (sub u)))"
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   182
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   183
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   184
\noindent
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   185
Here, @{const map} is the standard map function on lists, and @{const zip}
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   186
converts two parallel lists into a list of pairs. The @{const tplus} function is
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   187
primitively corecursive. Instead of @{text "corec (friend)"}, we could also have
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   188
used \keyw{primcorec} and @{command friend_of_corec}, as we did for
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   189
@{const ssum}.
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   190
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   191
Once @{const tplus} is registered as friendly, we can use it in the corecursive
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   192
call context:
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   193
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   194
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   195
    corec (friend) ttimes :: "('a :: {plus,times}) tree \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   196
      "ttimes t u = Node (lab t * lab u)
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   197
         (map (\<lambda>(t', u'). tplus (ttimes t u') (ttimes t' u)) (zip (sub t) (sub u)))"
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   198
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   199
text \<open>
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   200
All the syntactic convenience provided by \keyw{primcorec} is also supported by
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   201
@{command corec}, @{command corecursive}, and @{command friend_of_corec}. In
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   202
particular, nesting through the function type can be expressed using
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   203
@{text \<lambda>}-abstractions and function applications rather than through composition
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   204
(@{term "op \<circ>"}, the map function for @{text \<Rightarrow>}). For example:
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   205
\<close>
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   206
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   207
    codatatype 'a language =
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   208
      Lang (\<oo>: bool) (\<dd>: "'a \<Rightarrow> 'a language")
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   209
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   210
text \<open>\blankline\<close>
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   211
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   212
    corec (friend) Plus :: "'a language \<Rightarrow> 'a language \<Rightarrow> 'a language" where
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   213
      "Plus r s = Lang (\<oo> r \<or> \<oo> s) (\<lambda>a. Plus (\<dd> r a) (\<dd> s a))"
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   214
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   215
text \<open>\blankline\<close>
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   216
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   217
    corec (friend) Times :: "'a language \<Rightarrow> 'a language \<Rightarrow> 'a language" where
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   218
      "Times r s = Lang (\<oo> r \<and> \<oo> s)
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   219
         (\<lambda>a. if \<oo> r then Plus (Times (\<dd> r a) s) (\<dd> s a) else Times (\<dd> r a) s)"
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   220
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   221
text \<open>\blankline\<close>
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   222
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   223
    corec (friend) Star :: "'a language \<Rightarrow> 'a language" where
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   224
      "Star r = Lang True (\<lambda>a. Times (\<dd> r a) (Star r))"
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   225
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   226
text \<open>\blankline\<close>
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   227
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   228
    corec (friend) Inter :: "'a language \<Rightarrow> 'a language \<Rightarrow> 'a language" where
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   229
      "Inter r s = Lang (\<oo> r \<and> \<oo> s) (\<lambda>a. Inter (\<dd> r a) (\<dd> s a))"
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   230
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   231
text \<open>\blankline\<close>
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   232
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   233
    corec (friend) PLUS :: "'a language list \<Rightarrow> 'a language" where
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   234
      "PLUS xs = Lang (\<exists>x \<in> set xs. \<oo> x) (\<lambda>a. PLUS (map (\<lambda>r. \<dd> r a) xs))"
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   235
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   236
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   237
subsection \<open>Mixed Recursion--Corecursion
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   238
  \label{ssec:mixed-recursion-corecursion}\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   239
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   240
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   241
It is often convenient to let a corecursive function perform some finite
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   242
computation before producing a constructor. With mixed recursion--corecursion, a
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   243
finite number of unguarded recursive calls perform this calculation before
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   244
reaching a guarded corecursive call.
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   245
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   246
Intuitively, the unguarded recursive call can be unfolded to arbitrary finite
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   247
depth, ultimately yielding a purely corecursive definition. An example is the
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   248
@{term primes} function from Di Gianantonio and Miculan
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   249
@{cite "di-gianantonio-miculan-2003"}:
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   250
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   251
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   252
    corecursive primes :: "nat \<Rightarrow> nat \<Rightarrow> nat stream" where
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   253
      "primes m n =
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   254
       (if (m = 0 \<and> n > 1) \<or> coprime m n then
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   255
          SCons n (primes (m * n) (n + 1))
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   256
        else
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   257
          primes m (n + 1))"
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   258
      apply (relation "measure (\<lambda>(m, n).
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   259
        if n = 0 then 1 else if coprime m n then 0 else m - n mod m)")
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   260
       apply (auto simp: mod_Suc intro: Suc_lessI)
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   261
       apply (metis One_nat_def coprime_Suc_nat gcd.commute gcd_red_nat)
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   262
      by (metis diff_less_mono2 lessI mod_less_divisor)
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   263
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   264
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   265
\noindent
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   266
The @{command corecursive} command is a variant of @{command corec} that allows
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   267
us to specify a termination argument for any unguarded self-call.
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   268
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   269
When called with @{term "m = 1"} and @{term "n = 2"}, the @{const primes}
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   270
function computes the stream of prime numbers. The unguarded call in the
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   271
@{text else} branch increments @{term n} until it is coprime to the first
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   272
argument @{term m} (i.e., the greatest common divisor of @{term m} and
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   273
@{term n} is @{term 1}).
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   274
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   275
For any positive integers @{term m} and @{term n}, the numbers @{term m} and
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   276
@{term "m * n + 1"} are coprime, yielding an upper bound on the number of times
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   277
@{term n} is increased. Hence, the function will take the @{text else} branch at
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   278
most finitely often before taking the then branch and producing one constructor.
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   279
There is a slight complication when @{term "m = 0 \<and> n > 1"}: Without the first
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   280
disjunct in the @{text "if"} condition, the function could stall. (This corner
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   281
case was overlooked in the original example
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   282
@{cite "di-gianantonio-miculan-2003"}.)
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   283
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   284
In the following examples, termination is discharged automatically by
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   285
@{command corec} by invoking @{method lexicographic_order}:
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   286
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   287
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   288
    corec catalan :: "nat \<Rightarrow> nat stream" where
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   289
      "catalan n =
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   290
       (if n > 0 then ssum (catalan (n - 1)) (SCons 0 (catalan (n + 1)))
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   291
        else SCons 1 (catalan 1))"
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   292
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   293
    corec collatz :: "nat \<Rightarrow> nat stream" where
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   294
      "collatz n = (if even n \<and> n > 0 then collatz (n div 2)
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   295
         else SCons n (collatz (3 * n + 1)))"
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   296
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   297
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   298
A more elaborate case study, revolving around the filter function on lazy lists,
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   299
is presented in @{file "~~/src/HOL/Corec_Examples/LFilter.thy"}.
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   300
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   301
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   302
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   303
subsection \<open>Self-Friendship
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   304
  \label{ssec:self-friendship}\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   305
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   306
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   307
Paradoxically, the package allows us to simultaneously define a function and use
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   308
it as its own friend, as in the following definition of a ``skewed product'':
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   309
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   310
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   311
    corec (friend)
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   312
      sskew :: "('a :: {plus,times}) stream \<Rightarrow> 'a stream \<Rightarrow> 'a stream"
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   313
    where
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   314
      "sskew xs ys =
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   315
       SCons (shd xs * shd ys) (sskew (sskew xs (stl ys)) (sskew (stl xs) ys))"
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   316
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   317
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   318
\noindent
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   319
Such definitions, with nested self-calls on the right-hand side, cannot be
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   320
separated into a @{command corec} part and a @{command friend_of_corec} part.
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   321
\<close>
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   322
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   323
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   324
subsection \<open>Coinduction
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   325
  \label{ssec:coinduction}\<close>
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   326
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   327
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   328
Once a corecursive specification has been accepted, we normally want to reason
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   329
about it. The @{text codatatype} command generates a structural coinduction
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   330
principle that matches primitively corecursive functions. For nonprimitive
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   331
specifications, our package provides the more advanced proof principle of
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   332
\emph{coinduction up to congruence}---or simply \emph{coinduction up-to}.
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   333
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   334
The structural coinduction principle for @{typ "'a stream"}, called
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   335
@{thm [source] stream.coinduct}, is as follows:
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   336
%
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   337
\begin{indentblock}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   338
@{thm stream.coinduct[no_vars]}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   339
\end{indentblock}
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   340
%
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   341
Coinduction allows us to prove an equality @{text "l = r"} on streams by
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   342
providing a relation @{text R} that relates @{text l} and @{text r} (first
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   343
premise) and that constitutes a bisimulation (second premise). Streams that are
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   344
related by a bisimulation cannot be distinguished by taking observations (via
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   345
the selectors @{const shd} and @{const stl}); hence they must be equal.
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   346
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   347
The coinduction up-to principle after registering @{const sskew} as friendly is
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   348
available as @{thm [source] sskew.coinduct} and as one of the components of
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   349
the collection @{thm [source] stream.coinduct_upto}:
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   350
%
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   351
\begin{indentblock}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   352
@{thm sfsup.coinduct[no_vars]}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   353
\end{indentblock}
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   354
%
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   355
This rule is almost identical to structural coinduction, except that the
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   356
corecursive application of @{term R} is replaced by
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   357
@{term "stream.v5.congclp R"}. The @{const stream.v5.congclp} predicate is
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   358
equipped with the following introduction rules:
62745
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   359
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   360
\begin{indentblock}
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   361
\begin{description}
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   362
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   363
\item[@{thm [source] sskew.cong_base}\rm:] ~ \\
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   364
@{thm sskew.cong_base[no_vars]}
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   365
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   366
\item[@{thm [source] sskew.cong_refl}\rm:] ~ \\
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   367
@{thm sskew.cong_refl[no_vars]}
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   368
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   369
\item[@{thm [source] sskew.cong_sym}\rm:] ~ \\
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   370
@{thm sskew.cong_sym[no_vars]}
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   371
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   372
\item[@{thm [source] sskew.cong_trans}\rm:] ~ \\
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   373
@{thm sskew.cong_trans[no_vars]}
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   374
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   375
\item[@{thm [source] sskew.cong_SCons}\rm:] ~ \\
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   376
@{thm sskew.cong_SCons[no_vars]}
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   377
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   378
\item[@{thm [source] sskew.cong_ssum}\rm:] ~ \\
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   379
@{thm sskew.cong_ssum[no_vars]}
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   380
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   381
\item[@{thm [source] sskew.cong_sprod}\rm:] ~ \\
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   382
@{thm sskew.cong_sprod[no_vars]}
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   383
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   384
\item[@{thm [source] sskew.cong_sskew}\rm:] ~ \\
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   385
@{thm sskew.cong_sskew[no_vars]}
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   386
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   387
\end{description}
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   388
\end{indentblock}
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   389
%
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   390
The introduction rules are also available as
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   391
@{thm [source] sskew.cong_intros}.
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   392
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   393
Notice that there is no introduction rule corresponding to @{const sexp},
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   394
because @{const sexp} has a more restrictive result type than @{const sskew}
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   395
(@{typ "nat stream"} vs. @{typ "('a :: {plus,times}) stream"}.
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   396
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   397
Since the package maintains a set of incomparable corecursors, there is also a
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   398
set of associated coinduction principles and a set of sets of introduction
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   399
rules. A technically subtle point is to make Isabelle choose the right rules in
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   400
most situations. For this purpose, the package maintains the collection
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   401
@{thm [source] stream.coinduct_upto} of coinduction principles ordered by
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   402
increasing generality, which works well with Isabelle's philosophy of applying
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   403
the first rule that matches. For example, after registering @{const ssum} as a
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   404
friend, proving the equality @{term "l = r"} on @{typ "nat stream"} might
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   405
require coinduction principle for @{term "nat stream"}, which is up to
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   406
@{const ssum}.
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   407
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   408
The collection @{thm [source] stream.coinduct_upto} is guaranteed to be complete
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   409
and up to date with respect to the type instances of definitions considered so
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   410
far, but occasionally it may be necessary to take the union of two incomparable
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   411
coinduction principles. This can be done using the @{command coinduction_upto}
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   412
command. Consider the following definitions:
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   413
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   414
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   415
    codatatype ('a, 'b) tllist =
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   416
      TNil (terminal: 'b)
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   417
    | TCons (thd: 'a) (ttl: "('a, 'b) tllist")
62745
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   418
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   419
    corec (friend) square_elems :: "(nat, 'b) tllist \<Rightarrow> (nat, 'b) tllist" where
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   420
      "square_elems xs =
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   421
       (case xs of
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   422
         TNil z \<Rightarrow> TNil z
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   423
       | TCons y ys \<Rightarrow> TCons (y ^^ 2) (square_elems ys))"
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   424
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   425
    corec (friend) square_terminal :: "('a, int) tllist \<Rightarrow> ('a, int) tllist" where
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   426
      "square_terminal xs =
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   427
       (case xs of
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   428
         TNil z \<Rightarrow> TNil (z ^^ 2)
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   429
       | TCons y ys \<Rightarrow> TCons y (square_terminal ys))"
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   430
62745
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   431
text \<open>
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   432
At this point, @{thm [source] tllist.coinduct_upto} contains three variants of the
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   433
coinduction principles:
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   434
%
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   435
\begin{itemize}
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   436
\item @{typ "('a, int) tllist"} up to @{const TNil}, @{const TCons}, and
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   437
  @{const square_terminal};
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   438
\item @{typ "(nat, 'b) tllist"} up to @{const TNil}, @{const TCons}, and
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   439
  @{const square_elems};
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   440
\item @{typ "('a, 'b) tllist"} up to @{const TNil} and @{const TCons}.
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   441
\end{itemize}
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   442
%
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   443
The following variant is missing:
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   444
%
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   445
\begin{itemize}
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   446
\item @{typ "(nat, int) tllist"} up to @{const TNil}, @{const TCons},
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   447
  @{const square_elems}, and @{const square_terminal}.
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   448
\end{itemize}
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   449
%
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   450
To generate it, without having to define a new function with @{command corec},
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   451
we can use the following command:
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   452
\<close>
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   453
62745
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   454
    coinduction_upto nat_int_tllist: "(nat, int) tllist"
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   455
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   456
text \<open>
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   457
\noindent
62745
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   458
This produces the theorems @{thm [source] nat_int_tllist.coinduct_upto} and
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   459
@{thm [source] nat_int_tllist.cong_intros} (as well as the individually named
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   460
introduction rules) and extends @{thm [source] tllist.coinduct_upto} and
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   461
@{thm [source] tllist.cong_intros}.
62745
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   462
\<close>
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   463
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   464
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   465
subsection \<open>Uniqueness Reasoning
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   466
  \label{ssec:uniqueness-reasoning}\<close>
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   467
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   468
text \<open>
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   469
It is sometimes possible to achieve better automation by using a more
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   470
specialized proof method than coinduction. Uniqueness principles maintain a good
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   471
balance between expressiveness and automation. They exploit the property that a
62745
257a022f7e7b more 'corec' docs
blanchet
parents: 62742
diff changeset
   472
corecursive specification is the unique solution to a fixpoint equation.
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   473
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   474
The @{command corec}, @{command corecursive}, and @{command friend_of_corec}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   475
commands generate a property @{text f.unique} about the function of interest
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   476
@{term f} that can be used to prove that any function that satisfies
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   477
@{term f}'s corecursive specification must be equal to @{term f}. For example:
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   478
\[@{thm ssum.unique[no_vars]}\]
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   479
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   480
The uniqueness principles are not restricted to functions defined using
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   481
@{command corec} or @{command corecursive} or registered with
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   482
@{command friend_of_corec}. Suppose @{term "t x"} is an arbitrary term
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   483
depending on @{term x}. The @{method corec_unique} proof method, provided by our
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   484
tool, transforms subgoals of the form
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   485
\[@{term "(\<forall>x. f x = H x f) \<Longrightarrow> f x = t x"}\]
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   486
into
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   487
\[@{term "\<forall>x. t x = H x t"}\]
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   488
The higher-order functional @{term H} must be such that the equation
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   489
@{term "f x = H x f"} would be a valid @{command corec} specification---but
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   490
without nested self-calls or unguarded calls. Thus, @{method corec_unique}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   491
proves uniqueness of @{term t} with respect to the given fixpoint equation
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   492
regardless of how @{term t} was defined. For example:
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   493
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   494
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   495
    lemma
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   496
      fixes f :: "nat stream \<Rightarrow> nat stream \<Rightarrow> nat stream"
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   497
      assumes "\<forall>xs ys. f xs ys =
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   498
        SCons (shd ys * shd xs) (ssum (f xs (stl ys)) (f (stl xs) ys))"
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   499
      shows "f = sprod"
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   500
        using assms
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   501
      proof corec_unique
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   502
        show "sprod = (\<lambda>xs ys :: nat stream.
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   503
            SCons (shd ys * shd xs) (ssum (sprod xs (stl ys)) (sprod (stl xs) ys)))"
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   504
          apply (rule ext)+
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   505
          apply (subst sprod.code)
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   506
          by simp
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   507
      qed
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   508
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   509
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   510
section \<open>Command Syntax
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   511
  \label{sec:command-syntax}\<close>
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   512
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   513
subsection \<open>\keyw{corec} and \keyw{corecursive}
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   514
  \label{ssec:corec-and-corecursive}\<close>
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   515
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   516
text \<open>
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   517
\begin{matharray}{rcl}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   518
  @{command_def "corec"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   519
  @{command_def "corecursive"} & : & @{text "local_theory \<rightarrow> proof(prove)"}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   520
\end{matharray}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   521
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   522
@{rail \<open>
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   523
  (@@{command corec} | @@{command corecursive}) target? \<newline>
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   524
    @{syntax cr_options}? fix @'where' prop
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   525
  ;
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   526
  @{syntax_def cr_options}: '(' ((@{syntax plugins} | 'friend' | 'transfer') + ',') ')'
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   527
\<close>}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   528
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   529
\medskip
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   530
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   531
\noindent
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   532
The @{command corec} and @{command corecursive} commands introduce a corecursive
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   533
function over a codatatype.
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   534
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   535
The syntactic entity \synt{target} can be used to specify a local context,
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   536
\synt{fix} denotes name with an optional type signature, and \synt{prop} denotes
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   537
a HOL proposition @{cite "isabelle-isar-ref"}.
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   538
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   539
The optional target is optionally followed by a combination of the following
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   540
options:
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   541
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   542
\begin{itemize}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   543
\setlength{\itemsep}{0pt}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   544
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   545
\item
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   546
The @{text plugins} option indicates which plugins should be enabled
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   547
(@{text only}) or disabled (@{text del}). By default, all plugins are enabled.
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   548
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   549
\item
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   550
The @{text friend} option indicates that the defined function should be
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   551
registered as a friend. This gives rise to additional proof obligations.
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   552
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   553
\item
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   554
The @{text transfer} option indicates that an unconditional transfer rule
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   555
should be generated and proved @{text "by transfer_prover"}. The
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   556
@{text "[transfer_rule]"} attribute is set on the generated theorem.
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   557
\end{itemize}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   558
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   559
The @{command corec} command is an abbreviation for @{command corecursive} with
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   560
@{text "apply transfer_prover"} or @{text "apply lexicographic_order"} to
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   561
discharge any emerging proof obligations.
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   562
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   563
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   564
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   565
subsection \<open>\keyw{friend_of_corec}
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   566
  \label{ssec:friend-of-corec}\<close>
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   567
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   568
text \<open>
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   569
\begin{matharray}{rcl}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   570
  @{command_def "friend_of_corec"} & : & @{text "local_theory \<rightarrow> proof(prove)"}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   571
\end{matharray}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   572
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   573
@{rail \<open>
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   574
  @@{command friend_of_corec} target? \<newline>
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   575
    @{syntax foc_options}? fix @'where' prop
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   576
  ;
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   577
  @{syntax_def foc_options}: '(' ((@{syntax plugins} | 'transfer') + ',') ')'
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   578
\<close>}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   579
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   580
\medskip
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   581
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   582
\noindent
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   583
The @{command friend_of_corec} command registers a corecursive function as
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   584
friendly.
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   585
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   586
The syntactic entity \synt{target} can be used to specify a local context,
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   587
\synt{fix} denotes name with an optional type signature, and \synt{prop} denotes
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   588
a HOL proposition @{cite "isabelle-isar-ref"}.
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   589
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   590
The optional target is optionally followed by a combination of the following
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   591
options:
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   592
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   593
\begin{itemize}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   594
\setlength{\itemsep}{0pt}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   595
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   596
\item
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   597
The @{text plugins} option indicates which plugins should be enabled
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   598
(@{text only}) or disabled (@{text del}). By default, all plugins are enabled.
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   599
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   600
\item
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   601
The @{text transfer} option indicates that an unconditional transfer rule
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   602
should be generated and proved @{text "by transfer_prover"}. The
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   603
@{text "[transfer_rule]"} attribute is set on the generated theorem.
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   604
\end{itemize}
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   605
\<close>
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   606
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   607
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   608
subsection \<open>\keyw{coinduction_upto}
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   609
  \label{ssec:coinduction-upto}\<close>
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   610
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   611
text \<open>
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   612
\begin{matharray}{rcl}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   613
  @{command_def "coinduction_upto"} & : & @{text "local_theory \<rightarrow> local_theory"}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   614
\end{matharray}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   615
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   616
@{rail \<open>
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   617
  @@{command coinduction_upto} target? name ':' type
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   618
\<close>}
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   619
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   620
\medskip
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   621
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   622
\noindent
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   623
The @{command coinduction_upto} generates a coinduction up-to rule for a given
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   624
instance of a (possibly polymorphic) codatatype and notes the result with the
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   625
specified prefix.
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   626
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   627
The syntactic entity \synt{name} denotes an identifier and \synt{type} denotes a
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   628
type @{cite "isabelle-isar-ref"}.
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   629
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   630
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   631
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   632
section \<open>Generated Theorems
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   633
  \label{sec:generated-theorems}\<close>
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   634
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   635
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   636
...
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   637
\<close>
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   638
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   639
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   640
subsection \<open>\keyw{corec} and \keyw{corecursive}
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   641
  \label{ssec:corec-and-corecursive}\<close>
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   642
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   643
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   644
...
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   645
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   646
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   647
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   648
subsection \<open>\keyw{friend_of_corec}
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   649
  \label{ssec:friend-of-corec}\<close>
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   650
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   651
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   652
...
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   653
\<close>
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   654
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   655
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   656
subsection \<open>\keyw{coinduction_upto}
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   657
  \label{ssec:coinduction-upto}\<close>
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   658
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   659
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   660
...
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   661
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   662
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   663
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   664
section \<open>Proof Method
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   665
  \label{sec:proof-method}\<close>
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   666
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   667
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   668
...
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   669
\<close>
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   670
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   671
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   672
subsection \<open>\textit{corec_unique}
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   673
  \label{ssec:corec-unique}\<close>
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   674
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   675
text \<open>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   676
...
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   677
\<close>
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   678
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   679
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   680
section \<open>Known Bugs and Limitations
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   681
  \label{sec:known-bugs-and-limitations}\<close>
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   682
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   683
text \<open>
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   684
This section lists the known bugs and limitations of the corecursion package at
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   685
the time of this writing.
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   686
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   687
\begin{enumerate}
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   688
\setlength{\itemsep}{0pt}
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   689
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   690
\item
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   691
\emph{TODO.} TODO.
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   692
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   693
  * no mutual types
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   694
  * limitation on type of friend
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   695
  * unfinished tactics
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   696
  * polymorphism of corecUU_transfer
62742
bfb5a70e4319 more 'corec' documentation
blanchet
parents: 62739
diff changeset
   697
  * alternative views
62747
f65ef4723aca more 'corec' docs
blanchet
parents: 62745
diff changeset
   698
  * plugins
62739
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   699
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   700
\end{enumerate}
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   701
\<close>
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   702
628c97d39627 added sketchy 'corec' documentation
blanchet
parents:
diff changeset
   703
end