src/Doc/Codegen/Partial_Functions.thy
author nipkow
Fri, 16 Dec 2022 18:18:57 +0100
changeset 76650 a2c23c68f699
child 76673 059a68d21f0f
permissions -rw-r--r--
file with partial function docu
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
76650
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
     1
theory Partial_Functions
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
     2
imports Main "HOL-Library.Monad_Syntax"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
     3
begin
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
     4
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
     5
section \<open>Partial Functions \label{sec:partial}\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
     6
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
     7
text \<open>We demonstrate three approaches to defining executable partial recursive functions.
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
     8
The main points are the definitions of the functions and the inductive proofs about them.
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
     9
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    10
Our concrete example represents a typical termination problem: following a data structure that
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    11
may contain cycles. We want to follow a mapping from \<open>nat\<close> to \<open>nat\<close> to the end
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    12
(until we leave its domain). The mapping is represented by a list \<open>ns :: nat list\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    13
that maps \<open>n\<close> to \<open>ns ! n\<close>.
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    14
The details of the example are in some sense irrelevant but make the exposition more realistic.
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    15
However, we hide most proofs or show only the characteristic opening.\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    16
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    17
text \<open>The list representation of the mapping can be abstracted to a relation.
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    18
The order @{term "(ns!n, n)"} is the order that @{const wf} expects.\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    19
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    20
definition rel :: "nat list \<Rightarrow> (nat * nat) set" where
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    21
"rel ns = set(zip ns [0..<length ns])"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    22
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    23
lemma finite_rel[simp]: "finite(rel ns)"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    24
(*<*)by(simp add: rel_def)(*>*)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    25
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    26
text \<open>This relation should be acyclic
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    27
 to guarantee termination of the partial functions defined below.\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    28
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    29
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    30
subsection \<open>Tail recursion\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    31
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    32
text \<open>If a function is tail-recursive, an executable definition is easy:\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    33
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    34
partial_function (tailrec) follow :: "nat list \<Rightarrow> nat \<Rightarrow> nat" where
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    35
"follow ns n = (if n < length ns then follow ns (ns!n) else n)"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    36
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    37
text \<open>Informing the code generator:\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    38
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    39
declare follow.simps[code]
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    40
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    41
text \<open>Now @{const follow} is executable:\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    42
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    43
value "follow [1,2,3] 0"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    44
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    45
text \<open>For proofs about @{const follow} we need a @{const wf} relation on @{term "(ns,n)"} pairs
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    46
that decreases with each recursive call. The first component stays the same but must be acyclic.
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    47
The second component must decrease w.r.t @{const rel}:\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    48
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    49
definition "rel_follow = same_fst (acyclic o rel) rel"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    50
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    51
lemma wf_follow: "wf rel_follow"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    52
(*<*)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    53
by (auto intro: wf_same_fst simp: rel_follow_def finite_acyclic_wf)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    54
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    55
text \<open>A more explicit formulation of \<open>rel_follow\<close>:
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    56
The first component stays the same but must be acyclic.
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    57
The second component decreases w.r.t \<open>rel\<close>:\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    58
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    59
lemma rel_follow_step:
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    60
  "\<lbrakk> m < length ms; acyclic (rel ms) \<rbrakk> \<Longrightarrow> ((ms, ms ! m), (ms, m)) \<in> rel_follow"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    61
by(force simp: rel_follow_def rel_def in_set_zip)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    62
(*>*)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    63
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    64
text \<open>This is how you start an inductive proof about @{const follow} along @{const rel_follow}:\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    65
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    66
lemma "acyclic(rel ms) \<Longrightarrow> follow ms m = n \<Longrightarrow> length ms \<le> n"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    67
proof (induction "(ms,m)" arbitrary: m n rule: wf_induct_rule[OF wf_follow])
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    68
(*<*)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    69
  case 1
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    70
  thus ?case using follow.simps rel_follow_step by fastforce
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    71
qed
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    72
(*>*)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    73
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    74
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    75
subsection \<open>Option\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    76
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    77
text \<open>If the function is not tail-recursive, not all is lost: if we rewrite it to return an \<open>option\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    78
type, it can still be defined. In this case @{term "Some x"} represents the result \<open>x\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    79
and @{term None} represents represents nontermination.
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    80
For example, counting the length of the chain represented by \<open>ns\<close> can be defined like this:\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    81
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    82
partial_function (option) count :: "nat list \<Rightarrow> nat \<Rightarrow> nat option" where
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    83
"count ns n
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    84
= (if n < length ns then do {k \<leftarrow> count ns (ns!n); Some (k+1)} else Some 0)"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    85
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    86
text \<open>We use a Haskell-like \<open>do\<close> notation (import @{theory "HOL-Library.Monad_Syntax"})
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    87
to abbreviate the clumsy explicit
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    88
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    89
\noindent
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    90
@{term "case count ns (ns!n) of Some k \<Rightarrow> Some(k+1) | None \<Rightarrow> None"}.
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    91
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    92
\noindent
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    93
The branch \<open>None \<Rightarrow> None\<close> represents the requirement
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    94
that nontermination of a recursive call must lead to nontermination of the function.
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    95
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    96
Now we can prove that @{const count} terminates for all acyclic maps:\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    97
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    98
lemma "acyclic(rel ms) \<Longrightarrow> \<exists>k. count ms m = Some k"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
    99
