doc-src/TutorialI/Sets/Recur.thy
author paulson
Thu, 26 Oct 2000 11:27:48 +0200
changeset 10341 6eb91805a012
parent 10294 2ec9c808a8a7
child 16417 9bc16273c2d4
permissions -rw-r--r--
added the $Id:$ line
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10341
6eb91805a012 added the $Id:$ line
paulson
parents: 10294
diff changeset
     1
(* ID:         $Id$ *)
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
     2
theory Recur = Main:
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
     3
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
     4
ML "Pretty.setmargin 64"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
     5
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
     6
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
     7
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
     8
@{thm[display] mono_def[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
     9
\rulename{mono_def}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    10
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    11
@{thm[display] monoI[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    12
\rulename{monoI}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    13
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    14
@{thm[display] monoD[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    15
\rulename{monoD}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    16
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    17
@{thm[display] lfp_unfold[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    18
\rulename{lfp_unfold}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    19
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    20
@{thm[display] lfp_induct[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    21
\rulename{lfp_induct}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    22
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    23
@{thm[display] gfp_unfold[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    24
\rulename{gfp_unfold}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    25
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    26
@{thm[display] coinduct[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    27
\rulename{coinduct}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    28
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    29
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    30
text{*\noindent
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    31
A relation $<$ is
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    32
\bfindex{wellfounded} if it has no infinite descending chain $\cdots <
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    33
a@2 < a@1 < a@0$. Clearly, a function definition is total iff the set
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    34
of all pairs $(r,l)$, where $l$ is the argument on the left-hand side
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    35
of an equation and $r$ the argument of some recursive call on the
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    36
corresponding right-hand side, induces a wellfounded relation. 
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    37
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    38
The HOL library formalizes
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    39
some of the theory of wellfounded relations. For example
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    40
@{prop"wf r"}\index{*wf|bold} means that relation @{term[show_types]"r::('a*'a)set"} is
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    41
wellfounded.
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    42
Finally we should mention that HOL already provides the mother of all
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    43
inductions, \textbf{wellfounded
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    44
induction}\indexbold{induction!wellfounded}\index{wellfounded
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    45
induction|see{induction, wellfounded}} (@{thm[source]wf_induct}):
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    46
@{thm[display]wf_induct[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    47
where @{term"wf r"} means that the relation @{term r} is wellfounded
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    48
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    49
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    50
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    51
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    52
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    53
@{thm[display] wf_induct[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    54
\rulename{wf_induct}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    55
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    56
@{thm[display] less_than_iff[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    57
\rulename{less_than_iff}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    58
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    59
@{thm[display] inv_image_def[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    60
\rulename{inv_image_def}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    61
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    62
@{thm[display] measure_def[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    63
\rulename{measure_def}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    64
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    65
@{thm[display] wf_less_than[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    66
\rulename{wf_less_than}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    67
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    68
@{thm[display] wf_inv_image[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    69
\rulename{wf_inv_image}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    70
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    71
@{thm[display] wf_measure[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    72
\rulename{wf_measure}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    73
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    74
@{thm[display] lex_prod_def[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    75
\rulename{lex_prod_def}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    76
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    77
@{thm[display] wf_lex_prod[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    78
\rulename{wf_lex_prod}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    79
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    80
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    81
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    82
end
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    83