doc-src/TutorialI/Trie/Trie.thy
author nipkow
Wed, 19 Apr 2000 12:59:38 +0200
changeset 8749 2665170f104a
parent 8745 13b32661dde4
child 8771 026f37a86ea7
permissions -rw-r--r--
Adding generated files
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     1
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     2
theory Trie = Main:;
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     3
(*>*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     4
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     5
text{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     6
To minimize running time, each node of a trie should contain an array that maps
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     7
letters to subtries. We have chosen a (sometimes) more space efficient
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     8
representation where the subtries are held in an association list, i.e.\ a
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     9
list of (letter,trie) pairs.  Abstracting over the alphabet \isa{'a} and the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    10
values \isa{'v} we define a trie as follows:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    11
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    12
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    13
datatype ('a,'v)trie = Trie  "'v option"  "('a * ('a,'v)trie)list";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    14
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    15
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    16
The first component is the optional value, the second component the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    17
association list of subtries.  This is an example of nested recursion involving products,
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    18
which is fine because products are datatypes as well.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    19
We define two selector functions:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    20
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    21
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    22
consts value :: "('a,'v)trie \\<Rightarrow> 'v option"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    23
       alist :: "('a,'v)trie \\<Rightarrow> ('a * ('a,'v)trie)list";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    24
primrec "value(Trie ov al) = ov";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    25
primrec "alist(Trie ov al) = al";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    26
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    27
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    28
Association lists come with a generic lookup function:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    29
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    30
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    31
consts   assoc :: "('key * 'val)list \\<Rightarrow> 'key \\<Rightarrow> 'val option";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    32
primrec "assoc [] x = None"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    33
        "assoc (p#ps) x =
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    34
           (let (a,b) = p in if a=x then Some b else assoc ps x)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    35
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    36
text{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    37
Now we can define the lookup function for tries. It descends into the trie
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    38
examining the letters of the search string one by one. As
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    39
recursion on lists is simpler than on tries, let us express this as primitive
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    40
recursion on the search string argument:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    41
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    42
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    43
consts   lookup :: "('a,'v)trie \\<Rightarrow> 'a list \\<Rightarrow> 'v option";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    44
primrec "lookup t [] = value t"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    45
        "lookup t (a#as) = (case assoc (alist t) a of
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    46
                              None \\<Rightarrow> None
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    47
                            | Some at \\<Rightarrow> lookup at as)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    48
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    49
text{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    50
As a first simple property we prove that looking up a string in the empty
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    51
trie \isa{Trie~None~[]} always returns \isa{None}. The proof merely
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    52
distinguishes the two cases whether the search string is empty or not:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    53
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    54
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    55
lemma [simp]: "lookup (Trie None []) as = None";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    56
apply(cases "as", auto).;
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    57
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    58
text{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    59
Things begin to get interesting with the definition of an update function
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    60
that adds a new (string,value) pair to a trie, overwriting the old value
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    61
associated with that string:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    62
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    63
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    64
consts update :: "('a,'v)trie \\<Rightarrow> 'a list \\<Rightarrow> 'v \\<Rightarrow> ('a,'v)trie";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    65
primrec
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    66
  "update t []     v = Trie (Some v) (alist t)"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    67
  "update t (a#as) v =
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    68
     (let tt = (case assoc (alist t) a of
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    69
                  None \\<Rightarrow> Trie None [] | Some at \\<Rightarrow> at)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    70
      in Trie (value t) ((a,update tt as v)#alist t))";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    71
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    72
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    73
The base case is obvious. In the recursive case the subtrie
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    74
\isa{tt} associated with the first letter \isa{a} is extracted,
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    75
recursively updated, and then placed in front of the association list.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    76
The old subtrie associated with \isa{a} is still in the association list
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    77
but no longer accessible via \isa{assoc}. Clearly, there is room here for
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    78
optimizations!
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    79
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    80
Before we start on any proofs about \isa{update} we tell the simplifier to
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    81
expand all \isa{let}s and to split all \isa{case}-constructs over
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    82
options:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    83
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    84
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    85
theorems [simp] = Let_def;
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    86
theorems [split] = option.split;
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    87
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    88
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    89
The reason becomes clear when looking (probably after a failed proof
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    90
attempt) at the body of \isa{update}: it contains both
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    91
\isa{let} and a case distinction over type \isa{option}.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    92
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    93
Our main goal is to prove the correct interaction of \isa{update} and
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    94
\isa{lookup}:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    95
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    96
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    97
theorem "\\<forall>t v bs. lookup (update t as v) bs =
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    98
                    (if as=bs then Some v else lookup t bs)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    99
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   100
txt{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   101
Our plan is to induct on \isa{as}; hence the remaining variables are
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   102
quantified. From the definitions it is clear that induction on either
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   103
\isa{as} or \isa{bs} is required. The choice of \isa{as} is merely
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   104
guided by the intuition that simplification of \isa{lookup} might be easier
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   105
if \isa{update} has already been simplified, which can only happen if
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   106
\isa{as} is instantiated.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   107
The start of the proof is completely conventional:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   108
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   109
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   110
apply(induct_tac "as", auto);
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   111
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   112
txt{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   113
Unfortunately, this time we are left with three intimidating looking subgoals:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   114
\begin{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   115
~1.~\dots~{\isasymLongrightarrow}~lookup~\dots~bs~=~lookup~t~bs\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   116
~2.~\dots~{\isasymLongrightarrow}~lookup~\dots~bs~=~lookup~t~bs\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   117
~3.~\dots~{\isasymLongrightarrow}~lookup~\dots~bs~=~lookup~t~bs%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   118
\end{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   119
Clearly, if we want to make headway we have to instantiate \isa{bs} as
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   120
well now. It turns out that instead of induction, case distinction
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   121
suffices:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   122
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   123
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   124
apply(case_tac[!] bs);
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   125
apply(auto).;
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   126
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   127
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   128
Both \isaindex{case_tac} and \isaindex{induct_tac}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   129
take an optional first argument that specifies the range of subgoals they are
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   130
applied to, where \isa{!} means all subgoals, i.e.\ \isa{[1-3]}. Individual
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   131
subgoal number are also allowed.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   132
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   133
This proof may look surprisingly straightforward. However, note that this
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   134
comes at a cost: the proof script is unreadable because the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   135
intermediate proof states are invisible, and we rely on the (possibly
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   136
brittle) magic of \isa{auto} (after the induction) to split the remaining
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   137
goals up in such a way that case distinction on \isa{bs} makes sense and
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   138
solves the proof. Part~\ref{Isar} shows you how to write readable and stable
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   139
proofs.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   140
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   141
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   142
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   143
end
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   144
(*>*)