proof (induction "(ms,m)" arbitrary: ms m rule: wf_induct_rule[OF wf_follow])
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   100
(*<*)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   101
  case 1
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   102
  thus ?case
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   103
    by (metis bind.bind_lunit count.simps rel_follow_step)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   104
qed
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   105
(*>*)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   106
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   107
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   108
subsection \<open>Subtype\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   109
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   110
text \<open>In this approach we define a new type that contains only elements on which the function
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   111
in question terminates.
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   112
In our example that is the subtype of all \<open>ns :: nat list\<close> such that @{term "rel ns"} is acyclic.
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   113
Then @{const follow} can be defined as a total function on that subtype.\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   114
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   115
text \<open>The subtype is not empty:\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   116
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   117
lemma acyclic_rel_Nil: "acyclic(rel [])"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   118
(*<*)by (simp add: rel_def acyclic_def)(*>*)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   119
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   120
text \<open>Definition of subtype \<open>acyc\<close>:\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   121
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   122
typedef acyc = "{ns. acyclic(rel ns)}"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   123
morphisms rep_acyc abs_acyc
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   124
using acyclic_rel_Nil by auto
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   125
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   126
text \<open>This defines two functions \<open>rep_acyc :: acyc \<Rightarrow> nat list\<close> and \<open>abs_acyc :: nat list \<Rightarrow> acyc\<close>.
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   127
Function @{const abs_acyc} is only defined on acyclic lists and not executable for that reason.\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   128
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   129
setup_lifting type_definition_acyc
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   130
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   131
text \<open>This is how @{const follow} can be defined on @{typ acyc}:\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   132
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   133
function follow2 :: "acyc \<Rightarrow> nat \<Rightarrow> nat" where
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   134
"follow2 ac n
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   135
= (let ns = rep_acyc ac in if n < length ns then follow2 ac (ns!n) else n)"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   136
by pat_completeness auto
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   137
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   138
text \<open>Now we prepare for termination proof.
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   139
 Relation \<open>rel_follow2\<close> is almost identical to @{const rel_follow}.\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   140
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   141
definition "rel_follow2 = same_fst (acyclic o rel o rep_acyc) (rel o rep_acyc)"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   142
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   143
lemma wf_follow2: "wf rel_follow2"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   144
(*<*)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   145
by (auto intro: wf_same_fst simp add: rel_follow2_def finite_acyclic_wf)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   146
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   147
text \<open>A more explicit formulation of \<open>rel_follow2\<close>:\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   148
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   149
lemma rel_follow2_step:
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   150
  "\<lbrakk> ns = rep_acyc ac; m < length ns; acyclic (rel ns) \<rbrakk> \<Longrightarrow> ((ac, ns ! m), (ac, m)) \<in> rel_follow2"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   151
by(force simp add: rel_follow2_def rel_def in_set_zip)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   152
(*>*)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   153
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   154
text \<open>Here comes the actual termination proof:\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   155
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   156
termination follow2
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   157
proof
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   158
  show "wf rel_follow2"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   159
(*<*)    by (auto intro: wf_same_fst simp add: rel_follow2_def finite_acyclic_wf)(*>*)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   160
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   161
next
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   162
  show "\<And>ac n ns. \<lbrakk> ns = rep_acyc ac; n < length ns \<rbrakk>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   163
              \<Longrightarrow> ((ac, ns ! n), (ac, n)) \<in> rel_follow2"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   164
(*<*)    using rel_follow2_step rep_acyc by simp(*>*)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   165
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   166
qed
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   167
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   168
text \<open>Inductive proofs about @{const follow2} can now simply use computation induction:\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   169
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   170
lemma "follow2 ac m = n \<Longrightarrow> length (rep_acyc ac) \<le> n"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   171
proof (induction ac m arbitrary: n rule: follow2.induct)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   172
(*<*)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   173
  case 1
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   174
  thus ?case by (metis (full_types) follow2.simps linorder_not_le)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   175
