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