qed
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   176
(*>*)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   177
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   178
text \<open>A complication with the subtype approach is that injection into the subtype
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   179
(function \<open>abs_acyc\<close> in our example) is not executable. But to call @{const follow2},
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   180
we need an argument of type \<open>acyc\<close> and we need to obtain it in an executable manner.
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   181
There are two approaches.\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   182
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   183
text \<open>In the first approach we check wellformedness (i.e. acyclicity) explicitly.
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   184
This check needs to be executable (which @{const acyclic} and @{const rel} are).
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   185
If the check fails, @{term "[]"} is returned (which is acyclic).\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   186
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   187
lift_definition is_acyc :: "nat list \<Rightarrow> acyc" is 
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   188
  "\<lambda>ns. if acyclic(rel ns) then ns else []"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   189
(*<*)by (auto simp: acyclic_rel_Nil)(*>*)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   190
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   191
text \<open>This works because we can easily prove that for any \<open>ns\<close>,
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   192
 the \<open>\<lambda>\<close>-term produces an acyclic list.
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   193
But it requires the possibly expensive check @{prop "acyclic(rel ns)"}.\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   194
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   195
definition "follow_test ns n = follow2 (is_acyc ns) n"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   196
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   197
text \<open>The relation is acyclic (a chain):\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   198
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   199
value "follow_test [1,2,3] 1"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   200
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   201
text \<open>In the second approach, wellformedness of the argument is guaranteed by construction.
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   202
In our example \<open>[1..<n+1]\<close> represents an acyclic chain \<open>i \<mapsto> i+1\<close>\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   203
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   204
lemma acyclic_chain: "acyclic (rel [1..<n+1])"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   205
(*<*)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   206
apply(induction n)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   207
 apply (simp add: rel_def acyclic_def)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   208
apply (auto simp add: rel_def)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   209
by (metis atLeast_upt lessI lessThan_iff order_less_asym' rtranclE set_zip_rightD)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   210
(*>*)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   211
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   212
text \<open>\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   213
lift_definition acyc_chain :: "nat \<Rightarrow> acyc" is "\<lambda>n. [1..<n+1]"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   214
(*<*)by (use acyclic_chain in auto)(*>*)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   215
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   216
text \<open>\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   217
definition "follow_chain m n = follow2 (acyc_chain m) n"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   218
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   219
value "follow_chain 5 1"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   220
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   221
text \<open>The subtype approach requires neither tail-recursion nor \<open>option\<close> but you cannot easily modify
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   222
arguments of the subtype except via existing functions on the subtype. Otherwise you need to inject
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   223
some value into the subtype and that injection is not computable.
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   224
\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   225
(*<*)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   226
text \<open>The problem with subtypes: the Abs function must not be used explicitly.
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   227
This can be avoided by using \<open>lift_definition\<close>. Example:\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   228
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   229
typedef nat1 = "{n::nat. n > 0}"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   230
by auto
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   231
print_theorems
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   232
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   233
setup_lifting type_definition_nat1
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   234
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   235
(* just boiler plate *)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   236
lift_definition mk1 :: "nat \<Rightarrow> nat1" is 
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   237
  "\<lambda>n. if n>0 then n else 1"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   238
by auto
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   239
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   240
lift_definition g :: "nat1 \<Rightarrow> nat1" is
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   241
"\<lambda>n. if n \<ge> 2 then n-1 else n"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   242
by auto
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   243
print_theorems
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   244
text \<open>This generates
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   245
 \<open>g.rep_eq: Rep_nat1 (g x) = (if 2 \<le> Rep_nat1 x then Rep_nat1 x - 1 else Rep_nat1 x)\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   246
which is acceptable as an abstract code eqn. The manual definition of
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   247
 \<open>g n1 = (let n = Rep_nat1 n1 in if 2 \<le> n then Abs_nat1 (n - 1) else Abs_nat1 n)\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   248
looks non-executable but \<open>g.rep_eq\<close> can be derived from it.\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   249
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   250
value "g (mk1 3)"
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   251
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   252
text \<open>However, this trick does not work when passing an Abs-term as an argument,
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   253
eg in a recursive call, because the Abs-term is `hidden' by the function call.\<close>
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   254
(*>*)
a2c23c68f699 file with partial function docu
nipkow
parents:
diff changeset
   255
end