--- a/README.md Fri Jun 23 15:16:30 2017 +0000
+++ b/README.md Fri Jun 23 18:08:31 2017 +0200
@@ -1,5 +1,5 @@
# README #
-### Functional Data Structures ###
+### Functional Data Structures 2017 ###
-Repository of slides
\ No newline at end of file
+Private repository of slides etc
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/ROOT Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,27 @@
+session "HOL-FDS17" = HOL +
+ options [document = pdf, document_output="$THIS_DIR/FDS-generated"]
+ theories [document = false]
+ "~~/src/HOL/Library/Tree"
+ "../Public/Thys/Tree_Additions"
+ "~~/src/HOL/Data_Structures/Tree_Set"
+ "~~/src/HOL/Data_Structures/Tree23_Set"
+ "../Public/Thys/RBT_Set"
+ "../Public/Thys/Trie2"
+ "../Public/Thys/Leftist_Heap"
+ "$AFP/Priority_Queue_Braun/Priority_Queue_Braun"
+ "Slides/Tree_Notation"
+
+session Slides in Slides = "HOL-FDS17" +
+ options [document_output = tex, document = pdf, names_short = true]
+ theories[show_question_marks=false]
+ Sorting_Slides
+ Tree_Slides
+ Search_Trees
+ Tree23_Slides
+ RBT_Slides
+ Trie_Slides
+ Leftist_Heap_Slides
+ Braun_Tree_Slides
+ document_files
+ "root.tex"
+ "prelude.tex"
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/Braun_Tree_Slides.thy Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,114 @@
+(*<*)
+theory Braun_Tree_Slides
+imports
+ Tree_Notation
+ "$AFP/Priority_Queue_Braun/Priority_Queue_Braun"
+begin
+notation (output) insert_pq ("insert")
+
+lemma del_min_simp3: "l \<noteq> Leaf \<Longrightarrow>
+ del_min (Node l x r) = (let (y,l') = del_left l in sift_down r y l')"
+by(auto simp: neq_Leaf_iff)
+lemmas del_min_simps = del_min.simps(1,2) del_min_simp3
+(*>*)
+text_raw{*
+\end{isabellebody}%
+%-------------------------------------------------------------
+\section{Priority Queues Based on Braun Trees}
+%-------------------------------------------------------------
+\afpentry{Priority_Queue_Braun}
+%-------------------------------------------------------------
+\begin{frame}{What is a Braun tree?}
+@{const_typ braun}\\\smallskip\pause
+\high{@{thm[break,margin=60] braun.simps}}\\
+\bigskip\pause
+
+\textbf{Lemma} \high{@{thm height_size_braun}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Priority queue implementation}
+Implementation type: ordinary binary trees
+\bigskip\pause
+
+Invariants: @{const heap} and @{const braun}
+\bigskip\pause
+
+\begin{center}
+\bad{No \<open>merge\<close>} \pause --- \high{@{const insert_pq} and @{const del_min} defined explicitly}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{@{const insert_pq}}
+@{const_typ insert_pq}\\\smallskip\pause
+\high{@{thm [break,margin=60] insert_pq.simps}}
+\bigskip\pause
+
+Correctness and preservation of invariant straightforward.
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{@{const del_min}}
+@{const_typ del_min}\\\smallskip\pause
+\high{@{thm [break,margin=60] (concl) del_min_simps}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{@{const sift_down}}
+\high<1>{@{const_typ sift_down}}\\\smallskip\pause
+\high<2>{@{thm [break,margin=60] (sub) sift_down.simps(1)}}\\\smallskip\pause
+\high<3>{@{thm [break,margin=60] (sub) sift_down.simps(2)}}\\\smallskip\pause
+\high{@{thm [break,margin=60] (sub) sift_down.simps(3)}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Functional correctness proofs for deletion}{including preservation of @{const heap} and @{const braun}}
+Many lemmas, mostly straightforward
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Logarithmic complexity}
+Running time of @{const insert_pq}, @{const del_left}
+and @{const sift_down} (and therefore @{const del_min}) bounded by height
+\bigskip\pause
+
+Remember: @{thm height_size_braun}
+\medskip\pause
+
+\<open>\<Longrightarrow>\<close>
+\medskip
+
+Above running times logarithmic in size
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}[t]{Sorting with priority queue}
+\pause
+
+@{text "pq [] = empty"}\\
+@{text "pq (x#xs) = insert x (pq xs)"}
+\bigskip\pause
+
+@{prop[break,margin=50] "mins q = (if is_empty q then [] else get_min h # mins (del_min h))"}
+\bigskip\pause
+
+\high{@{prop"sort_pq = mins \<circ> pq"}}
+\bigskip\pause
+
+Complexity of \<open>sort\<close>: $O(n \log n)$ \\\pause
+if all priority queue functions have complexity $O(\log n)$
+%\bigskip\pause
+
+%Braun trees $\leadsto$ functional version of heap sort
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}[t]{Sorting with priority queue}
+
+@{text "pq [] = empty"}\\
+@{text "pq (x#xs) = insert x (pq xs)"}
+\bigskip\pause
+
+\bad{Not optimal.} \pause\high{Linear time possible.}
+
+\end{frame}
+%-------------------------------------------------------------
+
+\begin{isabellebody}%
+*}
+(*<*)
+end
+(*>*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/FDSlatex.thy Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,109 @@
+theory FDSlatex
+imports Complex_Main
+begin
+
+(* LOGIC *)
+
+notation (latex output)
+ If ("(\<^latex>\<open>{\\sf \<close>if\<^latex>\<open>}\<close> (_)/ \<^latex>\<open>{\\sf \<close>then\<^latex>\<open>}\<close> (_)/ \<^latex>\<open>{\\sf \<close>else\<^latex>\<open>}\<close> (_))" 10)
+
+syntax (latex output)
+
+ "_Let" :: "[letbinds, 'a] => 'a"
+ ("(\<^latex>\<open>{\\sf \<close>let\<^latex>\<open>}\<close> (_)/ \<^latex>\<open>{\\sf \<close>in\<^latex>\<open>}\<close> (_))" 10)
+
+ "_case_syntax":: "['a, cases_syn] => 'b"
+ ("(\<^latex>\<open>{\\sf \<close>case\<^latex>\<open>}\<close> _ \<^latex>\<open>{\\sf \<close>of\<^latex>\<open>}\<close>// _)" 10)
+
+ "_case2" :: "[case_syn, cases_syn] \<Rightarrow> cases_syn" ("_//| _")
+
+
+abbreviation (iffSpace)
+ iffSpace :: "[bool, bool] => bool" (infixr "\<^latex>\<open>\\ \<close>\<longleftrightarrow>\<^latex>\<open>\\ \<close>" 25)
+where
+ "iffSpace A B == A = B"
+
+(* DUMMY *)
+consts DUMMY :: 'a ("\<^latex>\<open>\\_\<close>")
+
+definition dummyid :: "'a \<Rightarrow> 'a" where
+"dummyid x = x"
+
+notation (latex output) dummyid ("_")
+
+translations "x" <= "CONST Real.real x"
+
+(* log *)
+notation (latex output)
+ log ("\<^latex>\<open>$\\log$\<close>\<^bsub>_\<^esub>")
+
+notation (latex output)
+ power ("(_\<^bsup>\<^latex>\<open>\\isa{\<close>_\<^latex>\<open>}\<close>\<^esup>)" [1000] 1000)
+
+(* THEOREMS *)
+(*
+notation (Rule output)
+ "Pure.imp" ("\<^latex>\<open>\\mbox{}\\inferrule{\\mbox{\<close>_\<^latex>\<open>}}\<close>\<^latex>\<open>{\\mbox{\<close>_\<^latex>\<open>}}\<close>")
+
+syntax (Rule output)
+ "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
+ ("\<^latex>\<open>\\mbox{}\\inferrule{\<close>_\<^latex>\<open>}\<close>\<^latex>\<open>{\\mbox{\<close>_\<^latex>\<open>}}\<close>")
+
+ "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms"
+ ("\<^latex>\<open>\\mbox{\<close>_\<^latex>\<open>}\\\\\<close>/ _")
+
+ "_asm" :: "prop \<Rightarrow> asms" ("\<^latex>\<open>\\mbox{\<close>_\<^latex>\<open>}\<close>")
+
+notation (Axiom output)
+ "Trueprop" ("\<^latex>\<open>\\mbox{}\\inferrule{\\mbox{}}{\\mbox{\<close>_\<^latex>\<open>}}\<close>")
+*)
+notation (IfThen output)
+ "Pure.imp" ("\<^latex>\<open>{\\textsf{}\<close>If\<^latex>\<open>\\,}\<close> _/ \<^latex>\<open>{\\normalsize \\,\<close>then\<^latex>\<open>\\,}\<close>/ _.")
+syntax (IfThen output)
+ "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
+ ("\<^latex>\<open>{\\sf \<close>If\<^latex>\<open>\\,}\<close> _ /\<^latex>\<open>{\\sf \\,\<close>then\<^latex>\<open>\\,}\<close>/ _.")
+ "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("\<^latex>\<open>\\mbox{\<close>_\<^latex>\<open>}\<close> /\<^latex>\<open>{\\sf \\,\<close>and\<^latex>\<open>\\,}\<close>/ _")
+ "_asm" :: "prop \<Rightarrow> asms" ("\<^latex>\<open>\\mbox{\<close>_\<^latex>\<open>}\<close>")
+
+notation (IfThenNoBox output)
+ "Pure.imp" ("\<^latex>\<open>{\\normalsize{}\<close>If\<^latex>\<open>\\,}\<close> _/ \<^latex>\<open>{\\normalsize \\,\<close>then\<^latex>\<open>\\,}\<close>/ _.")
+syntax (IfThenNoBox output)
+ "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
+ ("\<^latex>\<open>{\\normalsize{}\<close>If\<^latex>\<open>\\,}\<close> _ /\<^latex>\<open>{\\normalsize \\,\<close>then\<^latex>\<open>\\,}\<close>/ _.")
+ "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("_ /\<^latex>\<open>{\\normalsize \\,\<close>and\<^latex>\<open>\\,}\<close>/ _")
+ "_asm" :: "prop \<Rightarrow> asms" ("_")
+
+
+setup\<open>
+ let
+ fun pretty ctxt c =
+ let val tc = Proof_Context.read_const {proper = true, strict = false} ctxt c
+ in Pretty.block [Thy_Output.pretty_term ctxt tc, Pretty.str " ::",
+ Pretty.brk 1, Syntax.pretty_typ ctxt (fastype_of tc)]
+ end
+ in
+ Thy_Output.antiquotation @{binding "const_typ"}
+ (Scan.lift Args.embedded_inner_syntax)
+ (fn {source = src, context = ctxt, ...} => fn arg =>
+ Thy_Output.output ctxt
+ (Thy_Output.maybe_pretty_source pretty ctxt src [arg]))
+ end;
+\<close>
+
+setup\<open>
+let
+ fun dummy_pats (wrap $ (eq $ lhs $ rhs)) =
+ let
+ val rhs_vars = Term.add_vars rhs [];
+ fun dummy (v as Var (ixn as (_, T))) =
+ if member (op = ) rhs_vars ixn then v else Const (@{const_name DUMMY}, T)
+ | dummy (t $ u) = dummy t $ dummy u
+ | dummy (Abs (n, T, b)) = Abs (n, T, dummy b)
+ | dummy t = t;
+ in wrap $ (eq $ dummy lhs $ rhs) end
+in
+ Term_Style.setup @{binding dummy_pats} (Scan.succeed (K dummy_pats))
+end
+\<close>
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/Leftist_Heap_Slides.thy Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,285 @@
+(*<*)
+theory Leftist_Heap_Slides
+imports
+ FDSlatex
+ "../../Public/Thys/Leftist_Heap"
+begin
+
+notation size1 ("|_|\<^sub>1")
+(*>*)
+text_raw{*
+\end{isabellebody}%
+\newcommand{\revealbeamer}[1]{\uncover<+->{\invisible<handout>{#1}}}
+\section{Priority Queues}
+%-------------------------------------------------------------
+\begin{frame}{Priority queue informally}
+\begin{center}
+Collection of elements with priorities
+\end{center}
+\pause
+Operations:
+\begin{itemize}
+\pause
+\item empty
+\pause
+\item emptiness test
+\pause
+\item insert
+\pause
+\item get element with minimal priority
+\pause
+\item delete element with minimal priority
+\end{itemize}
+\bigskip\pause
+\begin{center}
+We focus on the priorities:\\
+\high{element = priority}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Priority queues are multisets}
+\begin{center}
+The same element can be contained \high{multiple times}\\ in a priority queue
+\pause
+
+$\Longrightarrow$
+
+The abstract view of a priority queue is a \high{multiset}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Multisets in Isabelle}
+
+Import \texttt{"Library/Multiset"}
+\end{frame}
+%-------------------------------------------------------------
+*}
+text_raw (in Priority_Queue) {*
+%-------------------------------------------------------------
+\begin{frame}{Interface of implementation}
+\begin{center}
+The type of elements (= priorities) \ \high{@{typ 'a}} \ is a linear order
+\end{center}
+\pause
+An implementation of a priority queue of elements of type \ @{typ 'a} \ must provide
+\begin{itemize}
+\pause
+\item An implementation type \ \high{@{typ 'q}}
+\pause
+\item \high{\<open>empty ::\<close> @{typ 'q}}
+\pause
+\item \high{\<open>is_empty ::\<close> @{typ "'q \<Rightarrow> bool"}}
+\pause
+\item \high{\<open>insert ::\<close> @{typ"'a \<Rightarrow> 'q \<Rightarrow> 'q"}}
+\pause
+\item \high{\<open>get_min ::\<close> @{typ"'q \<Rightarrow> 'a"}}
+\pause
+\item \high{\<open>del_min ::\<close> @{typ"'q \<Rightarrow> 'q"}}
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{More operations}
+\begin{itemize}
+\item \high{\<open>merge ::\<close> @{typ"'q \<Rightarrow> 'q \<Rightarrow> 'q"}}\\
+Often provided
+\pause
+\item decrease key/priority\\ Not easy in functional setting
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of implementation}
+A priority queue represents a \high{multiset} of priorities.
+
+Correctness proof requires:
+\begin{center}
+\begin{tabular}{ll}
+Abstraction function: & \high{@{term[source]"mset :: 'q \<Rightarrow> 'a multiset"}}
+\\\pause
+Invariant: & \high{@{term[source]"invar :: 'q \<Rightarrow> bool"}}
+\end{tabular}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of implementation}
+Must prove \ \high{@{prop"invar q"}} \<open>\<Longrightarrow>\<close>\medskip
+
+\pause
+\high{@{thm mset_empty}}\smallskip\\\pause
+\high{@{thm (concl) is_empty}}\smallskip\\\pause
+\high{@{thm (concl) mset_insert}}\smallskip\\\pause
+\high{@{thm (concl) mset_del_min}}\smallskip\\\pause
+\high{@{thm (prem 2) get_min} \<open>\<Longrightarrow>\<close>}\\
+\high{@{text "get_min q \<in> set q"} \<open>\<and>\<close>}
+\high{@{text[source] "(\<forall>x \<in> set q. get_min q \<le> x)"}}\\
+where \<open>set q = set_mset (mset q)\<close>
+\bigskip\pause
+
+\high{@{prop"invar empty"}}\\
+\high{@{thm (concl) invar_insert}}\\
+\high{@{thm (concl) invar_del_min}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Terminology}
+\begin{center}
+A tree is a \high{heap} if for every subtree\\
+the root is $\ge$ all elements in the subtrees.
+\end{center}
+\pause
+
+The term ``heap'' is frequently used synonymously with ``priority queue''.
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Priority queue via heap}
+\begin{itemize}
+\pause
+\item \<open>empty = \<langle>\<rangle>\<close>
+\pause
+\item \<open>is_empty h = (h = \<langle>\<rangle>)\<close>
+\pause
+\item \<open>get_min \<langle>_, a, _\<rangle> = a\<close>
+\pause
+\item Assume we have \<open>merge\<close>
+\pause
+\item @{text"insert a t = merge \<langle>\<langle>\<rangle>, a, \<langle>\<rangle>\<rangle> t"}
+\pause
+\item \<open>del_min \<langle>l, a, r\<rangle> = merge l r\<close>
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Priority queue via heap}
+A naive merge:
+\medskip\pause
+
+\<open>merge t\<^sub>1 t\<^sub>2 = (\<^latex>\<open>{\sf case}\<close> (t\<^sub>1,t\<^sub>2) \<^latex>\<open>{\sf of}\<close>\<close>\\
+\quad\<open>(\<langle>\<rangle>, _) \<Rightarrow> t\<^sub>2 |\<close>\\
+\quad@{text"(_, \<langle>\<rangle>) \<Rightarrow> t\<^sub>1 |"}\\\pause
+\quad@{text"(\<langle>l\<^sub>1,a\<^sub>1,r\<^sub>1\<rangle>, \<langle>l\<^sub>2,a\<^sub>2,r\<^sub>2\<rangle>) \<Rightarrow>"}\\
+\qquad@{text"\<^latex>\<open>{\\sf if}\<close> a\<^sub>1 \<le> a\<^sub>2 \<^latex>\<open>{\\sf then}\<close> \<langle>merge l\<^sub>1 r\<^sub>1, a\<^sub>1, t\<^sub>2\<rangle>"}\\
+\qquad@{text"\<^latex>\<open>{\\sf else}\<close> \<langle>t\<^sub>1, a\<^sub>2, merge l\<^sub>2 r\<^sub>2\<rangle>"}
+\bigskip\pause
+
+\bad{Challenge:} how to maintaining some kind of balance
+\end{frame}
+%-------------------------------------------------------------
+*}
+text_raw {*
+%-------------------------------------------------------------
+\section{Leftist Heap}
+%-------------------------------------------------------------
+\thyfile{Data\char`_Structures/Leftist\char`_Heap}{}
+%-------------------------------------------------------------
+\begin{frame}{Leftist tree informally}
+The \high{rank} of a tree is the depth of the rightmost leaf.
+\pause\bigskip
+
+In a \high{leftist tree}, the rank of every left child is $\ge$
+the rank of its right sibling
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Implementation type}
+\isakeyword{datatype}\\
+\quad \high{@{text "'a lheap = Leaf | Node nat ('a tree) 'a ('a tree)"}}
+\medskip\pause
+
+Abbreviations \high{@{const Leaf}} and \high{@{term "\<langle>h, l, a, r\<rangle>"}} as usual
+\bigskip\pause
+
+Abstraction function:\\
+@{const mset_tree} \<open>::\<close> @{typ "'a lheap \<Rightarrow> 'a multiset"}\\\smallskip\pause
+\high{@{thm mset_tree.simps(1)}}\\
+\high{@{thm[break,margin=50] (dummy_pats) mset_tree.simps(2)}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Heap}
+@{const heap} \<open>::\<close> @{typ "'a lheap \<Rightarrow> bool"}\\\smallskip\pause
+\high{@{thm heap.simps(1)}}\\
+\high{@{text"heap \<langle>_, l, a, r\<rangle> = (heap l \<and> heap r \<and>"}}\\
+\quad\high{@{text"(\<forall>x \<in># mset_tree l + mset_tree r. a \<le> x))"}}
+
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Leftist tree}
+@{const rank} \<open>::\<close> @{typ "'a lheap \<Rightarrow> nat"}\\\smallskip\pause
+\high{@{thm rank.simps(1)}}\\
+\high{@{thm (dummy_pats) rank.simps(2)}}
+\bigskip\pause
+
+Node @{term "Node n l a r"}: \<open>n = \<close> rank of node
+\bigskip\pause
+
+@{const ltree} \<open>::\<close> @{typ "'a lheap \<Rightarrow> bool"}\\\smallskip\pause
+\high{@{thm ltree.simps(1)}}\\
+\high{@{thm[break,margin=60] (dummy_pats) ltree.simps(2)}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Leftist heap invariant}
+\begin{center}
+\high{@{prop"invar h = (heap h \<and> ltree h)"}}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Why leftist tree?}
+\pause
+
+\textbf{Lemma} \high{@{thm pow2_rank_size1}}
+\bigskip\pause
+
+\textbf{Lemma} Execution time of @{term "merge t\<^sub>1 t\<^sub>2"} is bounded by
+\mbox{@{term"rank t\<^sub>1 + rank t\<^sub>2"}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{@{const merge}}
+Principle: descend on the right
+\smallskip\pause
+
+\high{@{thm[break] (sub) merge.simps(1,2)}}\\\pause
+\high{@{thm[break] (sub,dummy_pats) merge.simps(3)}}
+\bigskip\pause
+
+@{const node} \<open>::\<close> @{typ "'a lheap \<Rightarrow> 'a \<Rightarrow> 'a lheap \<Rightarrow> 'a lheap"}\\\smallskip\pause
+\high{@{thm[break,margin=65] (sub) node_def}}\\
+where @{thm (dummy_pats) rk.simps(2)}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{@{const merge}}
+@{thm[break] (sub,dummy_pats) merge.simps(3)}
+\bigskip\pause
+
+\high{\onslide<2->{Function @{const merge} terminates because}
+\onslide<3>{\revealbeamer{@{term"rank t\<^sub>1 + rank t\<^sub>2"}}}
+\onslide<2->{decreases with every recursive call.}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Functional correctness proofs}{including preservation of \<open>invar\<close>}
+Straightforward
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Logarithmic complexity}
+Complexity measures @{const t_merge}, @{const t_insert} @{const t_del_min}:\\
+ count calls of @{const merge}.
+\pause\bigskip
+
+\textbf{Lemma} @{thm t_merge_rank}
+\pause
+
+\textbf{Lemma} @{thm[break,margin=40] t_merge_log}
+\pause\bigskip
+
+\textbf{Lemma}\\ @{thm t_insert_log}
+\pause\bigskip
+
+\textbf{Lemma}\\ @{thm t_del_min_log}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}
+\begin{center}
+\bad{Can we avoid the rank info in each node?}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+
+\begin{isabellebody}%
+*}
+(*<*)
+end
+(*>*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/Leftist_Heap_Slides.thy~ Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,285 @@
+(*<*)
+theory Leftist_Heap_Slides
+imports
+ FDSlatex
+ "../Material/Thys/Leftist_Heap"
+begin
+
+notation size1 ("|_|\<^sub>1")
+(*>*)
+text_raw{*
+\end{isabellebody}%
+\newcommand{\revealbeamer}[1]{\uncover<+->{\invisible<handout>{#1}}}
+\section{Priority Queues}
+%-------------------------------------------------------------
+\begin{frame}{Priority queue informally}
+\begin{center}
+Collection of elements with priorities
+\end{center}
+\pause
+Operations:
+\begin{itemize}
+\pause
+\item empty
+\pause
+\item emptiness test
+\pause
+\item insert
+\pause
+\item get element with minimal priority
+\pause
+\item delete element with minimal priority
+\end{itemize}
+\bigskip\pause
+\begin{center}
+We focus on the priorities:\\
+\high{element = priority}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Priority queues are multisets}
+\begin{center}
+The same element can be contained \high{multiple times}\\ in a priority queue
+\pause
+
+$\Longrightarrow$
+
+The abstract view of a priority queue is a \high{multiset}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Multisets in Isabelle}
+
+Import \texttt{"Library/Multiset"}
+\end{frame}
+%-------------------------------------------------------------
+*}
+text_raw (in Priority_Queue) {*
+%-------------------------------------------------------------
+\begin{frame}{Interface of implementation}
+\begin{center}
+The type of elements (= priorities) \ \high{@{typ 'a}} \ is a linear order
+\end{center}
+\pause
+An implementation of a priority queue of elements of type \ @{typ 'a} \ must provide
+\begin{itemize}
+\pause
+\item An implementation type \ \high{@{typ 'q}}
+\pause
+\item \high{\<open>empty ::\<close> @{typ 'q}}
+\pause
+\item \high{\<open>is_empty ::\<close> @{typ "'q \<Rightarrow> bool"}}
+\pause
+\item \high{\<open>insert ::\<close> @{typ"'a \<Rightarrow> 'q \<Rightarrow> 'q"}}
+\pause
+\item \high{\<open>get_min ::\<close> @{typ"'q \<Rightarrow> 'a"}}
+\pause
+\item \high{\<open>del_min ::\<close> @{typ"'q \<Rightarrow> 'q"}}
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{More operations}
+\begin{itemize}
+\item \high{\<open>merge ::\<close> @{typ"'q \<Rightarrow> 'q \<Rightarrow> 'q"}}\\
+Often provided
+\pause
+\item decrease key/priority\\ Not easy in functional setting
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of implementation}
+A priority queue represents a \high{multiset} of priorities.
+
+Correctness proof requires:
+\begin{center}
+\begin{tabular}{ll}
+Abstraction function: & \high{@{term[source]"mset :: 'q \<Rightarrow> 'a multiset"}}
+\\\pause
+Invariant: & \high{@{term[source]"invar :: 'q \<Rightarrow> bool"}}
+\end{tabular}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of implementation}
+Must prove \ \high{@{prop"invar q"}} \<open>\<Longrightarrow>\<close>\medskip
+
+\pause
+\high{@{thm mset_empty}}\smallskip\\\pause
+\high{@{thm (concl) is_empty}}\smallskip\\\pause
+\high{@{thm (concl) mset_insert}}\smallskip\\\pause
+\high{@{thm (concl) mset_del_min}}\smallskip\\\pause
+\high{@{thm (prem 2) get_min} \<open>\<Longrightarrow>\<close>}\\
+\high{@{text "get_min q \<in> set q"} \<open>\<and>\<close>}
+\high{@{text[source] "(\<forall>x \<in> set q. get_min q \<le> x)"}}\\
+where \<open>set q = set_mset (mset q)\<close>
+\bigskip\pause
+
+\high{@{prop"invar empty"}}\\
+\high{@{thm (concl) invar_insert}}\\
+\high{@{thm (concl) invar_del_min}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Terminology}
+\begin{center}
+A tree is a \high{heap} if for every subtree\\
+the root is $\ge$ all elements in the subtrees.
+\end{center}
+\pause
+
+The term ``heap'' is frequently used synonymously with ``priority queue''.
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Priority queue via heap}
+\begin{itemize}
+\pause
+\item \<open>empty = \<langle>\<rangle>\<close>
+\pause
+\item \<open>is_empty h = (h = \<langle>\<rangle>)\<close>
+\pause
+\item \<open>get_min \<langle>_, a, _\<rangle> = a\<close>
+\pause
+\item Assume we have \<open>merge\<close>
+\pause
+\item @{text"insert a t = merge \<langle>\<langle>\<rangle>, a, \<langle>\<rangle>\<rangle> t"}
+\pause
+\item \<open>del_min \<langle>l, a, r\<rangle> = merge l r\<close>
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Priority queue via heap}
+A naive merge:
+\medskip\pause
+
+\<open>merge t\<^sub>1 t\<^sub>2 = (\<^latex>\<open>{\sf case}\<close> (t\<^sub>1,t\<^sub>2) \<^latex>\<open>{\sf of}\<close>\<close>\\
+\quad\<open>(\<langle>\<rangle>, _) \<Rightarrow> t\<^sub>2 |\<close>\\
+\quad@{text"(_, \<langle>\<rangle>) \<Rightarrow> t\<^sub>1 |"}\\\pause
+\quad@{text"(\<langle>l\<^sub>1,a\<^sub>1,r\<^sub>1\<rangle>, \<langle>l\<^sub>2,a\<^sub>2,r\<^sub>2\<rangle>) \<Rightarrow>"}\\
+\qquad@{text"\<^latex>\<open>{\\sf if}\<close> a\<^sub>1 \<le> a\<^sub>2 \<^latex>\<open>{\\sf then}\<close> \<langle>merge l\<^sub>1 r\<^sub>1, a\<^sub>1, t\<^sub>2\<rangle>"}\\
+\qquad@{text"\<^latex>\<open>{\\sf else}\<close> \<langle>t\<^sub>1, a\<^sub>2, merge l\<^sub>2 r\<^sub>2\<rangle>"}
+\bigskip\pause
+
+\bad{Challenge:} how to maintaining some kind of balance
+\end{frame}
+%-------------------------------------------------------------
+*}
+text_raw {*
+%-------------------------------------------------------------
+\section{Leftist Heap}
+%-------------------------------------------------------------
+\thyfile{Data\char`_Structures/Leftist\char`_Heap}{}
+%-------------------------------------------------------------
+\begin{frame}{Leftist tree informally}
+The \high{rank} of a tree is the depth of the rightmost leaf.
+\pause\bigskip
+
+In a \high{leftist tree}, the rank of every left child is $\ge$
+the rank of its right sibling
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Implementation type}
+\isakeyword{datatype}\\
+\quad \high{@{text "'a lheap = Leaf | Node nat ('a tree) 'a ('a tree)"}}
+\medskip\pause
+
+Abbreviations \high{@{const Leaf}} and \high{@{term "\<langle>h, l, a, r\<rangle>"}} as usual
+\bigskip\pause
+
+Abstraction function:\\
+@{const mset_tree} \<open>::\<close> @{typ "'a lheap \<Rightarrow> 'a multiset"}\\\smallskip\pause
+\high{@{thm mset_tree.simps(1)}}\\
+\high{@{thm[break,margin=50] (dummy_pats) mset_tree.simps(2)}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Heap}
+@{const heap} \<open>::\<close> @{typ "'a lheap \<Rightarrow> bool"}\\\smallskip\pause
+\high{@{thm heap.simps(1)}}\\
+\high{@{text"heap \<langle>_, l, a, r\<rangle> = (heap l \<and> heap r \<and>"}}\\
+\quad\high{@{text"(\<forall>x \<in># mset_tree l + mset_tree r. a \<le> x))"}}
+
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Leftist tree}
+@{const rank} \<open>::\<close> @{typ "'a lheap \<Rightarrow> nat"}\\\smallskip\pause
+\high{@{thm rank.simps(1)}}\\
+\high{@{thm (dummy_pats) rank.simps(2)}}
+\bigskip\pause
+
+Node @{term "Node n l a r"}: \<open>n = \<close> rank of node
+\bigskip\pause
+
+@{const ltree} \<open>::\<close> @{typ "'a lheap \<Rightarrow> bool"}\\\smallskip\pause
+\high{@{thm ltree.simps(1)}}\\
+\high{@{thm[break,margin=60] (dummy_pats) ltree.simps(2)}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Leftist heap invariant}
+\begin{center}
+\high{@{prop"invar h = (heap h \<and> ltree h)"}}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Why leftist tree?}
+\pause
+
+\textbf{Lemma} \high{@{thm pow2_rank_size1}}
+\bigskip\pause
+
+\textbf{Lemma} Execution time of @{term "merge t\<^sub>1 t\<^sub>2"} is bounded by
+\mbox{@{term"rank t\<^sub>1 + rank t\<^sub>2"}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{@{const merge}}
+Principle: descend on the right
+\smallskip\pause
+
+\high{@{thm[break] (sub) merge.simps(1,2)}}\\\pause
+\high{@{thm[break] (sub,dummy_pats) merge.simps(3)}}
+\bigskip\pause
+
+@{const node} \<open>::\<close> @{typ "'a lheap \<Rightarrow> 'a \<Rightarrow> 'a lheap \<Rightarrow> 'a lheap"}\\\smallskip\pause
+\high{@{thm[break,margin=65] (sub) node_def}}\\
+where @{thm (dummy_pats) rk.simps(2)}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{@{const merge}}
+@{thm[break] (sub,dummy_pats) merge.simps(3)}
+\bigskip\pause
+
+\high{\onslide<2->{Function @{const merge} terminates because}
+\onslide<3>{\revealbeamer{@{term"rank t\<^sub>1 + rank t\<^sub>2"}}}
+\onslide<2->{decreases with every recursive call.}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Functional correctness proofs}{including preservation of \<open>invar\<close>}
+Straightforward
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Logarithmic complexity}
+Complexity measures @{const t_merge}, @{const t_insert} @{const t_del_min}:\\
+ count calls of @{const merge}.
+\pause\bigskip
+
+\textbf{Lemma} @{thm t_merge_rank}
+\pause
+
+\textbf{Lemma} @{thm[break,margin=40] t_merge_log}
+\pause\bigskip
+
+\textbf{Lemma}\\ @{thm t_insert_log}
+\pause\bigskip
+
+\textbf{Lemma}\\ @{thm t_del_min_log}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}
+\begin{center}
+\bad{Can we avoid the rank info in each node?}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+
+\begin{isabellebody}%
+*}
+(*<*)
+end
+(*>*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/RBT_Slides.thy Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,234 @@
+(*<*)
+theory RBT_Slides
+imports
+ FDSlatex
+ "../../Public/Thys/RBT_Set"
+begin
+(*
+fun Bpl :: "'a rbt \<Rightarrow> nat set" where
+"Bpl Leaf = {0}" |
+"Bpl (Node c l _ r) =
+ (let Bs = Bpl l \<union> Bpl r
+ in if c = Black then {n+1 |n. n \<in> Bs} else Bs)"
+
+lemma invh_Bpl: "invh t \<Longrightarrow> Bpl t = {bheight t}"
+by(induction t) auto
+*)
+
+(*
+lemma combine_simp2: "combine t \<langle>\<rangle> = t"
+by (cases t) auto
+
+lemmas combine_simps = combine.simps(1) combine_simp2 combine.simps(3-)
+*)
+notation size1 ("|_|\<^sub>1")
+notation height ("h'(_')" 100)
+notation bheight ("bh'(_')" 100)
+(*>*)
+text_raw{*
+\end{isabellebody}%
+\section{Red-Black Trees}
+%-------------------------------------------------------------
+\thyfile{Thys/RBT\char`_Set}{}
+%-------------------------------------------------------------
+\begin{frame}{Relationship to 2-3-4 trees}
+Idea:
+\begin{tabular}[t]{l}
+ encode 2-3-4 trees as binary trees;\\\pause
+ use color to express grouping
+\end{tabular}
+\pause\bigskip
+
+\begin{tabular}{rcl}
+@{const Leaf} &$\approx$& @{const Leaf}\\\pause
+@{text"\<langle>t\<^sub>1,a,t\<^sub>2\<rangle>"} &$\approx$& @{text"\<langle>t\<^sub>1,a,t\<^sub>2\<rangle>"}\\\pause
+@{text"\<langle>t\<^sub>1,a,t\<^sub>2,b,t\<^sub>3\<rangle>"} &$\approx$&\pause \<open>\<langle>\<close>\bad{\<open>\<langle>t\<^sub>1,a,t\<^sub>2\<rangle>\<close>}\<open>,b,t\<^sub>3\<rangle>\<close>
+ \pause\<open>\<langle>t\<^sub>1,a,\<close>\bad{\<open>\<langle>t\<^sub>2,b,t\<^sub>3\<rangle>\<close>}\<open>\<rangle>\<close>\\\pause
+@{text"\<langle>t\<^sub>1,a,t\<^sub>2,b,t\<^sub>3,c,t\<^sub>4\<rangle>"} &$\approx$& \pause\<open>\<langle>\<close>\bad{\<open>\<langle>t\<^sub>1,a,t\<^sub>2\<rangle>\<close>}\<open>,b,\<close>\bad{\<open>\<langle>t\<^sub>3,c,t\<^sub>4\<rangle>\<close>}\<open>\<rangle>\<close>\\
+\end{tabular}
+\bigskip\pause
+
+\begin{center}
+\bad{Red} means ``I am part of a bigger node''
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Invariants}
+\begin{itemize}
+\pause
+\item The root is \pause\invisible<handout>{Black.}
+\pause
+\item Every @{const Leaf} is considered Black.
+\pause
+\item If a node is \bad{Red},
+ \pause\invisible<handout>{its children are Black.}
+\pause
+\item All paths from a node to a leaf have the same number of
+ \pause\invisible<handout>{Black nodes.}
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Red-black trees}
+@{datatype [display] color}
+\pause
+\isakeyword{datatype}\\
+\quad \high{@{text "'a rbt = Leaf | Node color ('a tree) 'a ('a tree)"}}
+\bigskip\pause
+
+Abbreviations:
+\smallskip
+\begin{center}
+\begin{tabular}{rcl}
+\high{@{const Leaf}} &\<open>\<equiv>\<close>& @{text Leaf}\\\pause
+\high{@{term "\<langle>c, l, a, r\<rangle>"}} &\<open>\<equiv>\<close>& @{term[source] "Node c l a r"}\medskip
+\\\pause
+\high{@{term "R l a r"}} &\<open>\<equiv>\<close>& @{term[source] "Node Red l a r"}\\\pause
+\high{@{term "B l a r"}} &\<open>\<equiv>\<close>& @{term[source] "Node Black l a r"}
+\end{tabular}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Color}
+\pause
+
+@{const color} \<open>::\<close> @{typ "'a rbt \<Rightarrow> color"}\\\smallskip\pause
+\high{@{thm color.simps(1)}}\\
+\high{@{thm (dummy_pats) color.simps(2)}}
+\bigskip\pause
+
+@{const paint} \<open>::\<close> @{typ "color \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt"}\\\smallskip\pause
+\high{@{thm paint.simps(1)}}\\
+\high{@{thm (dummy_pats) paint.simps(2)}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Invariants}
+\pause
+
+@{const rbt} \<open>::\<close> @{typ "'a rbt \<Rightarrow> bool"}\\\smallskip\pause
+\high{@{thm rbt_def}}
+\bigskip\pause
+
+@{const invc} \<open>::\<close> @{typ "'a rbt \<Rightarrow> bool"}\\\smallskip\pause
+\high{@{thm [break,margin=50] (dummy_pats) invc.simps}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Invariants}
+
+@{const invh} \<open>::\<close> @{typ "'a rbt \<Rightarrow> bool"}\\\smallskip\pause
+\high{@{thm [break,margin=55] (dummy_pats) invh.simps}}
+\bigskip\pause
+
+@{const bheight} \<open>::\<close> @{typ "'a rbt \<Rightarrow> nat"}\\\smallskip\pause
+\high{@{thm [break,margin=55] (dummy_pats) bheight.simps}}
+\end{frame}
+%-------------------------------------------------------------
+%\begin{frame}{Exercise}
+%\begin{center}
+%\bad{Is @ {const invh} what we want?}
+%\end{center}
+%\pause
+%
+%Define a function \ \high{@ {const Bpl} \<open>::\<close> @ {typ "'a rbt \<Rightarrow> nat set"}}\\
+%such that @ {term "Bpl t"} (``black path lengths'')
+%is the set of all \<open>n\<close> such that there is a path
+%from the root of \<open>t\<close> to a leaf that contains exactly \<open>n\<close> black nodes.
+%\pause\medskip
+%
+%Prove \ \high{@ {thm invh_Bpl}}
+%\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Logarithmic height}
+\begin{lemma}
+@{thm rbt_height_le}
+\end{lemma}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Insertion}
+
+@{const insert} \<open>::\<close> @{typ "'a \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt"}\\\smallskip\pause
+\high{@{thm insert_def}}
+\bigskip\pause
+
+@{const ins} \<open>::\<close> @{typ "'a \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt"}\\\smallskip\pause
+\high{@{thm ins.simps(1)}}\pause\vspace{-1ex}
+\high{@{thm[display] (dummy_pats) ins.simps(2)}}\pause\vspace{-1ex}
+\high{@{thm[display] (dummy_pats) ins.simps(3)}}
+
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Adjusting colors}
+@{const baliL}, @{const baliR} \<open>::\<close> @{typ "'a rbt \<Rightarrow> 'a \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt"}
+\begin{itemize}
+\pause
+\item Combine arguments \<open>l\<close> \<open>a\<close> \<open>r\<close> into tree, ideally \<open>\<langle>l, a, r\<rangle>\<close>
+\pause
+\item Treat invariant violation \bad{Red-Red} in \<open>l\<close>/\<open>r\<close>\\\pause
+\high{@{term (sub) "baliL (R (R t1 a1 t2) a2 t3) a3 t4"}\\
+\quad \<open>=\<close> @{term (sub) "R (B t1 a1 t2) a2 (B t3 a3 t4)"}}
+\pause
+
+\high{@{term (sub) "baliL (R t1 a1 (R t2 a2 t3)) a3 t4"}\\
+\quad \<open>=\<close> @{term (sub) "R (B t1 a1 t2) a2 (B t3 a3 t4)"}}
+\pause
+\item Principle: replace \bad{Red-Red} by \high{Red-Black}
+\pause
+\item Final equation: \\\pause
+\high{@{prop (sub) "baliL l a r = B l a r"}}
+\pause
+\item Symmetric: @{const baliR}
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+%\begin{frame}{Correctness via sorted lists}
+%\pause
+%\begin{lemma}
+%\high{@ {thm inorder_baliL}}\\
+%\high{@ {thm inorder_baliR}}
+%\end{lemma}
+%\pause
+%\begin{lemma}
+%\high{@ {thm[break,margin=60] inorder_ins}}
+%\end{lemma}
+%\pause
+%\begin{corollary}
+%\high{@ {thm[break,margin=60] inorder_insert}}
+%\end{corollary}
+%\bigskip\pause
+%
+%Proofs easy!
+%\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Preservation of invariant}
+\begin{theorem}
+\high{@{thm rbt_insert}}
+\end{theorem}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Deletion}
+
+\high{@{thm delete_def}}
+\bigskip\pause
+
+\high{@{thm [break,margin=55](dummy_pats) del.simps}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Deletion}
+Tricky functions: \high{@{const baldL}, @{const baldR}, @{const combine}}
+\bigskip\pause
+
+Tricky invariant preservation lemmas, eg
+
+\high{@{thm [display,margin=55] del_invc_invh}}
+\pause
+
+\begin{theorem}
+\high{@{thm rbt_delete}}
+\end{theorem}
+\end{frame}
+%-------------------------------------------------------------
+
+\begin{isabellebody}%
+*}
+(*<*)
+end
+(*>*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/RBT_Slides.thy~ Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,234 @@
+(*<*)
+theory RBT_Slides
+imports
+ FDSlatex
+ "../Material/Thys/RBT_Set"
+begin
+(*
+fun Bpl :: "'a rbt \<Rightarrow> nat set" where
+"Bpl Leaf = {0}" |
+"Bpl (Node c l _ r) =
+ (let Bs = Bpl l \<union> Bpl r
+ in if c = Black then {n+1 |n. n \<in> Bs} else Bs)"
+
+lemma invh_Bpl: "invh t \<Longrightarrow> Bpl t = {bheight t}"
+by(induction t) auto
+*)
+
+(*
+lemma combine_simp2: "combine t \<langle>\<rangle> = t"
+by (cases t) auto
+
+lemmas combine_simps = combine.simps(1) combine_simp2 combine.simps(3-)
+*)
+notation size1 ("|_|\<^sub>1")
+notation height ("h'(_')" 100)
+notation bheight ("bh'(_')" 100)
+(*>*)
+text_raw{*
+\end{isabellebody}%
+\section{Red-Black Trees}
+%-------------------------------------------------------------
+\thyfile{Thys/RBT\char`_Set}{}
+%-------------------------------------------------------------
+\begin{frame}{Relationship to 2-3-4 trees}
+Idea:
+\begin{tabular}[t]{l}
+ encode 2-3-4 trees as binary trees;\\\pause
+ use color to express grouping
+\end{tabular}
+\pause\bigskip
+
+\begin{tabular}{rcl}
+@{const Leaf} &$\approx$& @{const Leaf}\\\pause
+@{text"\<langle>t\<^sub>1,a,t\<^sub>2\<rangle>"} &$\approx$& @{text"\<langle>t\<^sub>1,a,t\<^sub>2\<rangle>"}\\\pause
+@{text"\<langle>t\<^sub>1,a,t\<^sub>2,b,t\<^sub>3\<rangle>"} &$\approx$&\pause \<open>\<langle>\<close>\bad{\<open>\<langle>t\<^sub>1,a,t\<^sub>2\<rangle>\<close>}\<open>,b,t\<^sub>3\<rangle>\<close>
+ \pause\<open>\<langle>t\<^sub>1,a,\<close>\bad{\<open>\<langle>t\<^sub>2,b,t\<^sub>3\<rangle>\<close>}\<open>\<rangle>\<close>\\\pause
+@{text"\<langle>t\<^sub>1,a,t\<^sub>2,b,t\<^sub>3,c,t\<^sub>4\<rangle>"} &$\approx$& \pause\<open>\<langle>\<close>\bad{\<open>\<langle>t\<^sub>1,a,t\<^sub>2\<rangle>\<close>}\<open>,b,\<close>\bad{\<open>\<langle>t\<^sub>3,c,t\<^sub>4\<rangle>\<close>}\<open>\<rangle>\<close>\\
+\end{tabular}
+\bigskip\pause
+
+\begin{center}
+\bad{Red} means ``I am part of a bigger node''
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Invariants}
+\begin{itemize}
+\pause
+\item The root is \pause\invisible<handout>{Black.}
+\pause
+\item Every @{const Leaf} is considered Black.
+\pause
+\item If a node is \bad{Red},
+ \pause\invisible<handout>{its children are Black.}
+\pause
+\item All paths from a node to a leaf have the same number of
+ \pause\invisible<handout>{Black nodes.}
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Red-black trees}
+@{datatype [display] color}
+\pause
+\isakeyword{datatype}\\
+\quad \high{@{text "'a rbt = Leaf | Node color ('a tree) 'a ('a tree)"}}
+\bigskip\pause
+
+Abbreviations:
+\smallskip
+\begin{center}
+\begin{tabular}{rcl}
+\high{@{const Leaf}} &\<open>\<equiv>\<close>& @{text Leaf}\\\pause
+\high{@{term "\<langle>c, l, a, r\<rangle>"}} &\<open>\<equiv>\<close>& @{term[source] "Node c l a r"}\medskip
+\\\pause
+\high{@{term "R l a r"}} &\<open>\<equiv>\<close>& @{term[source] "Node Red l a r"}\\\pause
+\high{@{term "B l a r"}} &\<open>\<equiv>\<close>& @{term[source] "Node Black l a r"}
+\end{tabular}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Color}
+\pause
+
+@{const color} \<open>::\<close> @{typ "'a rbt \<Rightarrow> color"}\\\smallskip\pause
+\high{@{thm color.simps(1)}}\\
+\high{@{thm (dummy_pats) color.simps(2)}}
+\bigskip\pause
+
+@{const paint} \<open>::\<close> @{typ "color \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt"}\\\smallskip\pause
+\high{@{thm paint.simps(1)}}\\
+\high{@{thm (dummy_pats) paint.simps(2)}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Invariants}
+\pause
+
+@{const rbt} \<open>::\<close> @{typ "'a rbt \<Rightarrow> bool"}\\\smallskip\pause
+\high{@{thm rbt_def}}
+\bigskip\pause
+
+@{const invc} \<open>::\<close> @{typ "'a rbt \<Rightarrow> bool"}\\\smallskip\pause
+\high{@{thm [break,margin=50] (dummy_pats) invc.simps}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Invariants}
+
+@{const invh} \<open>::\<close> @{typ "'a rbt \<Rightarrow> bool"}\\\smallskip\pause
+\high{@{thm [break,margin=55] (dummy_pats) invh.simps}}
+\bigskip\pause
+
+@{const bheight} \<open>::\<close> @{typ "'a rbt \<Rightarrow> nat"}\\\smallskip\pause
+\high{@{thm [break,margin=55] (dummy_pats) bheight.simps}}
+\end{frame}
+%-------------------------------------------------------------
+%\begin{frame}{Exercise}
+%\begin{center}
+%\bad{Is @ {const invh} what we want?}
+%\end{center}
+%\pause
+%
+%Define a function \ \high{@ {const Bpl} \<open>::\<close> @ {typ "'a rbt \<Rightarrow> nat set"}}\\
+%such that @ {term "Bpl t"} (``black path lengths'')
+%is the set of all \<open>n\<close> such that there is a path
+%from the root of \<open>t\<close> to a leaf that contains exactly \<open>n\<close> black nodes.
+%\pause\medskip
+%
+%Prove \ \high{@ {thm invh_Bpl}}
+%\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Logarithmic height}
+\begin{lemma}
+@{thm rbt_height_le}
+\end{lemma}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Insertion}
+
+@{const insert} \<open>::\<close> @{typ "'a \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt"}\\\smallskip\pause
+\high{@{thm insert_def}}
+\bigskip\pause
+
+@{const ins} \<open>::\<close> @{typ "'a \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt"}\\\smallskip\pause
+\high{@{thm ins.simps(1)}}\pause\vspace{-1ex}
+\high{@{thm[display] (dummy_pats) ins.simps(2)}}\pause\vspace{-1ex}
+\high{@{thm[display] (dummy_pats) ins.simps(3)}}
+
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Adjusting colors}
+@{const baliL}, @{const baliR} \<open>::\<close> @{typ "'a rbt \<Rightarrow> 'a \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt"}
+\begin{itemize}
+\pause
+\item Combine arguments \<open>l\<close> \<open>a\<close> \<open>r\<close> into tree, ideally \<open>\<langle>l, a, r\<rangle>\<close>
+\pause
+\item Treat invariant violation \bad{Red-Red} in \<open>l\<close>/\<open>r\<close>\\\pause
+\high{@{term (sub) "baliL (R (R t1 a1 t2) a2 t3) a3 t4"}\\
+\quad \<open>=\<close> @{term (sub) "R (B t1 a1 t2) a2 (B t3 a3 t4)"}}
+\pause
+
+\high{@{term (sub) "baliL (R t1 a1 (R t2 a2 t3)) a3 t4"}\\
+\quad \<open>=\<close> @{term (sub) "R (B t1 a1 t2) a2 (B t3 a3 t4)"}}
+\pause
+\item Principle: replace \bad{Red-Red} by \high{Red-Black}
+\pause
+\item Final equation: \\\pause
+\high{@{prop (sub) "baliL l a r = B l a r"}}
+\pause
+\item Symmetric: @{const baliR}
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+%\begin{frame}{Correctness via sorted lists}
+%\pause
+%\begin{lemma}
+%\high{@ {thm inorder_baliL}}\\
+%\high{@ {thm inorder_baliR}}
+%\end{lemma}
+%\pause
+%\begin{lemma}
+%\high{@ {thm[break,margin=60] inorder_ins}}
+%\end{lemma}
+%\pause
+%\begin{corollary}
+%\high{@ {thm[break,margin=60] inorder_insert}}
+%\end{corollary}
+%\bigskip\pause
+%
+%Proofs easy!
+%\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Preservation of invariant}
+\begin{theorem}
+\high{@{thm rbt_insert}}
+\end{theorem}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Deletion}
+
+\high{@{thm delete_def}}
+\bigskip\pause
+
+\high{@{thm [break,margin=55](dummy_pats) del.simps}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Deletion}
+Tricky functions: \high{@{const baldL}, @{const baldR}, @{const combine}}
+\bigskip\pause
+
+Tricky invariant preservation lemmas, eg
+
+\high{@{thm [display,margin=55] del_invc_invh}}
+\pause
+
+\begin{theorem}
+\high{@{thm rbt_delete}}
+\end{theorem}
+\end{frame}
+%-------------------------------------------------------------
+
+\begin{isabellebody}%
+*}
+(*<*)
+end
+(*>*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/Search_Trees.thy Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,237 @@
+(*<*)
+theory Search_Trees
+imports
+ Tree_Notation
+ "~~/src/HOL/Data_Structures/Tree_Set"
+begin
+
+lemma bst_simps2:
+ "bst \<langle>l, a, r\<rangle> \<longleftrightarrow> bst l \<and> bst r \<and> (\<forall>x\<in>set_tree l. x < a) \<and> (\<forall>x\<in>set_tree r. a < dummyid x)"
+by(simp add: dummyid_def)
+
+lemmas bst_simps = bst.simps(1) bst_simps2
+
+lemma delete_simps2: "delete x (Node l a r) =
+ (case cmp x a of
+ LT \<Rightarrow> Node (delete x l) a r |
+ GT \<Rightarrow> Node l a (delete x r) |
+ EQ \<Rightarrow> if r = Leaf then l else let (a',r') = del_min r in Node l a' (dummyid r'))"
+by(simp add: dummyid_def)
+
+(* TODO: SOS *)
+(*>*)
+text_raw{*
+\end{isabellebody}%
+%-------------------------------------------------------------
+\begin{frame}
+\begin{center}
+Most of the material focuses on\\ \high{BST}s = \high{binary search trees}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{BSTs represent sets}
+\pause
+Any tree represents a set:
+\medskip
+
+@{const_typ set_tree}\\\smallskip\pause
+\high{@{thm tree.set(1)}}\\
+\high{@{thm tree.set(2)[of l x r]}}
+\bigskip\pause
+
+A BST represents a set that can be searched in time $O(h(t))$
+\pause
+\begin{center}
+Function @{const set_tree} is called an \high{\emph{abstraction function}}
+because it maps the implementation \\ to the abstract mathematical object
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{@{const bst}}
+@{const_typ bst}\pause
+\high{@{thm[display,margin=40] bst_simps}}
+\vfill\pause
+Type @{typ 'a} must be in class @{class linorder} (@{typ[source]"'a :: linorder"})
+where @{class linorder} are \emph{linear orders} (also called \emph{total orders}).
+\bigskip\pause
+
+Note: @{typ nat}, @{typ int} and @{typ real} are in class @{class linorder}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Set interface}
+An implementation of sets of elements of type \ \high{@{typ 'a}} \ must provide
+\begin{itemize}
+\pause
+\item An implementation type \ \high{@{typ 's}}
+\pause
+\item \high{\<open>empty ::\<close> @{typ 's}}
+\pause
+\item \high{\<open>insert ::\<close> @{typ"'a \<Rightarrow> 's \<Rightarrow> 's"}}
+\pause
+\item \high{\<open>delete ::\<close> @{typ"'a \<Rightarrow> 's \<Rightarrow> 's"}}
+\pause
+\item \high{\<open>isin ::\<close> \quad @{typ"'s \<Rightarrow> 'a \<Rightarrow> bool"}}
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Map interface}
+Instead of a set, a search tree can also implement a \high{map} from @{typ 'a} to @{typ 'b}:
+\begin{itemize}
+\pause
+\item An implementation type \ \high{@{typ 'm}}
+\pause
+\item \high{\<open>empty ::\<close> @{typ 'm}}
+\pause
+\item \high{\<open>update ::\<close> @{typ"'a \<Rightarrow> 'b \<Rightarrow> 'm \<Rightarrow> 'm"}}
+\pause
+\item \high{\<open>delete ::\<close> @{typ"'a \<Rightarrow> 'm \<Rightarrow> 'm"}}
+\pause
+\item \high{\<open>lookup ::\<close> @{typ"'m \<Rightarrow> 'a \<Rightarrow> 'b option"}}
+\end{itemize}
+\pause\bigskip
+Sets are a special case of maps
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Comparison of elements}
+\pause
+\begin{center}
+\high{We assume that the element type @{typ 'a} is a linear order}
+\end{center}
+\pause
+Instead of using \<open><\<close> and \<open>\<le>\<close> directly:
+\high{@{datatype[display] cmp_val}}
+\pause
+\high{@{thm[break,margin=50] cmp_def}}
+\end{frame}
+%-------------------------------------------------------------
+\section{Unbalanced BST}
+%-------------------------------------------------------------
+\begin{frame}{Implementation}
+Implementation type: @{typ "'a tree"}
+\bigskip\pause
+
+\high{@{text"empty = Leaf"}}
+\pause
+
+\high{@{thm[display]insert.simps}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\mbox{Implementation}}
+\high{@{thm[display]isin.simps}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\mbox{Implementation}}
+\high{@{thm delete.simps(1)}}\\\pause
+\high{@{thm[break,margin=60]delete_simps2}}
+\pause
+\high{@{thm[display,margin=60]del_min.simps}}
+\end{frame}
+%-------------------------------------------------------------
+\subsection{Correctness}
+%-------------------------------------------------------------
+\begin{frame}{Why is this implementation correct?}
+\pause
+\begin{tabular}{@ {}lcccc @ {}}
+Because & \<open>empty\<close> & \<open>insert\<close> & \<open>delete\<close> & \<open>isin\<close>\\
+simulate & \<open>{}\<close> & \<open>\<union> {.}\<close> & \<open>- {.}\<close> & \<open>\<in>\<close>
+\end{tabular}
+\begin{center}
+\begin{tabular}{l}
+\pause
+\high{@{text "set_tree empty = {}"}}\\\pause
+\high{@{prop[source] "set_tree (insert x t) = set_tree t \<union> {x}"}}\\\pause
+\high{@{prop "set_tree (delete x t) = set_tree t - {x}"}}\\\pause
+\high{@{prop "isin t x = (x \<in> set_tree t)"}}\\
+\end{tabular}
+\end{center}
+\pause
+
+Under the assumption \pause \high{@{prop"bst t"}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Also: @{const bst} must be invariant}
+\begin{center}
+\begin{tabular}{l}
+\high{@{text "bst empty"}}\\
+\high{@{prop "bst t \<Longrightarrow> bst (insert x t)"}}\\
+\high{@{prop "bst t \<Longrightarrow> bst (delete x t)"}}
+\end{tabular}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\subsection{Correctness Proof Method Based on Sorted Lists}
+%-------------------------------------------------------------
+\begin{frame}{}
+@{const_typ sorted}\\\smallskip\pause
+\high{@{thm sorted.simps(1)}}\\\pause
+\high{@{thm sorted.simps(2)}}\\\pause
+\high{@{thm sorted.simps(3)}}\pause
+\begin{center}
+No duplicates!
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+*}
+text_raw (in Set_by_Ordered) {*
+%-------------------------------------------------------------
+\begin{frame}{Structural invariant}
+The proof method works not just for unbalanced trees.
+\pause
+
+We assume that there is some structural invariant on the search tree:
+\high{@{text [display] "inv : 's \<Rightarrow> bool"}}
+e.g. some balance criterion.
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of @{const insert}}
+\high{@{thm[display,margin=60] insert}}
+\pause
+where
+\high{@{const_typ[display] ins_list}}
+inserts an element into a sorted list.
+\bigskip\pause
+\begin{center}
+\high{Also covers preservation of @{const bst}}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of @{const delete}}
+\high{@{thm[display,margin=60] delete}}
+\pause
+where
+\high{@{const_typ[display] del_list}}
+deletes an element from a sorted list.
+\bigskip\pause
+\begin{center}
+\high{Also covers preservation of @{const bst}}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of @{const isin}}
+\high{@{thm[display,margin=60] isin}}
+\pause
+where
+\high{@{const_typ[display] elems}}
+converts a list into a set.
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness via sorted lists}
+\pause
+\begin{center}
+\high{Correctness proofs of all search trees\\
+ covered in this course\\
+can be automated.}
+\pause\medskip
+
+\bad{Except for the structural invariants.}
+\pause\medskip
+
+\bad{Therefore we concentrate on the latter from now on.}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{isabellebody}%
+*}
+(*<*)
+end
+(*>*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/Sorting_Slides.thy Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,205 @@
+(*<*)
+theory Sorting_Slides
+imports
+ FDSlatex
+ "../../Public/Thys/Sorting"
+begin
+
+definition dummyid :: "'a \<Rightarrow> 'a" where
+"dummyid x = x"
+
+notation (latex output) dummyid ("_")
+
+lemma sorted_Cons: "sorted (x # xs) = ((\<forall>y\<in>set xs. x \<le> dummyid y) & sorted xs)"
+by(simp add: dummyid_def)
+
+lemmas sorted_simps = sorted.simps(1) sorted_Cons
+
+lemma merge_simp2: "merge xs [] = xs"
+by (cases xs) auto
+(*>*)
+text_raw{*
+\end{isabellebody}%
+%-------------------------------------------------------------
+\section{Correctness}
+%-------------------------------------------------------------
+\begin{frame}
+\high{@{const sorted} \<open>::\<close> @{typ[source]"('a::linorder) list \<Rightarrow> bool"}}
+\pause
+\high{@{thm[display] sorted_simps}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of sorting}
+Specification of \<open>sort ::\<close> @{typ[source] "('a::linorder) list \<Rightarrow> 'a list"}:
+\begin{center}
+\high{@{prop "sorted (sort xs)"}}
+\end{center}
+\pause
+Is that it? \pause How about
+\begin{center}
+@{prop "set (sort xs) = set xs"}
+\end{center}
+\pause
+Better: every \<open>x\<close> occurs as often in @{term "sort xs"} as in \<open>xs\<close>.
+\pause\medskip
+
+More succinctly:
+\begin{center}
+\high{@{prop "mset (sort xs) = mset xs"}}
+\end{center}
+where \high{@{const_typ mset}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{What are multisets?}
+\pause
+\begin{center}
+\high{Sets with (possibly) repeated elements}
+\end{center}
+\pause
+Some operations:\medskip
+
+\begin{tabular}{rcl}
+\high{@{const Mempty}} &\<open>::\<close>& @{typ "'a multiset"}\\
+\high{@{const add_mset}} &\<open>::\<close>& @{typ "'a \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset"}\\
+\high{\<open>+\<close>} &\<open>::\<close>& @{typ "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset"}\\\pause
+\high{@{const mset}} &\<open>::\<close>& @{typ "'a list \<Rightarrow> 'a multiset"}\\
+\high{@{const set_mset}} &\<open>::\<close>& @{typ "'a multiset \<Rightarrow> 'a set"}\\
+\end{tabular}
+\end{frame}
+%-------------------------------------------------------------
+\section{Insertion Sort}
+%-------------------------------------------------------------
+\thyfile{Thys/Sorting}{Insertion sort correctness}
+%-------------------------------------------------------------
+\section{Time}
+%-------------------------------------------------------------
+\begin{frame}{Principle: Count function calls}
+\pause
+\begin{tabular}{@ {}ll@ {}}
+For every function & \ \<open>f :: \<tau>\<^sub>1 \<Rightarrow> ... \<Rightarrow> \<tau>\<^sub>n \<Rightarrow> \<tau>\<close>\\
+define a \concept{timing function} & \high{\<open>t_f :: \<tau>\<^sub>1 \<Rightarrow> ... \<Rightarrow> \<tau>\<^sub>n \<Rightarrow> nat\<close>}:
+\end{tabular}
+\smallskip\pause
+Translation of defining equations:
+\[
+\frac{e \leadsto \high{e'}}{f\,p_1 \dots p_n = e ~\leadsto~ \high{t\_f\,p_1 \dots p_n = e' + 1}}
+\]
+\pause
+
+Translation of expressions:
+\[ \frac{s_1 \leadsto \high{t_1} \quad \dots \quad s_k \leadsto \high{t_k}}
+ {g\,s_1 \dots s_k \leadsto \high{t_1 + \dots + t_k + t\_g\, s_1 \dots s_k}}
+\]
+\pause
+\begin{itemize}
+\item
+Variable $\leadsto$ 0, Constant $\leadsto$ 0
+\pause
+\item
+Constructor calls and primitive operations on @{typ bool} and numbers cost 1
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Example}
+\<open>app [] ys = ys\<close>\\
+\pause$\leadsto$\\
+\high{\<open>t_app [] ys = 0 + 1\<close>}\\
+\pause\bigskip
+
+\<open>app (x#xs) ys = x # app xs ys\<close>\\
+\pause$\leadsto$\\
+\high{\<open>t_app (x#xs) ys = 0 + (0 + 0 + t_app xs ys) + 1 + 1\<close>}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{A compact formulation of $e \leadsto t$}
+\pause
+\begin{center}
+\begin{tabular}{l}
+\high{\<open>t\<close>} \ is the sum of all \high{\<open>t_g s\<^sub>1 ... s\<^sub>k\<close>}\\
+such that \<open>g s\<^sub>1 ... s\<^sub>n\<close> is a subterm of \<open>e\<close>
+\end{tabular}
+\end{center}
+\pause\smallskip
+
+If $g$ is
+\begin{itemize}
+\item a constructor or
+\item a predefined function on @{typ bool} or numbers
+\end{itemize}
+then \<open>t_g ... = 1\<close>.
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\emph{if} and \emph{case}}
+\pause
+\begin{center}
+So far we model a call-by-value semantics
+\end{center}
+\pause
+Conditionals and case expressions are evaluated \bad{lazily}.
+\smallskip\pause
+
+Translation:
+\[ \frac{b \leadsto \high{t} \quad s_1 \leadsto \high{t_1} \quad s_2 \leadsto \high{t_2}}
+ {\textit{if } b \textit{ then } s_1 \textit{ else } s_2 \leadsto \high{t + (\textit{if } b \textit{ then } t_1 \textit{ else } t_2)}}
+\]
+\pause
+Similarly for \emph{case}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{$O(.)$ is enough}
+\pause
+$\Longrightarrow$ Reduce all additive constants to 1
+\pause
+\begin{example}
+\high{\<open>t_app (x#xs) ys = t_app xs ys + 1\<close>}
+\end{example}
+%\bigskip\pause
+
+%$\Longrightarrow$ May start with 0 instead of 1 in base case
+%\pause
+%\begin{example}
+%\high{\<open>t_app [] ys = 0\<close>}
+%\end{example}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Discussion}
+\begin{itemize}
+\pause
+\item
+The definition of \<open>t_f\<close> from \<open>f\<close> can be automated.
+\pause
+\item
+The correctness of \<open>t_f\<close> could be proved w.r.t.\\ a semantics
+that counts computation steps.
+\pause
+\item
+Precise complexity bounds (as opposed to $O(.)$)
+would require a formal model of (at least) the compiler and the hardware.
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\thyfile{Thys/Sorting}{Insertion sort complexity}
+%-------------------------------------------------------------
+\section{Merge Sort}
+%-------------------------------------------------------------
+\begin{frame}
+@{const_typ merge}\smallskip\\
+\pause
+\high{@{thm merge.simps(1)}}\\
+\high{@{thm merge_simp2}}\\
+\pause
+\high{@{thm[break,margin=60] merge.simps(3)}}
+\bigskip\pause
+
+@{const_typ msort}\smallskip\\
+\pause
+\high{@{thm[break,margin=60] msort.simps}}
+\end{frame}
+%-------------------------------------------------------------
+\thyfile{Thys/Sorting}{Merge sort}
+%-------------------------------------------------------------
+\begin{isabellebody}%
+*}
+(*<*)
+end
+(*>*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/Sorting_Slides.thy~ Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,205 @@
+(*<*)
+theory Sorting_Slides
+imports
+ FDSlatex
+ "../Material/Thys/Sorting"
+begin
+
+definition dummyid :: "'a \<Rightarrow> 'a" where
+"dummyid x = x"
+
+notation (latex output) dummyid ("_")
+
+lemma sorted_Cons: "sorted (x # xs) = ((\<forall>y\<in>set xs. x \<le> dummyid y) & sorted xs)"
+by(simp add: dummyid_def)
+
+lemmas sorted_simps = sorted.simps(1) sorted_Cons
+
+lemma merge_simp2: "merge xs [] = xs"
+by (cases xs) auto
+(*>*)
+text_raw{*
+\end{isabellebody}%
+%-------------------------------------------------------------
+\section{Correctness}
+%-------------------------------------------------------------
+\begin{frame}
+\high{@{const sorted} \<open>::\<close> @{typ[source]"('a::linorder) list \<Rightarrow> bool"}}
+\pause
+\high{@{thm[display] sorted_simps}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of sorting}
+Specification of \<open>sort ::\<close> @{typ[source] "('a::linorder) list \<Rightarrow> 'a list"}:
+\begin{center}
+\high{@{prop "sorted (sort xs)"}}
+\end{center}
+\pause
+Is that it? \pause How about
+\begin{center}
+@{prop "set (sort xs) = set xs"}
+\end{center}
+\pause
+Better: every \<open>x\<close> occurs as often in @{term "sort xs"} as in \<open>xs\<close>.
+\pause\medskip
+
+More succinctly:
+\begin{center}
+\high{@{prop "mset (sort xs) = mset xs"}}
+\end{center}
+where \high{@{const_typ mset}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{What are multisets?}
+\pause
+\begin{center}
+\high{Sets with (possibly) repeated elements}
+\end{center}
+\pause
+Some operations:\medskip
+
+\begin{tabular}{rcl}
+\high{@{const Mempty}} &\<open>::\<close>& @{typ "'a multiset"}\\
+\high{@{const add_mset}} &\<open>::\<close>& @{typ "'a \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset"}\\
+\high{\<open>+\<close>} &\<open>::\<close>& @{typ "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset"}\\\pause
+\high{@{const mset}} &\<open>::\<close>& @{typ "'a list \<Rightarrow> 'a multiset"}\\
+\high{@{const set_mset}} &\<open>::\<close>& @{typ "'a multiset \<Rightarrow> 'a set"}\\
+\end{tabular}
+\end{frame}
+%-------------------------------------------------------------
+\section{Insertion Sort}
+%-------------------------------------------------------------
+\thyfile{Thys/Sorting}{Insertion sort correctness}
+%-------------------------------------------------------------
+\section{Time}
+%-------------------------------------------------------------
+\begin{frame}{Principle: Count function calls}
+\pause
+\begin{tabular}{@ {}ll@ {}}
+For every function & \ \<open>f :: \<tau>\<^sub>1 \<Rightarrow> ... \<Rightarrow> \<tau>\<^sub>n \<Rightarrow> \<tau>\<close>\\
+define a \concept{timing function} & \high{\<open>t_f :: \<tau>\<^sub>1 \<Rightarrow> ... \<Rightarrow> \<tau>\<^sub>n \<Rightarrow> nat\<close>}:
+\end{tabular}
+\smallskip\pause
+Translation of defining equations:
+\[
+\frac{e \leadsto \high{e'}}{f\,p_1 \dots p_n = e ~\leadsto~ \high{t\_f\,p_1 \dots p_n = e' + 1}}
+\]
+\pause
+
+Translation of expressions:
+\[ \frac{s_1 \leadsto \high{t_1} \quad \dots \quad s_k \leadsto \high{t_k}}
+ {g\,s_1 \dots s_k \leadsto \high{t_1 + \dots + t_k + t\_g\, s_1 \dots s_k}}
+\]
+\pause
+\begin{itemize}
+\item
+Variable $\leadsto$ 0, Constant $\leadsto$ 0
+\pause
+\item
+Constructor calls and primitive operations on @{typ bool} and numbers cost 1
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Example}
+\<open>app [] ys = ys\<close>\\
+\pause$\leadsto$\\
+\high{\<open>t_app [] ys = 0 + 1\<close>}\\
+\pause\bigskip
+
+\<open>app (x#xs) ys = x # app xs ys\<close>\\
+\pause$\leadsto$\\
+\high{\<open>t_app (x#xs) ys = 0 + (0 + 0 + t_app xs ys) + 1 + 1\<close>}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{A compact formulation of $e \leadsto t$}
+\pause
+\begin{center}
+\begin{tabular}{l}
+\high{\<open>t\<close>} \ is the sum of all \high{\<open>t_g s\<^sub>1 ... s\<^sub>k\<close>}\\
+such that \<open>g s\<^sub>1 ... s\<^sub>n\<close> is a subterm of \<open>e\<close>
+\end{tabular}
+\end{center}
+\pause\smallskip
+
+If $g$ is
+\begin{itemize}
+\item a constructor or
+\item a predefined function on @{typ bool} or numbers
+\end{itemize}
+then \<open>t_g ... = 1\<close>.
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\emph{if} and \emph{case}}
+\pause
+\begin{center}
+So far we model a call-by-value semantics
+\end{center}
+\pause
+Conditionals and case expressions are evaluated \bad{lazily}.
+\smallskip\pause
+
+Translation:
+\[ \frac{b \leadsto \high{t} \quad s_1 \leadsto \high{t_1} \quad s_2 \leadsto \high{t_2}}
+ {\textit{if } b \textit{ then } s_1 \textit{ else } s_2 \leadsto \high{t + (\textit{if } b \textit{ then } t_1 \textit{ else } t_2)}}
+\]
+\pause
+Similarly for \emph{case}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{$O(.)$ is enough}
+\pause
+$\Longrightarrow$ Reduce all additive constants to 1
+\pause
+\begin{example}
+\high{\<open>t_app (x#xs) ys = t_app xs ys + 1\<close>}
+\end{example}
+%\bigskip\pause
+
+%$\Longrightarrow$ May start with 0 instead of 1 in base case
+%\pause
+%\begin{example}
+%\high{\<open>t_app [] ys = 0\<close>}
+%\end{example}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Discussion}
+\begin{itemize}
+\pause
+\item
+The definition of \<open>t_f\<close> from \<open>f\<close> can be automated.
+\pause
+\item
+The correctness of \<open>t_f\<close> could be proved w.r.t.\\ a semantics
+that counts computation steps.
+\pause
+\item
+Precise complexity bounds (as opposed to $O(.)$)
+would require a formal model of (at least) the compiler and the hardware.
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\thyfile{Thys/Sorting}{Insertion sort complexity}
+%-------------------------------------------------------------
+\section{Merge Sort}
+%-------------------------------------------------------------
+\begin{frame}
+@{const_typ merge}\smallskip\\
+\pause
+\high{@{thm merge.simps(1)}}\\
+\high{@{thm merge_simp2}}\\
+\pause
+\high{@{thm[break,margin=60] merge.simps(3)}}
+\bigskip\pause
+
+@{const_typ msort}\smallskip\\
+\pause
+\high{@{thm[break,margin=60] msort.simps}}
+\end{frame}
+%-------------------------------------------------------------
+\thyfile{Thys/Sorting}{Merge sort}
+%-------------------------------------------------------------
+\begin{isabellebody}%
+*}
+(*<*)
+end
+(*>*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/Tree23_Slides.thy Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,209 @@
+(*<*)
+theory Tree23_Slides
+imports
+ FDSlatex
+ "~~/src/HOL/Data_Structures/Tree23_Set"
+begin
+
+notation height ("h'(_')" 100)
+
+lemma del_simp4: "l \<noteq> Leaf \<Longrightarrow> r \<noteq> Leaf \<Longrightarrow> del x (Node2 l a r) =
+ (case cmp x a of
+ LT \<Rightarrow> node21 (del x l) a r |
+ GT \<Rightarrow> node22 l a (del x r) |
+ EQ \<Rightarrow> let (a',t) = del_min r in node22 l a' t)"
+by(cases l) auto
+(*>*)
+text_raw{*
+\end{isabellebody}%
+%-------------------------------------------------------------
+\section{2-3 Trees}
+%-------------------------------------------------------------
+\thyfile{HOL/Data\char`_Structures/Tree23\char`_Set}{}
+%-------------------------------------------------------------
+\begin{frame}{2-3 Trees}
+\high{@{datatype[display,margin=55] tree23}}
+\pause
+
+Abbreviations:
+\smallskip
+\begin{center}
+\begin{tabular}{rcl}
+\high{@{term "\<langle>l, a, r\<rangle>"}} &\<open>\<equiv>\<close>& @{text "Node2 l a r"}\\\pause
+\high{@{term "\<langle>l, a, m, b, r\<rangle>"}} &\<open>\<equiv>\<close>& @{text "Node3 l a m b r"}
+\end{tabular}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{@{const isin}}
+\high{@{thm[display,margin=25]isin.simps(3)}}
+\pause
+
+Assumes the usual ordering invariant
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Invariant @{const bal}}
+All leaves are at the same level:
+\pause
+\high{@{thm[display,margin=55]bal.simps(1)}}
+\pause
+\high{@{thm bal.simps(2)[of _ DUMMY]}}
+\pause
+\high{@{thm[display,margin=55]bal.simps(3)[of _ DUMMY _ DUMMY]}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Insertion}
+The idea:
+\begin{center}
+\begin{tabular}{rcl}
+\pause
+\<open>Leaf\<close> & $\leadsto$ & @{const Node2} \\
+\pause
+@{const Node2} & $\leadsto$ & @{const Node3} \\
+\pause
+@{const Node3} & $\leadsto$ & \bad{overflow}, pass 1 element back up \\
+\end{tabular}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Insertion}
+Two possible return values:
+\pause
+\begin{itemize}
+\item tree accommodates new element\\ without increasing height:
+ \high{@{term "T\<^sub>i t"}}
+\pause
+\item tree overflows: \bad{@{term "Up\<^sub>i l x r"}}
+\end{itemize}
+\pause
+@{datatype[display,margin=40] up\<^sub>i}
+\pause
+@{const_typ tree\<^sub>i}\\
+\pause\smallskip
+@{thm[break] tree\<^sub>i.simps}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Insertion}
+@{const_typ insert}\\\smallskip\pause
+\high{@{thm insert_def}}
+\bigskip\pause
+
+@{const_typ ins}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Insertion}
+\high{@{thm ins.simps(1)}}
+\smallskip\pause
+
+\high{@{thm (lhs,sub) ins.simps(2)} \<open>=\<close>}\\\pause
+\high{@{thm [break,margin=60] (rhs,sub) ins.simps(2)}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Insertion}
+\high{@{thm (lhs,sub) ins.simps(3)} \<open>=\<close>}\\\pause
+\high{@{thm[break,margin=52] (rhs,sub) ins.simps(3)}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Insertion preserves @{const bal}}
+\begin{lemma}
+\high{\alt<1-4>{@{prop"bal t \<Longrightarrow> bal (tree\<^sub>i(ins a t))"}}{@{thm bal_ins}}}
+\smallskip
+
+\onslide<3->
+where \<open>h ::\<close> @{typ "'a up\<^sub>i \<Rightarrow> nat"}\\
+\onslide<4->
+@{thm [break]height_up\<^sub>i.simps}
+\end{lemma}
+\onslide<2->
+\textbf{Proof} by induction on \<open>t\<close>
+\bigskip\onslide<6->
+\begin{corollary}
+\high{@{thm bal_insert}}
+\end{corollary}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Deletion}
+The idea:
+\begin{center}
+\begin{tabular}{rcl}
+\pause
+@{const Node3} & $\leadsto$ & @{const Node2} \\
+\pause
+@{const Node2} & $\leadsto$ & \bad{underflow}, height -1
+\end{tabular}
+\end{center}
+\pause
+Underflow: merge with siblings on the way up
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Deletion}
+Two possible return values:
+\begin{itemize}
+\pause
+\item height unchanged: \high{@{term "T\<^sub>d t"}}
+\pause
+\item height decreased by 1: \bad{@{term "Up\<^sub>d t"}}
+\end{itemize}
+\pause
+@{datatype[display,margin=60] up\<^sub>d}
+\pause
+@{thm[display] tree\<^sub>d.simps}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Deletion}
+@{const_typ delete}\\\smallskip\pause
+\high{@{thm delete_def}}
+\bigskip\pause
+
+@{const_typ del}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Deletion}
+@{thm del.simps(1)}
+\smallskip\pause
+
+@{thm[break,margin=60] del.simps(2)}
+\smallskip\pause
+
+@{thm (lhs) del.simps(3)} \<open>= ...\<close>
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}
+\high{@{thm[display,margin=60] (concl) del_simp4}}
+\pause
+@{thm[display,margin=60] (sub) node21.simps}
+\pause
+@{thm[display,margin=60] (sub) node22.simps}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Deletion preserves @{const bal}}
+\begin{lemma}
+\high{@{thm bal_tree\<^sub>d_del}}
+\end{lemma}
+
+\begin{corollary}
+\high{@{thm bal_delete}}
+\end{corollary}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Beyond 2-3 trees}
+\pause
+\high{\isakeyword{datatype} @{text "'a tree234 ="}\\
+\quad \<open>Leaf | Node2 ... | Node3 ... | Node4 ...\<close>}
+\pause
+
+\begin{center}
+Like 2-3 tress, but with many more cases
+\end{center}
+\pause
+The general case:
+\begin{center}
+ B-trees and (a,b) trees
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{isabellebody}%
+*}
+(*<*)
+end
+(*>*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/Tree_Notation.thy Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,18 @@
+(*<*)
+theory Tree_Notation
+imports
+ FDSlatex
+ "~~/src/HOL/Library/Tree"
+begin
+
+notation height ("h'(_')" 100)
+
+notation min_height ("mh'(_')" 100)
+
+abbreviation (latex output) tree_size :: "'a tree \<Rightarrow> nat" ("|_|") where
+"tree_size \<equiv> size"
+
+notation size1 ("|_|\<^sub>1")
+
+end
+(*>*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/Tree_Slides.thy Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,227 @@
+(*<*)
+theory Tree_Slides
+imports
+ Tree_Notation
+ "../../Public/Thys/Tree_Additions"
+begin
+
+lemma size_Node: "size(Node l a r) = size l + size r + 1"
+by simp
+
+lemmas size_simps = tree.size(3) size_Node
+
+(* TODO: SOS *)
+(*>*)
+text_raw{*
+\end{isabellebody}%
+\section{Binary Trees}
+%-------------------------------------------------------------
+\ttfilepage{HOL/Library/Tree.thy\\ Thys/Tree\char`\_Additions.thy}{}
+%-------------------------------------------------------------
+\begin{frame}{Binary trees}
+\begin{center}
+\isakeyword{datatype} \high{@{text "'a tree = Leaf | Node ('a tree) 'a ('a tree)"}}
+\end{center}
+\pause
+
+Abbreviations:
+\begin{tabular}{rcl}
+\high{@{const Leaf}} &\<open>\<equiv>\<close>& @{text Leaf}\\\pause
+\high{@{term "\<langle>l, a, r\<rangle>"}} &\<open>\<equiv>\<close>& @{text "Node l a r"}
+\end{tabular}
+%\pause\medskip
+
+%Terminology:
+%\begin{tabular}{rcl}
+%\high{leaf} & means & @ {text Leaf}\\\pause
+%\high{node} & means & @ {const Node}\\\pause
+%\high{leaf node} & means & @ {term"Node Leaf a Leaf"}
+%\end{tabular}
+\pause\vfill
+\begin{center}
+Most of the time: \high{tree = binary tree}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\section{Basic Functions}
+%-------------------------------------------------------------
+\begin{frame}{Tree traversal}
+@{const_typ inorder}\\\smallskip\pause
+\high{@{thm inorder.simps(1)}}\\
+\high{@{thm inorder.simps(2)}}
+\bigskip\pause
+
+@{const_typ preorder}\\\smallskip\pause
+\high{@{thm preorder.simps(1)}}\\
+\high{@{thm preorder.simps(2)}}
+\bigskip\pause
+
+@{const_typ postorder}\\\smallskip\pause
+\high{@{thm postorder.simps(1)}}\\
+\high{@{thm postorder.simps(2)}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Size}
+
+@{const size} \<open>::\<close> @{typ "'a tree \<Rightarrow> nat"}\\\smallskip\pause
+\high{@{thm size_simps(1)}}\\
+\high{@{thm size_simps(2)[of _ DUMMY]}}
+\bigskip\pause
+
+@{const_typ size1}\\\pause
+\high{@{thm size1_def}}\\
+\bigskip\pause
+
+\<open>\<Longrightarrow>\<close>\\
+@{thm size1_simps(1)}\\
+@{thm size1_simps(2)}
+\bigskip\pause
+
+\textbf{Lemma} The number of leaves in \<open>t\<close> is @{term"size1 t"}.
+\bigskip\pause
+
+\bad{Warning: \<open>|.|\<close> and \<open>|.|\<^sub>1\<close> only on slides}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Height}
+
+@{const height} \<open>::\<close> @{typ "'a tree \<Rightarrow> nat"}\\\smallskip\pause
+\high{@{thm height_tree.simps(1)}}\\
+\high{@{thm height_tree.simps(2)[of l DUMMY r]}}
+\pause
+\begin{center}
+\bad{Warning: \<open>h(.)\<close> only on slides}
+\end{center}
+\pause
+
+\textbf{Lemma} @{thm height_le_size_tree}
+\bigskip\pause
+
+\textbf{Lemma} @{thm size1_height}
+\end{frame}
+%-------------------------------------------------------------
+%\thyfile{Thys/Tree\char`\_Additions}{@{thm size1_height}}
+%-------------------------------------------------------------
+\begin{frame}{Minimal height}
+@{const_typ min_height}\\\smallskip\pause
+\high{@{thm min_height.simps(1)}}\\
+\high{@{thm min_height.simps(2)[of l DUMMY r]}}
+\pause
+\begin{center}
+\bad{Warning: \<open>mh(.)\<close> only on slides}
+\end{center}
+\pause
+
+\textbf{Lemma} @{thm min_height_le_height}
+\bigskip\pause
+
+\textbf{Lemma} @{thm min_height_size1}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Internal path length}
+\pause
+\begin{center}
+The sum of the lengths of the (simpl) paths\\ from the root to any other node
+\end{center}
+\pause
+
+@{const_typ ipl}\\\smallskip\pause
+\high{@{thm ipl.simps(1)}}\\
+\high{@{thm ipl.simps(2)[of l DUMMY r]}}
+\vfill\pause
+\begin{center}
+Why relevant?
+\pause\bigskip
+
+Upper bound?
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\section{Complete and Balanced Trees}
+%-------------------------------------------------------------
+\begin{frame}{Complete tree}
+@{const_typ complete}\\\pause
+\high{@{thm[break,margin=50] (dummy_pats) complete.simps}}
+\bigskip\pause
+
+\textbf{Lemma} @{thm complete_iff_height}
+\bigskip\pause
+
+\textbf{Lemma} @{thm size1_if_complete}
+%Proof by induction; induction-free proof: 2^mh(t) <= size1 t <= 2^h(t)
+\bigskip\pause
+
+\textbf{Lemma} @{thm complete_if_size1_height}
+\pause
+
+\textbf{Lemma} @{thm complete_if_size1_min_height}
+\bigskip\pause
+
+\textbf{Corollary} @{thm size1_height_if_incomplete}
+\pause
+
+\textbf{Corollary} @{thm min_height_size1_if_incomplete}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Complete tree: @{const ipl}}
+\textbf{Lemma} A complete tree of height \<open>h\<close> has internal path length
+\high{@{term "(h - 2)*2^h + 2"}}.
+\bigskip\pause
+
+In a search tree, finding the node labelled \<open>x\<close>
+takes as many steps as the path from the root to \<open>x\<close> is long.
+\pause
+
+Thus the average time to find an element that is in the tree is
+\high{@{term "ipl t"} \<open>/\<close> @{term "size(t::'a tree)"}}.
+\bigskip\pause
+
+\textbf{Lemma} Let \<open>t\<close> be a complete search tree of height \<open>h\<close>.
+\pause
+The average time to find a random element that is in the tree
+is asymptotically \high{\<open>h - 2\<close>} (as \<open>h\<close> approaches \<open>\<infinity>\<close>):\pause
+\begin{center}
+\high{@{term "ipl t"} \<open>/\<close> @{term "size(t::'a tree)"} \<open>\<sim>\<close> @{term"h-2::int"}}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Complete tree: @{const ipl}}
+A problem: \high{@{term "(h - 2)*2^h + 2"}} is only correct if interpreted
+over type \high{@{typ int}}, not \bad{@{typ nat}}.
+\bigskip\pause
+
+Correct version:
+
+\textbf{Lemma} @{thm (prem 1) ipl_if_complete_int} \<open>\<Longrightarrow>\<close>
+\mbox{@{thm (concl) ipl_if_complete_int}}
+\vfill\pause
+
+We do not cover the Isabelle formalization of limits.
+
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Balanced tree}
+@{const_typ balanced}\\\smallskip\pause
+\high{@{thm balanced_def}}
+\bigskip\pause
+
+Balanced trees have optimal height:\\
+\textbf{Lemma} @{thm[mode=IfThen] balanced_optimal}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Warning}
+\begin{itemize}
+\item The terms \bad{\emph{complete}} and \bad{\emph{balanced}}\\
+ are not defined uniquely in the literature.
+\bigskip\pause
+
+\item For example,\\
+Knuth calls \bad{\emph{complete}} what we call \bad{\emph{balanced}}.
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{isabellebody}%
+*}
+(*<*)
+end
+(*>*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/Tree_Slides.thy~ Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,227 @@
+(*<*)
+theory Tree_Slides
+imports
+ Tree_Notation
+ "../Material/Thys/Tree_Additions"
+begin
+
+lemma size_Node: "size(Node l a r) = size l + size r + 1"
+by simp
+
+lemmas size_simps = tree.size(3) size_Node
+
+(* TODO: SOS *)
+(*>*)
+text_raw{*
+\end{isabellebody}%
+\section{Binary Trees}
+%-------------------------------------------------------------
+\ttfilepage{HOL/Library/Tree.thy\\ Thys/Tree\char`\_Additions.thy}{}
+%-------------------------------------------------------------
+\begin{frame}{Binary trees}
+\begin{center}
+\isakeyword{datatype} \high{@{text "'a tree = Leaf | Node ('a tree) 'a ('a tree)"}}
+\end{center}
+\pause
+
+Abbreviations:
+\begin{tabular}{rcl}
+\high{@{const Leaf}} &\<open>\<equiv>\<close>& @{text Leaf}\\\pause
+\high{@{term "\<langle>l, a, r\<rangle>"}} &\<open>\<equiv>\<close>& @{text "Node l a r"}
+\end{tabular}
+%\pause\medskip
+
+%Terminology:
+%\begin{tabular}{rcl}
+%\high{leaf} & means & @ {text Leaf}\\\pause
+%\high{node} & means & @ {const Node}\\\pause
+%\high{leaf node} & means & @ {term"Node Leaf a Leaf"}
+%\end{tabular}
+\pause\vfill
+\begin{center}
+Most of the time: \high{tree = binary tree}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\section{Basic Functions}
+%-------------------------------------------------------------
+\begin{frame}{Tree traversal}
+@{const_typ inorder}\\\smallskip\pause
+\high{@{thm inorder.simps(1)}}\\
+\high{@{thm inorder.simps(2)}}
+\bigskip\pause
+
+@{const_typ preorder}\\\smallskip\pause
+\high{@{thm preorder.simps(1)}}\\
+\high{@{thm preorder.simps(2)}}
+\bigskip\pause
+
+@{const_typ postorder}\\\smallskip\pause
+\high{@{thm postorder.simps(1)}}\\
+\high{@{thm postorder.simps(2)}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Size}
+
+@{const size} \<open>::\<close> @{typ "'a tree \<Rightarrow> nat"}\\\smallskip\pause
+\high{@{thm size_simps(1)}}\\
+\high{@{thm size_simps(2)[of _ DUMMY]}}
+\bigskip\pause
+
+@{const_typ size1}\\\pause
+\high{@{thm size1_def}}\\
+\bigskip\pause
+
+\<open>\<Longrightarrow>\<close>\\
+@{thm size1_simps(1)}\\
+@{thm size1_simps(2)}
+\bigskip\pause
+
+\textbf{Lemma} The number of leaves in \<open>t\<close> is @{term"size1 t"}.
+\bigskip\pause
+
+\bad{Warning: \<open>|.|\<close> and \<open>|.|\<^sub>1\<close> only on slides}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Height}
+
+@{const height} \<open>::\<close> @{typ "'a tree \<Rightarrow> nat"}\\\smallskip\pause
+\high{@{thm height_tree.simps(1)}}\\
+\high{@{thm height_tree.simps(2)[of l DUMMY r]}}
+\pause
+\begin{center}
+\bad{Warning: \<open>h(.)\<close> only on slides}
+\end{center}
+\pause
+
+\textbf{Lemma} @{thm height_le_size_tree}
+\bigskip\pause
+
+\textbf{Lemma} @{thm size1_height}
+\end{frame}
+%-------------------------------------------------------------
+%\thyfile{Thys/Tree\char`\_Additions}{@{thm size1_height}}
+%-------------------------------------------------------------
+\begin{frame}{Minimal height}
+@{const_typ min_height}\\\smallskip\pause
+\high{@{thm min_height.simps(1)}}\\
+\high{@{thm min_height.simps(2)[of l DUMMY r]}}
+\pause
+\begin{center}
+\bad{Warning: \<open>mh(.)\<close> only on slides}
+\end{center}
+\pause
+
+\textbf{Lemma} @{thm min_height_le_height}
+\bigskip\pause
+
+\textbf{Lemma} @{thm min_height_size1}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Internal path length}
+\pause
+\begin{center}
+The sum of the lengths of the (simpl) paths\\ from the root to any other node
+\end{center}
+\pause
+
+@{const_typ ipl}\\\smallskip\pause
+\high{@{thm ipl.simps(1)}}\\
+\high{@{thm ipl.simps(2)[of l DUMMY r]}}
+\vfill\pause
+\begin{center}
+Why relevant?
+\pause\bigskip
+
+Upper bound?
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\section{Complete and Balanced Trees}
+%-------------------------------------------------------------
+\begin{frame}{Complete tree}
+@{const_typ complete}\\\pause
+\high{@{thm[break,margin=50] (dummy_pats) complete.simps}}
+\bigskip\pause
+
+\textbf{Lemma} @{thm complete_iff_height}
+\bigskip\pause
+
+\textbf{Lemma} @{thm size1_if_complete}
+%Proof by induction; induction-free proof: 2^mh(t) <= size1 t <= 2^h(t)
+\bigskip\pause
+
+\textbf{Lemma} @{thm complete_if_size1_height}
+\pause
+
+\textbf{Lemma} @{thm complete_if_size1_min_height}
+\bigskip\pause
+
+\textbf{Corollary} @{thm size1_height_if_incomplete}
+\pause
+
+\textbf{Corollary} @{thm min_height_size1_if_incomplete}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Complete tree: @{const ipl}}
+\textbf{Lemma} A complete tree of height \<open>h\<close> has internal path length
+\high{@{term "(h - 2)*2^h + 2"}}.
+\bigskip\pause
+
+In a search tree, finding the node labelled \<open>x\<close>
+takes as many steps as the path from the root to \<open>x\<close> is long.
+\pause
+
+Thus the average time to find an element that is in the tree is
+\high{@{term "ipl t"} \<open>/\<close> @{term "size(t::'a tree)"}}.
+\bigskip\pause
+
+\textbf{Lemma} Let \<open>t\<close> be a complete search tree of height \<open>h\<close>.
+\pause
+The average time to find a random element that is in the tree
+is asymptotically \high{\<open>h - 2\<close>} (as \<open>h\<close> approaches \<open>\<infinity>\<close>):\pause
+\begin{center}
+\high{@{term "ipl t"} \<open>/\<close> @{term "size(t::'a tree)"} \<open>\<sim>\<close> @{term"h-2::int"}}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Complete tree: @{const ipl}}
+A problem: \high{@{term "(h - 2)*2^h + 2"}} is only correct if interpreted
+over type \high{@{typ int}}, not \bad{@{typ nat}}.
+\bigskip\pause
+
+Correct version:
+
+\textbf{Lemma} @{thm (prem 1) ipl_if_complete_int} \<open>\<Longrightarrow>\<close>
+\mbox{@{thm (concl) ipl_if_complete_int}}
+\vfill\pause
+
+We do not cover the Isabelle formalization of limits.
+
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Balanced tree}
+@{const_typ balanced}\\\smallskip\pause
+\high{@{thm balanced_def}}
+\bigskip\pause
+
+Balanced trees have optimal height:\\
+\textbf{Lemma} @{thm[mode=IfThen] balanced_optimal}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Warning}
+\begin{itemize}
+\item The terms \bad{\emph{complete}} and \bad{\emph{balanced}}\\
+ are not defined uniquely in the literature.
+\bigskip\pause
+
+\item For example,\\
+Knuth calls \bad{\emph{complete}} what we call \bad{\emph{balanced}}.
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{isabellebody}%
+*}
+(*<*)
+end
+(*>*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/Trie_Slides.thy Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,223 @@
+(*<*)
+theory Trie_Slides
+imports
+ FDSlatex
+ "../../Public/Thys/Trie2"
+begin
+(*>*)
+text_raw{*
+\end{isabellebody}%
+%-------------------------------------------------------------
+\section{Tries and Patricia Tries}
+%-------------------------------------------------------------
+\ttfilepage{Thys/Trie1\\Thys/Trie2}{}
+%-------------------------------------------------------------
+\begin{frame}{Trie}
+\framesubtitle{[Fredkin, CACM 1960]}
+Name: \emph{reTRIEval}
+\pause
+\begin{itemize}
+\item
+Tries are search trees indexed by lists.
+\pause
+\item
+Each node maps the next element in the list\\ to a subtrie.
+\pause
+\item
+For simplicity we consider only binary tries:
+\high{@{datatype[display] trie}}
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Patricia trie}
+\framesubtitle{[Morrison, JACM 1968]}
+Sequences of edges without branching are compressed
+\pause
+\high{@{datatype[display,margin=60] ptrie}}
+\pause
+
+Name: PATRICIA = \emph{Practical Algorithm To Retrieve Information Coded in Alphanumeric}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{From trie to Patricia trie}
+\framesubtitle{and beyond}
+\pause
+\begin{center}
+\Large An exercise in data refinement
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\subsection{Data Refinement Basics}
+%-------------------------------------------------------------
+\begin{frame}{Data refinement}
+How to implement an abstract type \<open>A\<close> by a concrete (``refined'') type \<open>C\<close>
+with \high{invariant \<open>inv\<^sub>C :: C \<Rightarrow> bool\<close>}:
+\begin{itemize}
+\pause
+\item
+Define an \high{abstraction function \<open>abs\<^sub>C :: C \<Rightarrow> A\<close>}
+\pause
+\item Implement each interface function \<open>f\<^sub>A\<close> on \<open>A\<close>\\ by some \<open>f\<^sub>C\<close> on \<open>C\<close>
+\pause
+\item Prove that each \<open>f\<^sub>C\<close> implements \<open>f\<^sub>A\<close>:\\
+\high{\<open>inv\<^sub>C t \<Longrightarrow> abs\<^sub>C (f\<^sub>C t) = f\<^sub>A (abs\<^sub>C t)\<close>}\\\pause
+and preserves the invariant:\\
+\high{\<open>inv\<^sub>C t \<Longrightarrow> inv\<^sub>C (f\<^sub>C t)\<close>}
+\end{itemize}
+\pause
+The \<open>f\<close>\,s may take more parameters\\\pause
+ and may return values not of type \<open>A\<close>/\<open>C\<close>.
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Data refinement}
+\begin{center}
+\high{Multiple refinement steps can help\\
+to reduce the complexity of correctness proofs}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\subsection{Trie}
+%-------------------------------------------------------------
+\begin{frame}{Trie}
+\high{@{datatype[display] trie}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Trie @{const isin}}
+@{thm isin.simps(1)}\pause
+@{thm[display,margin=60] isin.simps(2)}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Trie @{const insert}}
+@{thm insert.simps(1)}\\\pause
+@{thm insert.simps(2)}\\\pause
+@{thm[display,margin=60] insert.simps(3)}\pause
+@{thm[break,margin=60] insert.simps(4)}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of \<open>trie\<close> w.r.t. \<open>set\<close>}
+\begin{center}
+\high{@{thm isin_insert}}
+\end{center}
+\pause
+Note: @{const isin} doubles as abstraction function
+\end{frame}
+%-------------------------------------------------------------
+\subsection{Patricia Trie}
+%-------------------------------------------------------------
+\begin{frame}{Patricia trie}
+\high{@{datatype[display,margin=60] ptrie}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{@{const isinP}}
+@{thm isinP.simps(1)}\pause
+@{thm[display,margin=60] isinP.simps(2)}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Splitting lists}
+\high{@{prop "split xs ys = (zs,xs',ys')"}}\\\pause
+iff \<open>zs\<close> is the longest common prefix of \<open>xs\<close> and \<open>ys\<close>\\\pause
+and \<open>xs'\<close>/\<open>ys'\<close> is the remainder of \<open>xs\<close>/\<open>ys\<close>
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{@{const insertP}}
+\pause
+\small
+@{thm[break,margin=60] insertP.simps(2)}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{@{const_typ abs_ptrie}}
+\pause
+\high{@{thm abs_ptrie.simps(1)}}
+\pause
+\high{@{thm[display] abs_ptrie.simps(2)}}
+\pause
+@{const_typ prefix_trie}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of @{typ ptrie} w.r.t. @{typ trie}}
+\begin{center}
+\high{@{thm isinP}}
+\bigskip\pause
+
+\high{@{thm abs_ptrie_insertP}}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\subsection{Patricia Trie for Words}
+%-------------------------------------------------------------
+\begin{frame}{Patricia trie for words}
+\framesubtitle{[Okasaki \& Gill, 1998]}
+\high{@{datatype[display,margin=50] ptrieW}}
+\bigskip\pause
+
+@{const_typ to_bl}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{@{const isinW}}
+@{thm isinW.simps(1)}\pause
+@{thm[display,margin=60] isinW.simps(2)}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of @{typ ptrie} w.r.t. @{typ trie}}
+\begin{center}
+\high{@{thm[display,margin=60] isinW_correct}}
+\end{center}
+\pause
+Proof uses intermedidate version of Patricia trie for lists
+where nodes contain complete prefixes (@{const isinL}).
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Refinement hierarchy}
+\begin{center}
+\begin{tabular}{cl}
+@{typ trie}\\
+$\uparrow$ & @{const abs_ptrie}\\
+@{typ ptrie}\\
+$\uparrow$ & @{const abs_ptrieL}\\
+@{typ ptrieL}\\
+$\uparrow$ & @{const abs_ptrieW}\\
+@{typ ptrieW}\\[1ex]
+\multicolumn{2}{l}{\pause where @{typ ptrieL} \<open>=\<close> @{typ ptrie}}
+\end{tabular}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\mbox{@{const_typ inv_ptrieW}}}
+\pause
+\high{@{thm inv_ptrieW.simps(1)}}
+\pause
+\high{@{thm[display,margin=60] inv_ptrieW.simps(2)}}
+\pause
+@{thm[display,margin=50] inv_pref_def}
+\pause
+@{thm[break,margin=60] inv_mask_def}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\mbox{@{const_typ abs_ptrieW}}}
+\pause
+\high{@{thm abs_ptrieW.simps(1)}}
+\pause
+\high{@{thm[display] abs_ptrieW.simps(2)}}
+\pause
+@{thm[display,margin=60] abs_pmb_def}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\mbox{@{text insertW}}}
+\begin{center}
+\Large\bad{?}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Further space optimization?}
+\pause
+\textbf{datatype} \<open>ptrieW2 =\<close>\\
+\quad \<open>LeafW2\<close> ~$|$\\
+\quad \<open>SingW2 (32 word)\<close> ~$|$\\
+\quad \<open>NodeW2 (32 word) (32 word) (ptrieW2 \<times> ptrieW2)\<close>
+\end{frame}
+%-------------------------------------------------------------
+\begin{isabellebody}%
+*}
+(*<*)
+end
+(*>*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/Trie_Slides.thy~ Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,223 @@
+(*<*)
+theory Trie_Slides
+imports
+ FDSlatex
+ "../Material/Thys/Trie2"
+begin
+(*>*)
+text_raw{*
+\end{isabellebody}%
+%-------------------------------------------------------------
+\section{Tries and Patricia Tries}
+%-------------------------------------------------------------
+\ttfilepage{Thys/Trie1\\Thys/Trie2}{}
+%-------------------------------------------------------------
+\begin{frame}{Trie}
+\framesubtitle{[Fredkin, CACM 1960]}
+Name: \emph{reTRIEval}
+\pause
+\begin{itemize}
+\item
+Tries are search trees indexed by lists.
+\pause
+\item
+Each node maps the next element in the list\\ to a subtrie.
+\pause
+\item
+For simplicity we consider only binary tries:
+\high{@{datatype[display] trie}}
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Patricia trie}
+\framesubtitle{[Morrison, JACM 1968]}
+Sequences of edges without branching are compressed
+\pause
+\high{@{datatype[display,margin=60] ptrie}}
+\pause
+
+Name: PATRICIA = \emph{Practical Algorithm To Retrieve Information Coded in Alphanumeric}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{From trie to Patricia trie}
+\framesubtitle{and beyond}
+\pause
+\begin{center}
+\Large An exercise in data refinement
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\subsection{Data Refinement Basics}
+%-------------------------------------------------------------
+\begin{frame}{Data refinement}
+How to implement an abstract type \<open>A\<close> by a concrete (``refined'') type \<open>C\<close>
+with \high{invariant \<open>inv\<^sub>C :: C \<Rightarrow> bool\<close>}:
+\begin{itemize}
+\pause
+\item
+Define an \high{abstraction function \<open>abs\<^sub>C :: C \<Rightarrow> A\<close>}
+\pause
+\item Implement each interface function \<open>f\<^sub>A\<close> on \<open>A\<close>\\ by some \<open>f\<^sub>C\<close> on \<open>C\<close>
+\pause
+\item Prove that each \<open>f\<^sub>C\<close> implements \<open>f\<^sub>A\<close>:\\
+\high{\<open>inv\<^sub>C t \<Longrightarrow> abs\<^sub>C (f\<^sub>C t) = f\<^sub>A (abs\<^sub>C t)\<close>}\\\pause
+and preserves the invariant:\\
+\high{\<open>inv\<^sub>C t \<Longrightarrow> inv\<^sub>C (f\<^sub>C t)\<close>}
+\end{itemize}
+\pause
+The \<open>f\<close>\,s may take more parameters\\\pause
+ and may return values not of type \<open>A\<close>/\<open>C\<close>.
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Data refinement}
+\begin{center}
+\high{Multiple refinement steps can help\\
+to reduce the complexity of correctness proofs}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\subsection{Trie}
+%-------------------------------------------------------------
+\begin{frame}{Trie}
+\high{@{datatype[display] trie}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Trie @{const isin}}
+@{thm isin.simps(1)}\pause
+@{thm[display,margin=60] isin.simps(2)}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Trie @{const insert}}
+@{thm insert.simps(1)}\\\pause
+@{thm insert.simps(2)}\\\pause
+@{thm[display,margin=60] insert.simps(3)}\pause
+@{thm[break,margin=60] insert.simps(4)}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of \<open>trie\<close> w.r.t. \<open>set\<close>}
+\begin{center}
+\high{@{thm isin_insert}}
+\end{center}
+\pause
+Note: @{const isin} doubles as abstraction function
+\end{frame}
+%-------------------------------------------------------------
+\subsection{Patricia Trie}
+%-------------------------------------------------------------
+\begin{frame}{Patricia trie}
+\high{@{datatype[display,margin=60] ptrie}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{@{const isinP}}
+@{thm isinP.simps(1)}\pause
+@{thm[display,margin=60] isinP.simps(2)}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Splitting lists}
+\high{@{prop "split xs ys = (zs,xs',ys')"}}\\\pause
+iff \<open>zs\<close> is the longest common prefix of \<open>xs\<close> and \<open>ys\<close>\\\pause
+and \<open>xs'\<close>/\<open>ys'\<close> is the remainder of \<open>xs\<close>/\<open>ys\<close>
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{@{const insertP}}
+\pause
+\small
+@{thm[break,margin=60] insertP.simps(2)}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{@{const_typ abs_ptrie}}
+\pause
+\high{@{thm abs_ptrie.simps(1)}}
+\pause
+\high{@{thm[display] abs_ptrie.simps(2)}}
+\pause
+@{const_typ prefix_trie}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of @{typ ptrie} w.r.t. @{typ trie}}
+\begin{center}
+\high{@{thm isinP}}
+\bigskip\pause
+
+\high{@{thm abs_ptrie_insertP}}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\subsection{Patricia Trie for Words}
+%-------------------------------------------------------------
+\begin{frame}{Patricia trie for words}
+\framesubtitle{[Okasaki \& Gill, 1998]}
+\high{@{datatype[display,margin=50] ptrieW}}
+\bigskip\pause
+
+@{const_typ to_bl}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{@{const isinW}}
+@{thm isinW.simps(1)}\pause
+@{thm[display,margin=60] isinW.simps(2)}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of @{typ ptrie} w.r.t. @{typ trie}}
+\begin{center}
+\high{@{thm[display,margin=60] isinW_correct}}
+\end{center}
+\pause
+Proof uses intermedidate version of Patricia trie for lists
+where nodes contain complete prefixes (@{const isinL}).
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Refinement hierarchy}
+\begin{center}
+\begin{tabular}{cl}
+@{typ trie}\\
+$\uparrow$ & @{const abs_ptrie}\\
+@{typ ptrie}\\
+$\uparrow$ & @{const abs_ptrieL}\\
+@{typ ptrieL}\\
+$\uparrow$ & @{const abs_ptrieW}\\
+@{typ ptrieW}\\[1ex]
+\multicolumn{2}{l}{\pause where @{typ ptrieL} \<open>=\<close> @{typ ptrie}}
+\end{tabular}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\mbox{@{const_typ inv_ptrieW}}}
+\pause
+\high{@{thm inv_ptrieW.simps(1)}}
+\pause
+\high{@{thm[display,margin=60] inv_ptrieW.simps(2)}}
+\pause
+@{thm[display,margin=50] inv_pref_def}
+\pause
+@{thm[break,margin=60] inv_mask_def}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\mbox{@{const_typ abs_ptrieW}}}
+\pause
+\high{@{thm abs_ptrieW.simps(1)}}
+\pause
+\high{@{thm[display] abs_ptrieW.simps(2)}}
+\pause
+@{thm[display,margin=60] abs_pmb_def}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\mbox{@{text insertW}}}
+\begin{center}
+\Large\bad{?}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Further space optimization?}
+\pause
+\textbf{datatype} \<open>ptrieW2 =\<close>\\
+\quad \<open>LeafW2\<close> ~$|$\\
+\quad \<open>SingW2 (32 word)\<close> ~$|$\\
+\quad \<open>NodeW2 (32 word) (32 word) (ptrieW2 \<times> ptrieW2)\<close>
+\end{frame}
+%-------------------------------------------------------------
+\begin{isabellebody}%
+*}
+(*<*)
+end
+(*>*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/document/prelude.tex Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,104 @@
+%Beamer configuration
+
+% Include page numbers
+\setbeamertemplate{footline}{\hfill\insertframenumber\hspace*{2mm}\vspace*{2mm}}
+
+\usetheme{Pittsburgh}
+
+% switch off the fancy navigation symbols
+\setbeamertemplate{navigation symbols}{}
+
+\usefonttheme[onlymath]{serif}
+\usepackage{isabelle,isabellesym}
+\usepackage{tikz}
+\usetikzlibrary{matrix}
+
+\isabellestyle{it}
+
+% for uniform font size
+\renewcommand{\isastyleminor}{\rmfamily\itshape}
+\renewcommand{\isastyle}{\isastyleminor}
+%\renewcommand{\isastyletext}{\normalsize\sf}
+
+% typeset underscore
+\renewcommand{\isacharunderscore}{\_}%
+% normal _ has no bold version, use tt _
+% both for theory snippets and for \isakeyword uses
+\renewcommand{\isacharunderscorekeyword}{\texttt{\char`_}}%
+\let\isakeywordold=\isakeyword
+\renewcommand{\isakeyword}[1]{\isakeywordold{\def\_{{\texttt{\char`_}}}#1}}
+
+\renewcommand{\isacharbackquoteopen}{\isacharbackquote}
+\renewcommand{\isacharbackquoteclose}{\isacharbackquote}
+\renewcommand{\isasymnabla}{\raisebox{2pt}{\isamath{\normalbigtriangledown}}}
+
+\let\isacharprimeR=\isacharprime
+\newcommand{\isactrlisubprime}[1]{{$\sb{#1}\sp{\mskip3mu\prime}$}}
+
+\newcommand{\xsymbol}[1]{\texttt{\char`\\\char`<#1>}}
+
+\definecolor{darkblue}{rgb}{0.0,0.0,0.5}
+\newcommand<>{\high}[1]{{\color#2{blue}{#1}}}
+\newcommand<>{\bad}[1]{{\color#2{red}{#1}}}
+\newcommand{\gray}[1]{\textcolor{gray}{#1}}
+
+\newcommand{\concept}[1]{\high{\emph{#1}}}
+\newcommand{\important}{{\Large \textbf{!}}}
+
+\newcommand{\filepage}[2]
+{\title{\textcolor{darkblue}{#1}}
+\author{#2}\institute{}\date{}
+\maketitle}
+
+\newcommand{\ttfilepage}[2]
+{\filepage{\texttt{#1}}{#2}}
+
+\newcommand{\thyfile}[2]
+{\ttfilepage{#1.thy}{#2}}
+
+\newcommand{\demofile}[2]
+{\thyfile{#1\char`\_Demo}{#2}}
+
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+\setbeamersize{text margin left=0.5cm,text margin right=0.5cm}
+\renewcommand\partname{Chapter}
+
+\setbeamertemplate{part page}
+{
+ \begin{centering}
+ {\usebeamerfont{part name}\usebeamercolor[fg]{part name}\partname~\insertpartnumber}
+ \vskip1em\par
+ \begin{beamercolorbox}[sep=16pt,center]{part title}
+ \usebeamerfont{part title}\insertpart\par
+ \end{beamercolorbox}
+ \end{centering}
+}
+
+\newcommand{\mypartpage}[2]{
+\begin{frame}
+ \begin{centering}
+ {\usebeamerfont{part name}\usebeamercolor[fg]{part name}Part #1}
+ \vskip1em\par
+ \begin{beamercolorbox}[sep=16pt,center]{part title}
+ \usebeamerfont{part title}#2\par
+ \end{beamercolorbox}
+ \end{centering}
+\end{frame}}
+
+\AtBeginPart{\frame{\partpage}\frame{\setcounter{tocdepth}{1}\tableofcontents}}
+
+\AtBeginSection[]
+{
+\begin{frame}
+\setcounter{tocdepth}{1}
+\tableofcontents[currentsection]
+\end{frame}
+}
+
+\AtBeginSubsection[]
+{
+\begin{frame}
+\setcounter{tocdepth}{2}
+\tableofcontents[sectionstyle=show/hide,subsectionstyle=show/shaded/hide]
+\end{frame}
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/document/root.tex Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,42 @@
+\documentclass[14pt
+,handout%
+]{beamer}
+
+\input{prelude}
+
+\renewcommand{\isadigit}[1]{\ensuremath{#1}}
+
+\newcommand{\afpentry}[1]
+{\filepage{Archive of Formal Proofs}
+{\url{https://www.isa-afp.org/entries/#1.shtml}}}
+
+\begin{document}
+
+\title{Functional Data Structures}
+\subtitle{with Isabelle/HOL}
+\author{Tobias Nipkow}
+\institute[]{%
+ Fakult\"at f\"ur Informatik\\
+ Technische Universit\"at M\"unchen}
+\date{\the\year-\the\month-\the\day}
+\frame{\titlepage}
+\subtitle{}
+
+\mypartpage{II}{Functional Data Structures}
+\setcounter{part}{5}
+\part{Sorting}
+\input{Sorting_Slides}
+\part{Binary Trees}
+\input{Tree_Slides}
+\part{Search Trees}
+\input{Search_Trees}
+\input{Tree23_Slides}
+\input{RBT_Slides}
+\input{Trie_Slides}
+%\part{Priority Queues}
+%\input{Leftist_Heap_Slides}
+%\input{Braun_Tree_Slides}
+
+%Amortized:
+%\input{Skew_Heap_Slides}
+\end{document}
Binary file Slides/tex/document.pdf has changed
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/tex/document/Braun_Tree_Slides.tex Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,159 @@
+%
+\begin{isabellebody}%
+\setisabellecontext{Braun{\isacharunderscore}Tree{\isacharunderscore}Slides}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isatagtheory
+%
+\endisatagtheory
+{\isafoldtheory}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\end{isabellebody}%
+%-------------------------------------------------------------
+\section{Priority Queues Based on Braun Trees}
+%-------------------------------------------------------------
+\afpentry{Priority_Queue_Braun}
+%-------------------------------------------------------------
+\begin{frame}{What is a Braun tree?}
+\isa{braun\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ tree\ {\isasymRightarrow}\ bool}\\\smallskip\pause
+\high{\isa{braun\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ True\isasep\isanewline%
+braun\ {\isasymlangle}l{\isacharcomma}\ x{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\isanewline
+{\isacharparenleft}{\isacharbar}r{\isacharbar}\ {\isasymle}\ {\isacharbar}l{\isacharbar}\ {\isasymand}\ {\isacharbar}l{\isacharbar}\ {\isasymle}\ Suc\ {\isacharbar}r{\isacharbar}\ {\isasymand}\ braun\ l\ {\isasymand}\ braun\ r{\isacharparenright}}}\\
+\bigskip\pause
+
+\textbf{Lemma} \high{\isa{braun\ t\ {\isasymLongrightarrow}\ {\isadigit{2}}\isactrlbsup \isa{h{\isacharparenleft}t{\isacharparenright}}\isactrlesup \ {\isasymle}\ {\isadigit{2}}\ {\isacharasterisk}\ {\isacharbar}t{\isacharbar}\ {\isacharplus}\ {\isadigit{1}}}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Priority queue implementation}
+Implementation type: ordinary binary trees
+\bigskip\pause
+
+Invariants: \isa{heap} and \isa{braun}
+\bigskip\pause
+
+\begin{center}
+\bad{No \isa{merge}} \pause --- \high{\isa{insert} and \isa{del{\isacharunderscore}min} defined explicitly}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\isa{insert}}
+\isa{insert\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ tree\ {\isasymRightarrow}\ {\isacharprime}a\ tree}\\\smallskip\pause
+\high{\isa{insert\ a\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ {\isasymlangle}{\isasymlangle}{\isasymrangle}{\isacharcomma}\ a{\isacharcomma}\ {\isasymlangle}{\isasymrangle}{\isasymrangle}\isasep\isanewline%
+insert\ a\ {\isasymlangle}l{\isacharcomma}\ x{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf if}\ a\ {\isacharless}\ x\ {\sf then}\ {\isasymlangle}insert\ x\ r{\isacharcomma}\ a{\isacharcomma}\ l{\isasymrangle}\ {\sf else}\ {\isasymlangle}insert\ a\ r{\isacharcomma}\ x{\isacharcomma}\ l{\isasymrangle}{\isacharparenright}}}
+\bigskip\pause
+
+Correctness and preservation of invariant straightforward.
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\isa{del{\isacharunderscore}min}}
+\isa{del{\isacharunderscore}min\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ tree\ {\isasymRightarrow}\ {\isacharprime}a\ tree}\\\smallskip\pause
+\high{\isa{del{\isacharunderscore}min\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ {\isasymlangle}{\isasymrangle}\isasep\isanewline%
+del{\isacharunderscore}min\ {\isasymlangle}{\isasymlangle}{\isasymrangle}{\isacharcomma}\ x{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\ {\isasymlangle}{\isasymrangle}\isasep\isanewline%
+del{\isacharunderscore}min\ {\isasymlangle}l{\isacharcomma}\ x{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf let}\ {\isacharparenleft}y{\isacharcomma}\ l{\isacharprime}{\isacharparenright}\ {\isacharequal}\ del{\isacharunderscore}left\ l\ {\sf in}\ sift{\isacharunderscore}down\ r\ y\ l{\isacharprime}{\isacharparenright}}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\isa{sift{\isacharunderscore}down}}
+\high<1>{\isa{sift{\isacharunderscore}down\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ tree\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ tree\ {\isasymRightarrow}\ {\isacharprime}a\ tree}}\\\smallskip\pause
+\high<2>{\isa{sift{\isacharunderscore}down\ {\isasymlangle}{\isasymrangle}\ a\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ {\isasymlangle}{\isasymlangle}{\isasymrangle}{\isacharcomma}\ a{\isacharcomma}\ {\isasymlangle}{\isasymrangle}{\isasymrangle}}}\\\smallskip\pause
+\high<3>{\isa{sift{\isacharunderscore}down\ {\isasymlangle}{\isasymlangle}{\isasymrangle}{\isacharcomma}\ x{\isacharcomma}\ {\isasymlangle}{\isasymrangle}{\isasymrangle}\ a\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf if}\ a\ {\isasymle}\ x\ {\sf then}\ {\isasymlangle}{\isasymlangle}{\isasymlangle}{\isasymrangle}{\isacharcomma}\ x{\isacharcomma}\ {\isasymlangle}{\isasymrangle}{\isasymrangle}{\isacharcomma}\ a{\isacharcomma}\ {\isasymlangle}{\isasymrangle}{\isasymrangle}\isanewline
+\isaindent{{\isacharparenleft}}{\sf else}\ {\isasymlangle}{\isasymlangle}{\isasymlangle}{\isasymrangle}{\isacharcomma}\ a{\isacharcomma}\ {\isasymlangle}{\isasymrangle}{\isasymrangle}{\isacharcomma}\ x{\isacharcomma}\ {\isasymlangle}{\isasymrangle}{\isasymrangle}{\isacharparenright}}}\\\smallskip\pause
+\high{\isa{sift{\isacharunderscore}down\ {\isasymlangle}l\isactrlsub {\isadigit{1}}{\isacharcomma}\ x\isactrlsub {\isadigit{1}}{\isacharcomma}\ r\isactrlsub {\isadigit{1}}{\isasymrangle}\ a\ {\isasymlangle}l\isactrlsub {\isadigit{2}}{\isacharcomma}\ x\isactrlsub {\isadigit{2}}{\isacharcomma}\ r\isactrlsub {\isadigit{2}}{\isasymrangle}\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf if}\ a\ {\isasymle}\ x\isactrlsub {\isadigit{1}}\ {\isasymand}\ a\ {\isasymle}\ x\isactrlsub {\isadigit{2}}\ {\sf then}\ {\isasymlangle}{\isasymlangle}l\isactrlsub {\isadigit{1}}{\isacharcomma}\ x\isactrlsub {\isadigit{1}}{\isacharcomma}\ r\isactrlsub {\isadigit{1}}{\isasymrangle}{\isacharcomma}\ a{\isacharcomma}\ {\isasymlangle}l\isactrlsub {\isadigit{2}}{\isacharcomma}\ x\isactrlsub {\isadigit{2}}{\isacharcomma}\ r\isactrlsub {\isadigit{2}}{\isasymrangle}{\isasymrangle}\isanewline
+\isaindent{{\isacharparenleft}}{\sf else}\ {\sf if}\ x\isactrlsub {\isadigit{1}}\ {\isasymle}\ x\isactrlsub {\isadigit{2}}\ {\sf then}\ {\isasymlangle}sift{\isacharunderscore}down\ l\isactrlsub {\isadigit{1}}\ a\ r\isactrlsub {\isadigit{1}}{\isacharcomma}\ x\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymlangle}l\isactrlsub {\isadigit{2}}{\isacharcomma}\ x\isactrlsub {\isadigit{2}}{\isacharcomma}\ r\isactrlsub {\isadigit{2}}{\isasymrangle}{\isasymrangle}\isanewline
+\isaindent{{\isacharparenleft}{\sf else}\ }{\sf else}\ {\isasymlangle}{\isasymlangle}l\isactrlsub {\isadigit{1}}{\isacharcomma}\ x\isactrlsub {\isadigit{1}}{\isacharcomma}\ r\isactrlsub {\isadigit{1}}{\isasymrangle}{\isacharcomma}\ x\isactrlsub {\isadigit{2}}{\isacharcomma}\ sift{\isacharunderscore}down\ l\isactrlsub {\isadigit{2}}\ a\ r\isactrlsub {\isadigit{2}}{\isasymrangle}{\isacharparenright}}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Functional correctness proofs for deletion}{including preservation of \isa{heap} and \isa{braun}}
+Many lemmas, mostly straightforward
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Logarithmic complexity}
+Running time of \isa{insert}, \isa{del{\isacharunderscore}left}
+and \isa{sift{\isacharunderscore}down} (and therefore \isa{del{\isacharunderscore}min}) bounded by height
+\bigskip\pause
+
+Remember: \isa{braun\ t\ {\isasymLongrightarrow}\ {\isadigit{2}}\isactrlbsup \isa{h{\isacharparenleft}t{\isacharparenright}}\isactrlesup \ {\isasymle}\ {\isadigit{2}}\ {\isacharasterisk}\ {\isacharbar}t{\isacharbar}\ {\isacharplus}\ {\isadigit{1}}}
+\medskip\pause
+
+\isa{{\isasymLongrightarrow}}
+\medskip
+
+Above running times logarithmic in size
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}[t]{Sorting with priority queue}
+\pause
+
+\isa{pq\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ empty}\\
+\isa{pq\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ insert\ x\ {\isacharparenleft}pq\ xs{\isacharparenright}}
+\bigskip\pause
+
+\isa{mins\ q\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf if}\ is{\isacharunderscore}empty\ q\ {\sf then}\ {\isacharbrackleft}{\isacharbrackright}\isanewline
+\isaindent{{\isacharparenleft}}{\sf else}\ get{\isacharunderscore}min\ h\ {\isacharhash}\ mins\ {\isacharparenleft}del{\isacharunderscore}min\ h{\isacharparenright}{\isacharparenright}}
+\bigskip\pause
+
+\high{\isa{sort{\isacharunderscore}pq\ {\isacharequal}\ mins\ {\isasymcirc}\ pq}}
+\bigskip\pause
+
+Complexity of \isa{sort}: $O(n \log n)$ \\\pause
+if all priority queue functions have complexity $O(\log n)$
+%\bigskip\pause
+
+%Braun trees $\leadsto$ functional version of heap sort
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}[t]{Sorting with priority queue}
+
+\isa{pq\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ empty}\\
+\isa{pq\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ insert\ x\ {\isacharparenleft}pq\ xs{\isacharparenright}}
+\bigskip\pause
+
+\bad{Not optimal.} \pause\high{Linear time possible.}
+
+\end{frame}
+%-------------------------------------------------------------
+
+\begin{isabellebody}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isatagtheory
+%
+\endisatagtheory
+{\isafoldtheory}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\end{isabellebody}%
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: "root"
+%%% End:
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/tex/document/Leftist_Heap_Slides.tex Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,315 @@
+%
+\begin{isabellebody}%
+\setisabellecontext{Leftist{\isacharunderscore}Heap{\isacharunderscore}Slides}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isatagtheory
+%
+\endisatagtheory
+{\isafoldtheory}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\end{isabellebody}%
+\newcommand{\revealbeamer}[1]{\uncover<+->{\invisible<handout>{#1}}}
+\section{Priority Queues}
+%-------------------------------------------------------------
+\begin{frame}{Priority queue informally}
+\begin{center}
+Collection of elements with priorities
+\end{center}
+\pause
+Operations:
+\begin{itemize}
+\pause
+\item empty
+\pause
+\item emptiness test
+\pause
+\item insert
+\pause
+\item get element with minimal priority
+\pause
+\item delete element with minimal priority
+\end{itemize}
+\bigskip\pause
+\begin{center}
+We focus on the priorities:\\
+\high{element = priority}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Priority queues are multisets}
+\begin{center}
+The same element can be contained \high{multiple times}\\ in a priority queue
+\pause
+
+$\Longrightarrow$
+
+The abstract view of a priority queue is a \high{multiset}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Multisets in Isabelle}
+
+Import \texttt{"Library/Multiset"}
+\end{frame}
+%-------------------------------------------------------------
+%
+%-------------------------------------------------------------
+\begin{frame}{Interface of implementation}
+\begin{center}
+The type of elements (= priorities) \ \high{\isa{{\isacharprime}a}} \ is a linear order
+\end{center}
+\pause
+An implementation of a priority queue of elements of type \ \isa{{\isacharprime}a} \ must provide
+\begin{itemize}
+\pause
+\item An implementation type \ \high{\isa{{\isacharprime}q}}
+\pause
+\item \high{\isa{empty\ {\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}q}}
+\pause
+\item \high{\isa{is{\isacharunderscore}empty\ {\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}q\ {\isasymRightarrow}\ bool}}
+\pause
+\item \high{\isa{insert\ {\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}q\ {\isasymRightarrow}\ {\isacharprime}q}}
+\pause
+\item \high{\isa{get{\isacharunderscore}min\ {\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}q\ {\isasymRightarrow}\ {\isacharprime}a}}
+\pause
+\item \high{\isa{del{\isacharunderscore}min\ {\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}q\ {\isasymRightarrow}\ {\isacharprime}q}}
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{More operations}
+\begin{itemize}
+\item \high{\isa{merge\ {\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}q\ {\isasymRightarrow}\ {\isacharprime}q\ {\isasymRightarrow}\ {\isacharprime}q}}\\
+Often provided
+\pause
+\item decrease key/priority\\ Not easy in functional setting
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of implementation}
+A priority queue represents a \high{multiset} of priorities.
+
+Correctness proof requires:
+\begin{center}
+\begin{tabular}{ll}
+Abstraction function: & \high{\isa{{\isachardoublequote}mset\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}q\ {\isasymRightarrow}\ {\isacharprime}a\ multiset{\isachardoublequote}}}
+\\\pause
+Invariant: & \high{\isa{{\isachardoublequote}invar\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}q\ {\isasymRightarrow}\ bool{\isachardoublequote}}}
+\end{tabular}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of implementation}
+Must prove \ \high{\isa{invar\ q}} \isa{{\isasymLongrightarrow}}\medskip
+
+\pause
+\high{\isa{mset\ empty\ {\isacharequal}\ {\isacharbraceleft}{\isacharhash}{\isacharbraceright}}}\smallskip\\\pause
+\high{\isa{is{\isacharunderscore}empty\ q\ {\isacharequal}\ {\isacharparenleft}mset\ q\ {\isacharequal}\ {\isacharbraceleft}{\isacharhash}{\isacharbraceright}{\isacharparenright}}}\smallskip\\\pause
+\high{\isa{mset\ {\isacharparenleft}insert\ x\ q{\isacharparenright}\ {\isacharequal}\ mset\ q\ {\isacharplus}\ {\isacharbraceleft}{\isacharhash}x{\isacharhash}{\isacharbraceright}}}\smallskip\\\pause
+\high{\isa{mset\ {\isacharparenleft}del{\isacharunderscore}min\ q{\isacharparenright}\ {\isacharequal}\ mset\ q\ {\isacharminus}\ {\isacharbraceleft}{\isacharhash}get{\isacharunderscore}min\ q{\isacharhash}{\isacharbraceright}}}\smallskip\\\pause
+\high{\isa{mset\ q\ {\isasymnoteq}\ {\isacharbraceleft}{\isacharhash}{\isacharbraceright}} \isa{{\isasymLongrightarrow}}}\\
+\high{\isa{get{\isacharunderscore}min\ q\ {\isasymin}\ set\ q} \isa{{\isasymand}}}
+\high{\isa{{\isacharparenleft}{\isasymforall}x\ {\isasymin}\ set\ q{\isachardot}\ get{\isacharunderscore}min\ q\ {\isasymle}\ x{\isacharparenright}}}\\
+where \isa{set\ q\ {\isacharequal}\ set{\isacharunderscore}mset\ {\isacharparenleft}mset\ q{\isacharparenright}}
+\bigskip\pause
+
+\high{\isa{invar\ empty}}\\
+\high{\isa{invar\ {\isacharparenleft}insert\ x\ q{\isacharparenright}}}\\
+\high{\isa{invar\ {\isacharparenleft}del{\isacharunderscore}min\ q{\isacharparenright}}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Terminology}
+\begin{center}
+A tree is a \high{heap} if for every subtree\\
+the root is $\ge$ all elements in the subtrees.
+\end{center}
+\pause
+
+The term ``heap'' is frequently used synonymously with ``priority queue''.
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Priority queue via heap}
+\begin{itemize}
+\pause
+\item \isa{empty\ {\isacharequal}\ {\isasymlangle}{\isasymrangle}}
+\pause
+\item \isa{is{\isacharunderscore}empty\ h\ {\isacharequal}\ {\isacharparenleft}h\ {\isacharequal}\ {\isasymlangle}{\isasymrangle}{\isacharparenright}}
+\pause
+\item \isa{get{\isacharunderscore}min\ {\isasymlangle}{\isacharunderscore}{\isacharcomma}\ a{\isacharcomma}\ {\isacharunderscore}{\isasymrangle}\ {\isacharequal}\ a}
+\pause
+\item Assume we have \isa{merge}
+\pause
+\item \isa{insert\ a\ t\ {\isacharequal}\ merge\ {\isasymlangle}{\isasymlangle}{\isasymrangle}{\isacharcomma}\ a{\isacharcomma}\ {\isasymlangle}{\isasymrangle}{\isasymrangle}\ t}
+\pause
+\item \isa{del{\isacharunderscore}min\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\ merge\ l\ r}
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Priority queue via heap}
+A naive merge:
+\medskip\pause
+
+\isa{merge\ t\isactrlsub {\isadigit{1}}\ t\isactrlsub {\isadigit{2}}\ {\isacharequal}\ {\isacharparenleft}{\sf case}\ {\isacharparenleft}t\isactrlsub {\isadigit{1}}{\isacharcomma}t\isactrlsub {\isadigit{2}}{\isacharparenright}\ {\sf of}}\\
+\quad\isa{{\isacharparenleft}{\isasymlangle}{\isasymrangle}{\isacharcomma}\ {\isacharunderscore}{\isacharparenright}\ {\isasymRightarrow}\ t\isactrlsub {\isadigit{2}}\ {\isacharbar}}\\
+\quad\isa{{\isacharparenleft}{\isacharunderscore}{\isacharcomma}\ {\isasymlangle}{\isasymrangle}{\isacharparenright}\ {\isasymRightarrow}\ t\isactrlsub {\isadigit{1}}\ {\isacharbar}}\\\pause
+\quad\isa{{\isacharparenleft}{\isasymlangle}l\isactrlsub {\isadigit{1}}{\isacharcomma}a\isactrlsub {\isadigit{1}}{\isacharcomma}r\isactrlsub {\isadigit{1}}{\isasymrangle}{\isacharcomma}\ {\isasymlangle}l\isactrlsub {\isadigit{2}}{\isacharcomma}a\isactrlsub {\isadigit{2}}{\isacharcomma}r\isactrlsub {\isadigit{2}}{\isasymrangle}{\isacharparenright}\ {\isasymRightarrow}}\\
+\qquad\isa{{\sf if}\ a\isactrlsub {\isadigit{1}}\ {\isasymle}\ a\isactrlsub {\isadigit{2}}\ {\sf then}\ {\isasymlangle}merge\ l\isactrlsub {\isadigit{1}}\ r\isactrlsub {\isadigit{1}}{\isacharcomma}\ a\isactrlsub {\isadigit{1}}{\isacharcomma}\ t\isactrlsub {\isadigit{2}}{\isasymrangle}}\\
+\qquad\isa{{\sf else}\ {\isasymlangle}t\isactrlsub {\isadigit{1}}{\isacharcomma}\ a\isactrlsub {\isadigit{2}}{\isacharcomma}\ merge\ l\isactrlsub {\isadigit{2}}\ r\isactrlsub {\isadigit{2}}{\isasymrangle}}
+\bigskip\pause
+
+\bad{Challenge:} how to maintaining some kind of balance
+\end{frame}
+%-------------------------------------------------------------
+%
+%-------------------------------------------------------------
+\section{Leftist Heap}
+%-------------------------------------------------------------
+\thyfile{Data\char`_Structures/Leftist\char`_Heap}{}
+%-------------------------------------------------------------
+\begin{frame}{Leftist tree informally}
+The \high{rank} of a tree is the depth of the rightmost leaf.
+\pause\bigskip
+
+In a \high{leftist tree}, the rank of every left child is $\ge$
+the rank of its right sibling
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Implementation type}
+\isakeyword{datatype}\\
+\quad \high{\isa{{\isacharprime}a\ lheap\ {\isacharequal}\ Leaf\ {\isacharbar}\ Node\ nat\ {\isacharparenleft}{\isacharprime}a\ tree{\isacharparenright}\ {\isacharprime}a\ {\isacharparenleft}{\isacharprime}a\ tree{\isacharparenright}}}
+\medskip\pause
+
+Abbreviations \high{\isa{{\isasymlangle}{\isasymrangle}}} and \high{\isa{{\isasymlangle}h{\isacharcomma}\ l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}}} as usual
+\bigskip\pause
+
+Abstraction function:\\
+\isa{mset{\isacharunderscore}tree} \isa{{\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}a\ lheap\ {\isasymRightarrow}\ {\isacharprime}a\ multiset}\\\smallskip\pause
+\high{\isa{mset{\isacharunderscore}tree\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ {\isacharbraceleft}{\isacharhash}{\isacharbraceright}}}\\
+\high{\isa{mset{\isacharunderscore}tree\ {\isasymlangle}\_{\isacharcomma}\ l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\isanewline
+{\isacharbraceleft}{\isacharhash}a{\isacharhash}{\isacharbraceright}\ {\isacharplus}\ mset{\isacharunderscore}tree\ l\ {\isacharplus}\ mset{\isacharunderscore}tree\ r}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Heap}
+\isa{heap} \isa{{\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}a\ lheap\ {\isasymRightarrow}\ bool}\\\smallskip\pause
+\high{\isa{heap\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ True}}\\
+\high{\isa{heap\ {\isasymlangle}{\isacharunderscore}{\isacharcomma}\ l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\ {\isacharparenleft}heap\ l\ {\isasymand}\ heap\ r\ {\isasymand}}}\\
+\quad\high{\isa{{\isacharparenleft}{\isasymforall}x\ {\isasymin}{\isacharhash}\ mset{\isacharunderscore}tree\ l\ {\isacharplus}\ mset{\isacharunderscore}tree\ r{\isachardot}\ a\ {\isasymle}\ x{\isacharparenright}{\isacharparenright}}}
+
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Leftist tree}
+\isa{rank} \isa{{\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}a\ lheap\ {\isasymRightarrow}\ nat}\\\smallskip\pause
+\high{\isa{rank\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ {\isadigit{0}}}}\\
+\high{\isa{rank\ {\isasymlangle}\_{\isacharcomma}\ \_{\isacharcomma}\ \_{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\ rank\ r\ {\isacharplus}\ {\isadigit{1}}}}
+\bigskip\pause
+
+Node \isa{{\isasymlangle}n{\isacharcomma}\ l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}}: \isa{n\ {\isacharequal}} rank of node
+\bigskip\pause
+
+\isa{ltree} \isa{{\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}a\ lheap\ {\isasymRightarrow}\ bool}\\\smallskip\pause
+\high{\isa{ltree\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ True}}\\
+\high{\isa{ltree\ {\isasymlangle}n{\isacharcomma}\ l{\isacharcomma}\ \_{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\isanewline
+{\isacharparenleft}n\ {\isacharequal}\ rank\ r\ {\isacharplus}\ {\isadigit{1}}\ {\isasymand}\ rank\ r\ {\isasymle}\ rank\ l\ {\isasymand}\ ltree\ l\ {\isasymand}\ ltree\ r{\isacharparenright}}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Leftist heap invariant}
+\begin{center}
+\high{\isa{invar\ h\ {\isacharequal}\ {\isacharparenleft}heap\ h\ {\isasymand}\ ltree\ h{\isacharparenright}}}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Why leftist tree?}
+\pause
+
+\textbf{Lemma} \high{\isa{ltree\ t\ {\isasymLongrightarrow}\ {\isadigit{2}}\isactrlbsup \isa{rank\ t}\isactrlesup \ {\isasymle}\ {\isacharbar}t{\isacharbar}\isactrlsub {\isadigit{1}}}}
+\bigskip\pause
+
+\textbf{Lemma} Execution time of \isa{merge\ t\isactrlsub {\isadigit{1}}\ t\isactrlsub {\isadigit{2}}} is bounded by
+\mbox{\isa{rank\ t\isactrlsub {\isadigit{1}}\ {\isacharplus}\ rank\ t\isactrlsub {\isadigit{2}}}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\isa{merge}}
+Principle: descend on the right
+\smallskip\pause
+
+\high{\isa{merge\ {\isasymlangle}{\isasymrangle}\ t\isactrlsub {\isadigit{2}}\ {\isacharequal}\ t\isactrlsub {\isadigit{2}}\isasep\isanewline%
+merge\ {\isasymlangle}v{\isacharcomma}\ va{\isacharcomma}\ vb{\isacharcomma}\ vc{\isasymrangle}\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ {\isasymlangle}v{\isacharcomma}\ va{\isacharcomma}\ vb{\isacharcomma}\ vc{\isasymrangle}}}\\\pause
+\high{\isa{merge\ {\isasymlangle}n\isactrlsub {\isadigit{1}}{\isacharcomma}\ l\isactrlsub {\isadigit{1}}{\isacharcomma}\ a\isactrlsub {\isadigit{1}}{\isacharcomma}\ r\isactrlsub {\isadigit{1}}{\isasymrangle}\ {\isasymlangle}n\isactrlsub {\isadigit{2}}{\isacharcomma}\ l\isactrlsub {\isadigit{2}}{\isacharcomma}\ a\isactrlsub {\isadigit{2}}{\isacharcomma}\ r\isactrlsub {\isadigit{2}}{\isasymrangle}\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf if}\ a\isactrlsub {\isadigit{1}}\ {\isasymle}\ a\isactrlsub {\isadigit{2}}\ {\sf then}\ node\ l\isactrlsub {\isadigit{1}}\ a\isactrlsub {\isadigit{1}}\ {\isacharparenleft}merge\ r\isactrlsub {\isadigit{1}}\ {\isasymlangle}n\isactrlsub {\isadigit{2}}{\isacharcomma}\ l\isactrlsub {\isadigit{2}}{\isacharcomma}\ a\isactrlsub {\isadigit{2}}{\isacharcomma}\ r\isactrlsub {\isadigit{2}}{\isasymrangle}{\isacharparenright}\isanewline
+\isaindent{{\isacharparenleft}}{\sf else}\ node\ l\isactrlsub {\isadigit{2}}\ a\isactrlsub {\isadigit{2}}\ {\isacharparenleft}merge\ r\isactrlsub {\isadigit{2}}\ {\isasymlangle}n\isactrlsub {\isadigit{1}}{\isacharcomma}\ l\isactrlsub {\isadigit{1}}{\isacharcomma}\ a\isactrlsub {\isadigit{1}}{\isacharcomma}\ r\isactrlsub {\isadigit{1}}{\isasymrangle}{\isacharparenright}{\isacharparenright}}}
+\bigskip\pause
+
+\isa{node} \isa{{\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}a\ lheap\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ lheap\ {\isasymRightarrow}\ {\isacharprime}a\ lheap}\\\smallskip\pause
+\high{\isa{node\ l\ a\ r\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf let}\ rl\ {\isacharequal}\ rk\ l{\isacharsemicolon}\ rr\ {\isacharequal}\ rk\ r\isanewline
+\isaindent{{\isacharparenleft}}{\sf in}\ {\sf if}\ rr\ {\isasymle}\ rl\ {\sf then}\ {\isasymlangle}rr\ {\isacharplus}\ {\isadigit{1}}{\isacharcomma}\ l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\ {\sf else}\ {\isasymlangle}rl\ {\isacharplus}\ {\isadigit{1}}{\isacharcomma}\ r{\isacharcomma}\ a{\isacharcomma}\ l{\isasymrangle}{\isacharparenright}}}\\
+where \isa{rk\ {\isasymlangle}n{\isacharcomma}\ \_{\isacharcomma}\ \_{\isacharcomma}\ \_{\isasymrangle}\ {\isacharequal}\ n}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\isa{merge}}
+\isa{merge\ {\isasymlangle}n\isactrlsub {\isadigit{1}}{\isacharcomma}\ l\isactrlsub {\isadigit{1}}{\isacharcomma}\ a\isactrlsub {\isadigit{1}}{\isacharcomma}\ r\isactrlsub {\isadigit{1}}{\isasymrangle}\ {\isasymlangle}n\isactrlsub {\isadigit{2}}{\isacharcomma}\ l\isactrlsub {\isadigit{2}}{\isacharcomma}\ a\isactrlsub {\isadigit{2}}{\isacharcomma}\ r\isactrlsub {\isadigit{2}}{\isasymrangle}\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf if}\ a\isactrlsub {\isadigit{1}}\ {\isasymle}\ a\isactrlsub {\isadigit{2}}\ {\sf then}\ node\ l\isactrlsub {\isadigit{1}}\ a\isactrlsub {\isadigit{1}}\ {\isacharparenleft}merge\ r\isactrlsub {\isadigit{1}}\ {\isasymlangle}n\isactrlsub {\isadigit{2}}{\isacharcomma}\ l\isactrlsub {\isadigit{2}}{\isacharcomma}\ a\isactrlsub {\isadigit{2}}{\isacharcomma}\ r\isactrlsub {\isadigit{2}}{\isasymrangle}{\isacharparenright}\isanewline
+\isaindent{{\isacharparenleft}}{\sf else}\ node\ l\isactrlsub {\isadigit{2}}\ a\isactrlsub {\isadigit{2}}\ {\isacharparenleft}merge\ r\isactrlsub {\isadigit{2}}\ {\isasymlangle}n\isactrlsub {\isadigit{1}}{\isacharcomma}\ l\isactrlsub {\isadigit{1}}{\isacharcomma}\ a\isactrlsub {\isadigit{1}}{\isacharcomma}\ r\isactrlsub {\isadigit{1}}{\isasymrangle}{\isacharparenright}{\isacharparenright}}
+\bigskip\pause
+
+\high{\onslide<2->{Function \isa{merge} terminates because}
+\onslide<3>{\revealbeamer{\isa{rank\ t\isactrlsub {\isadigit{1}}\ {\isacharplus}\ rank\ t\isactrlsub {\isadigit{2}}}}}
+\onslide<2->{decreases with every recursive call.}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Functional correctness proofs}{including preservation of \isa{invar}}
+Straightforward
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Logarithmic complexity}
+Complexity measures \isa{t{\isacharunderscore}merge}, \isa{t{\isacharunderscore}insert} \isa{t{\isacharunderscore}del{\isacharunderscore}min}:\\
+ count calls of \isa{merge}.
+\pause\bigskip
+
+\textbf{Lemma} \isa{t{\isacharunderscore}merge\ l\ r\ {\isasymle}\ rank\ l\ {\isacharplus}\ rank\ r\ {\isacharplus}\ {\isadigit{1}}}
+\pause
+
+\textbf{Lemma} \isa{{\isasymlbrakk}ltree\ l{\isacharsemicolon}\ ltree\ r{\isasymrbrakk}\isanewline
+{\isasymLongrightarrow}\ t{\isacharunderscore}merge\ l\ r\ {\isasymle}\ $\log$\isactrlbsub {\isadigit{2}}\isactrlesub \ {\isacharbar}l{\isacharbar}\isactrlsub {\isadigit{1}}\ {\isacharplus}\ $\log$\isactrlbsub {\isadigit{2}}\isactrlesub \ {\isacharbar}r{\isacharbar}\isactrlsub {\isadigit{1}}\ {\isacharplus}\ {\isadigit{1}}}
+\pause\bigskip
+
+\textbf{Lemma}\\ \isa{ltree\ t\ {\isasymLongrightarrow}\ t{\isacharunderscore}insert\ x\ t\ {\isasymle}\ $\log$\isactrlbsub {\isadigit{2}}\isactrlesub \ {\isacharbar}t{\isacharbar}\isactrlsub {\isadigit{1}}\ {\isacharplus}\ {\isadigit{2}}}
+\pause\bigskip
+
+\textbf{Lemma}\\ \isa{ltree\ t\ {\isasymLongrightarrow}\ t{\isacharunderscore}del{\isacharunderscore}min\ t\ {\isasymle}\ {\isadigit{2}}\ {\isacharasterisk}\ $\log$\isactrlbsub {\isadigit{2}}\isactrlesub \ {\isacharbar}t{\isacharbar}\isactrlsub {\isadigit{1}}\ {\isacharplus}\ {\isadigit{1}}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}
+\begin{center}
+\bad{Can we avoid the rank info in each node?}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+
+\begin{isabellebody}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isatagtheory
+%
+\endisatagtheory
+{\isafoldtheory}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\end{isabellebody}%
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: "root"
+%%% End:
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/tex/document/RBT_Slides.tex Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,277 @@
+%
+\begin{isabellebody}%
+\setisabellecontext{RBT{\isacharunderscore}Slides}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isatagtheory
+%
+\endisatagtheory
+{\isafoldtheory}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\end{isabellebody}%
+\section{Red-Black Trees}
+%-------------------------------------------------------------
+\thyfile{Thys/RBT\char`_Set}{}
+%-------------------------------------------------------------
+\begin{frame}{Relationship to 2-3-4 trees}
+Idea:
+\begin{tabular}[t]{l}
+ encode 2-3-4 trees as binary trees;\\\pause
+ use color to express grouping
+\end{tabular}
+\pause\bigskip
+
+\begin{tabular}{rcl}
+\isa{{\isasymlangle}{\isasymrangle}} &$\approx$& \isa{{\isasymlangle}{\isasymrangle}}\\\pause
+\isa{{\isasymlangle}t\isactrlsub {\isadigit{1}}{\isacharcomma}a{\isacharcomma}t\isactrlsub {\isadigit{2}}{\isasymrangle}} &$\approx$& \isa{{\isasymlangle}t\isactrlsub {\isadigit{1}}{\isacharcomma}a{\isacharcomma}t\isactrlsub {\isadigit{2}}{\isasymrangle}}\\\pause
+\isa{{\isasymlangle}t\isactrlsub {\isadigit{1}}{\isacharcomma}a{\isacharcomma}t\isactrlsub {\isadigit{2}}{\isacharcomma}b{\isacharcomma}t\isactrlsub {\isadigit{3}}{\isasymrangle}} &$\approx$&\pause \isa{{\isasymlangle}}\bad{\isa{{\isasymlangle}t\isactrlsub {\isadigit{1}}{\isacharcomma}a{\isacharcomma}t\isactrlsub {\isadigit{2}}{\isasymrangle}}}\isa{{\isacharcomma}b{\isacharcomma}t\isactrlsub {\isadigit{3}}{\isasymrangle}}
+ \pause\isa{{\isasymlangle}t\isactrlsub {\isadigit{1}}{\isacharcomma}a{\isacharcomma}}\bad{\isa{{\isasymlangle}t\isactrlsub {\isadigit{2}}{\isacharcomma}b{\isacharcomma}t\isactrlsub {\isadigit{3}}{\isasymrangle}}}\isa{{\isasymrangle}}\\\pause
+\isa{{\isasymlangle}t\isactrlsub {\isadigit{1}}{\isacharcomma}a{\isacharcomma}t\isactrlsub {\isadigit{2}}{\isacharcomma}b{\isacharcomma}t\isactrlsub {\isadigit{3}}{\isacharcomma}c{\isacharcomma}t\isactrlsub {\isadigit{4}}{\isasymrangle}} &$\approx$& \pause\isa{{\isasymlangle}}\bad{\isa{{\isasymlangle}t\isactrlsub {\isadigit{1}}{\isacharcomma}a{\isacharcomma}t\isactrlsub {\isadigit{2}}{\isasymrangle}}}\isa{{\isacharcomma}b{\isacharcomma}}\bad{\isa{{\isasymlangle}t\isactrlsub {\isadigit{3}}{\isacharcomma}c{\isacharcomma}t\isactrlsub {\isadigit{4}}{\isasymrangle}}}\isa{{\isasymrangle}}\\
+\end{tabular}
+\bigskip\pause
+
+\begin{center}
+\bad{Red} means ``I am part of a bigger node''
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Invariants}
+\begin{itemize}
+\pause
+\item The root is \pause\invisible<handout>{Black.}
+\pause
+\item Every \isa{{\isasymlangle}{\isasymrangle}} is considered Black.
+\pause
+\item If a node is \bad{Red},
+ \pause\invisible<handout>{its children are Black.}
+\pause
+\item All paths from a node to a leaf have the same number of
+ \pause\invisible<handout>{Black nodes.}
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Red-black trees}
+%
+\begin{isabelle}%
+\isacommand{datatype}\ color\ {\isacharequal}\ Red\ {\isacharbar}\ Black%
+\end{isabelle}
+\pause
+\isakeyword{datatype}\\
+\quad \high{\isa{{\isacharprime}a\ rbt\ {\isacharequal}\ Leaf\ {\isacharbar}\ Node\ color\ {\isacharparenleft}{\isacharprime}a\ tree{\isacharparenright}\ {\isacharprime}a\ {\isacharparenleft}{\isacharprime}a\ tree{\isacharparenright}}}
+\bigskip\pause
+
+Abbreviations:
+\smallskip
+\begin{center}
+\begin{tabular}{rcl}
+\high{\isa{{\isasymlangle}{\isasymrangle}}} &\isa{{\isasymequiv}}& \isa{Leaf}\\\pause
+\high{\isa{{\isasymlangle}c{\isacharcomma}\ l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}}} &\isa{{\isasymequiv}}& \isa{{\isachardoublequote}Node\ c\ l\ a\ r{\isachardoublequote}}\medskip
+\\\pause
+\high{\isa{R\ l\ a\ r}} &\isa{{\isasymequiv}}& \isa{{\isachardoublequote}Node\ Red\ l\ a\ r{\isachardoublequote}}\\\pause
+\high{\isa{B\ l\ a\ r}} &\isa{{\isasymequiv}}& \isa{{\isachardoublequote}Node\ Black\ l\ a\ r{\isachardoublequote}}
+\end{tabular}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Color}
+\pause
+
+\isa{color} \isa{{\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}a\ rbt\ {\isasymRightarrow}\ color}\\\smallskip\pause
+\high{\isa{color\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ Black}}\\
+\high{\isa{color\ {\isasymlangle}c{\isacharcomma}\ \_{\isacharcomma}\ \_{\isacharcomma}\ \_{\isasymrangle}\ {\isacharequal}\ c}}
+\bigskip\pause
+
+\isa{paint} \isa{{\isacharcolon}{\isacharcolon}} \isa{color\ {\isasymRightarrow}\ {\isacharprime}a\ rbt\ {\isasymRightarrow}\ {\isacharprime}a\ rbt}\\\smallskip\pause
+\high{\isa{paint\ c\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ {\isasymlangle}{\isasymrangle}}}\\
+\high{\isa{paint\ c\ {\isasymlangle}\_{\isacharcomma}\ l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\ {\isasymlangle}c{\isacharcomma}\ l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Invariants}
+\pause
+
+\isa{rbt} \isa{{\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}a\ rbt\ {\isasymRightarrow}\ bool}\\\smallskip\pause
+\high{\isa{rbt\ t\ {\isacharequal}\ {\isacharparenleft}invc\ t\ {\isasymand}\ invh\ t\ {\isasymand}\ color\ t\ {\isacharequal}\ Black{\isacharparenright}}}
+\bigskip\pause
+
+\isa{invc} \isa{{\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}a\ rbt\ {\isasymRightarrow}\ bool}\\\smallskip\pause
+\high{\isa{invc\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ True\isasep\isanewline%
+invc\ {\isasymlangle}c{\isacharcomma}\ l{\isacharcomma}\ \_{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\isanewline
+{\isacharparenleft}invc\ l\ {\isasymand}\isanewline
+\isaindent{{\isacharparenleft}}invc\ r\ {\isasymand}\isanewline
+\isaindent{{\isacharparenleft}}{\isacharparenleft}c\ {\isacharequal}\ Red\ {\isasymlongrightarrow}\ color\ l\ {\isacharequal}\ Black\ {\isasymand}\ color\ r\ {\isacharequal}\ Black{\isacharparenright}{\isacharparenright}}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Invariants}
+
+\isa{invh} \isa{{\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}a\ rbt\ {\isasymRightarrow}\ bool}\\\smallskip\pause
+\high{\isa{invh\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ True\isasep\isanewline%
+invh\ {\isasymlangle}\_{\isacharcomma}\ l{\isacharcomma}\ \_{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\ {\isacharparenleft}invh\ l\ {\isasymand}\ invh\ r\ {\isasymand}\ bh{\isacharparenleft}l{\isacharparenright}\ {\isacharequal}\ bh{\isacharparenleft}r{\isacharparenright}{\isacharparenright}}}
+\bigskip\pause
+
+\isa{bheight} \isa{{\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}a\ rbt\ {\isasymRightarrow}\ nat}\\\smallskip\pause
+\high{\isa{bh{\isacharparenleft}{\isasymlangle}{\isasymrangle}{\isacharparenright}\ {\isacharequal}\ {\isadigit{0}}\isasep\isanewline%
+bh{\isacharparenleft}{\isasymlangle}c{\isacharcomma}\ l{\isacharcomma}\ \_{\isacharcomma}\ \_{\isasymrangle}{\isacharparenright}\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf if}\ c\ {\isacharequal}\ Black\ {\sf then}\ bh{\isacharparenleft}l{\isacharparenright}\ {\isacharplus}\ {\isadigit{1}}\ {\sf else}\ bh{\isacharparenleft}l{\isacharparenright}{\isacharparenright}}}
+\end{frame}
+%-------------------------------------------------------------
+%\begin{frame}{Exercise}
+%\begin{center}
+%\bad{Is @ {const invh} what we want?}
+%\end{center}
+%\pause
+%
+%Define a function \ \high{@ {const Bpl} \isa{{\isacharcolon}{\isacharcolon}} @ {typ "'a rbt \<Rightarrow> nat set"}}\\
+%such that @ {term "Bpl t"} (``black path lengths'')
+%is the set of all \isa{n} such that there is a path
+%from the root of \isa{t} to a leaf that contains exactly \isa{n} black nodes.
+%\pause\medskip
+%
+%Prove \ \high{@ {thm invh_Bpl}}
+%\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Logarithmic height}
+\begin{lemma}
+\isa{rbt\ t\ {\isasymLongrightarrow}\ h{\isacharparenleft}t{\isacharparenright}\ {\isasymle}\ {\isadigit{2}}\ {\isacharasterisk}\ $\log$\isactrlbsub {\isadigit{2}}\isactrlesub \ {\isacharbar}t{\isacharbar}\isactrlsub {\isadigit{1}}}
+\end{lemma}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Insertion}
+
+\isa{insert} \isa{{\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ rbt\ {\isasymRightarrow}\ {\isacharprime}a\ rbt}\\\smallskip\pause
+\high{\isa{insert\ x\ t\ {\isacharequal}\ paint\ Black\ {\isacharparenleft}ins\ x\ t{\isacharparenright}}}
+\bigskip\pause
+
+\isa{ins} \isa{{\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ rbt\ {\isasymRightarrow}\ {\isacharprime}a\ rbt}\\\smallskip\pause
+\high{\isa{ins\ x\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ R\ {\isasymlangle}{\isasymrangle}\ x\ {\isasymlangle}{\isasymrangle}}}\pause\vspace{-1ex}
+\high{%
+\begin{isabelle}%
+ins\ x\ {\isacharparenleft}B\ l\ a\ r{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\sf case}\ cmp\ x\ a\ {\sf of}\isanewline
+\isaindent{ins\ x\ {\isacharparenleft}B\ l\ a\ r{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}}\ \ LT\ {\isasymRightarrow}\ baliL\ {\isacharparenleft}ins\ x\ l{\isacharparenright}\ a\ r\isanewline
+\isaindent{ins\ x\ {\isacharparenleft}B\ l\ a\ r{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}}{\isacharbar}\ EQ\ {\isasymRightarrow}\ B\ l\ a\ r\isanewline
+\isaindent{ins\ x\ {\isacharparenleft}B\ l\ a\ r{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}}{\isacharbar}\ GT\ {\isasymRightarrow}\ baliR\ l\ a\ {\isacharparenleft}ins\ x\ r{\isacharparenright}{\isacharparenright}%
+\end{isabelle}}\pause\vspace{-1ex}
+\high{%
+\begin{isabelle}%
+ins\ x\ {\isacharparenleft}R\ l\ a\ r{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\sf case}\ cmp\ x\ a\ {\sf of}\isanewline
+\isaindent{ins\ x\ {\isacharparenleft}R\ l\ a\ r{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}}\ \ LT\ {\isasymRightarrow}\ R\ {\isacharparenleft}ins\ x\ l{\isacharparenright}\ a\ r\isanewline
+\isaindent{ins\ x\ {\isacharparenleft}R\ l\ a\ r{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}}{\isacharbar}\ EQ\ {\isasymRightarrow}\ R\ l\ a\ r\isanewline
+\isaindent{ins\ x\ {\isacharparenleft}R\ l\ a\ r{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}}{\isacharbar}\ GT\ {\isasymRightarrow}\ R\ l\ a\ {\isacharparenleft}ins\ x\ r{\isacharparenright}{\isacharparenright}%
+\end{isabelle}}
+
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Adjusting colors}
+\isa{baliL}, \isa{baliR} \isa{{\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}a\ rbt\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ rbt\ {\isasymRightarrow}\ {\isacharprime}a\ rbt}
+\begin{itemize}
+\pause
+\item Combine arguments \isa{l} \isa{a} \isa{r} into tree, ideally \isa{{\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}}
+\pause
+\item Treat invariant violation \bad{Red-Red} in \isa{l}/\isa{r}\\\pause
+\high{\isa{baliL\ {\isacharparenleft}R\ {\isacharparenleft}R\ t\isactrlsub {\isadigit{1}}\ a\isactrlsub {\isadigit{1}}\ t\isactrlsub {\isadigit{2}}{\isacharparenright}\ a\isactrlsub {\isadigit{2}}\ t\isactrlsub {\isadigit{3}}{\isacharparenright}\ a\isactrlsub {\isadigit{3}}\ t\isactrlsub {\isadigit{4}}}\\
+\quad \isa{{\isacharequal}} \isa{R\ {\isacharparenleft}B\ t\isactrlsub {\isadigit{1}}\ a\isactrlsub {\isadigit{1}}\ t\isactrlsub {\isadigit{2}}{\isacharparenright}\ a\isactrlsub {\isadigit{2}}\ {\isacharparenleft}B\ t\isactrlsub {\isadigit{3}}\ a\isactrlsub {\isadigit{3}}\ t\isactrlsub {\isadigit{4}}{\isacharparenright}}}
+\pause
+
+\high{\isa{baliL\ {\isacharparenleft}R\ t\isactrlsub {\isadigit{1}}\ a\isactrlsub {\isadigit{1}}\ {\isacharparenleft}R\ t\isactrlsub {\isadigit{2}}\ a\isactrlsub {\isadigit{2}}\ t\isactrlsub {\isadigit{3}}{\isacharparenright}{\isacharparenright}\ a\isactrlsub {\isadigit{3}}\ t\isactrlsub {\isadigit{4}}}\\
+\quad \isa{{\isacharequal}} \isa{R\ {\isacharparenleft}B\ t\isactrlsub {\isadigit{1}}\ a\isactrlsub {\isadigit{1}}\ t\isactrlsub {\isadigit{2}}{\isacharparenright}\ a\isactrlsub {\isadigit{2}}\ {\isacharparenleft}B\ t\isactrlsub {\isadigit{3}}\ a\isactrlsub {\isadigit{3}}\ t\isactrlsub {\isadigit{4}}{\isacharparenright}}}
+\pause
+\item Principle: replace \bad{Red-Red} by \high{Red-Black}
+\pause
+\item Final equation: \\\pause
+\high{\isa{baliL\ l\ a\ r\ {\isacharequal}\ B\ l\ a\ r}}
+\pause
+\item Symmetric: \isa{baliR}
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+%\begin{frame}{Correctness via sorted lists}
+%\pause
+%\begin{lemma}
+%\high{@ {thm inorder_baliL}}\\
+%\high{@ {thm inorder_baliR}}
+%\end{lemma}
+%\pause
+%\begin{lemma}
+%\high{@ {thm[break,margin=60] inorder_ins}}
+%\end{lemma}
+%\pause
+%\begin{corollary}
+%\high{@ {thm[break,margin=60] inorder_insert}}
+%\end{corollary}
+%\bigskip\pause
+%
+%Proofs easy!
+%\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Preservation of invariant}
+\begin{theorem}
+\high{\isa{rbt\ t\ {\isasymLongrightarrow}\ rbt\ {\isacharparenleft}insert\ x\ t{\isacharparenright}}}
+\end{theorem}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Deletion}
+
+\high{\isa{delete\ x\ t\ {\isacharequal}\ paint\ Black\ {\isacharparenleft}del\ x\ t{\isacharparenright}}}
+\bigskip\pause
+
+\high{\isa{del\ \_\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ {\isasymlangle}{\isasymrangle}\isasep\isanewline%
+del\ x\ {\isasymlangle}\_{\isacharcomma}\ l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf case}\ cmp\ x\ a\ {\sf of}\isanewline
+\isaindent{{\isacharparenleft}}\ \ LT\ {\isasymRightarrow}\isanewline
+\isaindent{{\isacharparenleft}\ \ \ \ }{\sf if}\ l\ {\isasymnoteq}\ {\isasymlangle}{\isasymrangle}\ {\isasymand}\ color\ l\ {\isacharequal}\ Black\isanewline
+\isaindent{{\isacharparenleft}\ \ \ \ }{\sf then}\ baldL\ {\isacharparenleft}del\ x\ l{\isacharparenright}\ a\ r\ {\sf else}\ R\ {\isacharparenleft}del\ x\ l{\isacharparenright}\ a\ r\isanewline
+\isaindent{{\isacharparenleft}}{\isacharbar}\ EQ\ {\isasymRightarrow}\ combine\ l\ r\isanewline
+\isaindent{{\isacharparenleft}}{\isacharbar}\ GT\ {\isasymRightarrow}\isanewline
+\isaindent{{\isacharparenleft}{\isacharbar}\ \ \ }{\sf if}\ r\ {\isasymnoteq}\ {\isasymlangle}{\isasymrangle}\ {\isasymand}\ color\ r\ {\isacharequal}\ Black\isanewline
+\isaindent{{\isacharparenleft}{\isacharbar}\ \ \ }{\sf then}\ baldR\ l\ a\ {\isacharparenleft}del\ x\ r{\isacharparenright}\ {\sf else}\ R\ l\ a\ {\isacharparenleft}del\ x\ r{\isacharparenright}{\isacharparenright}}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Deletion}
+Tricky functions: \high{\isa{baldL}, \isa{baldR}, \isa{combine}}
+\bigskip\pause
+
+Tricky invariant preservation lemmas, eg
+
+\high{%
+\begin{isabelle}%
+{\isasymlbrakk}invh\ t{\isacharsemicolon}\ invc\ t{\isasymrbrakk}\isanewline
+{\isasymLongrightarrow}\ invh\ {\isacharparenleft}del\ x\ t{\isacharparenright}\ {\isasymand}\isanewline
+\isaindent{{\isasymLongrightarrow}\ }{\isacharparenleft}color\ t\ {\isacharequal}\ Red\ {\isasymand}\isanewline
+\isaindent{{\isasymLongrightarrow}\ {\isacharparenleft}}bh{\isacharparenleft}del\ x\ t{\isacharparenright}\ {\isacharequal}\ bh{\isacharparenleft}t{\isacharparenright}\ {\isasymand}\ invc\ {\isacharparenleft}del\ x\ t{\isacharparenright}\ {\isasymor}\isanewline
+\isaindent{{\isasymLongrightarrow}\ {\isacharparenleft}}color\ t\ {\isacharequal}\ Black\ {\isasymand}\isanewline
+\isaindent{{\isasymLongrightarrow}\ {\isacharparenleft}}bh{\isacharparenleft}del\ x\ t{\isacharparenright}\ {\isacharequal}\ bh{\isacharparenleft}t{\isacharparenright}\ {\isacharminus}\ {\isadigit{1}}\ {\isasymand}\ invc{\isadigit{2}}\ {\isacharparenleft}del\ x\ t{\isacharparenright}{\isacharparenright}%
+\end{isabelle}}
+\pause
+
+\begin{theorem}
+\high{\isa{rbt\ t\ {\isasymLongrightarrow}\ rbt\ {\isacharparenleft}delete\ k\ t{\isacharparenright}}}
+\end{theorem}
+\end{frame}
+%-------------------------------------------------------------
+
+\begin{isabellebody}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isatagtheory
+%
+\endisatagtheory
+{\isafoldtheory}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\end{isabellebody}%
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: "root"
+%%% End:
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/tex/document/Search_Trees.tex Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,331 @@
+%
+\begin{isabellebody}%
+\setisabellecontext{Search{\isacharunderscore}Trees}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isatagtheory
+%
+\endisatagtheory
+{\isafoldtheory}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\end{isabellebody}%
+%-------------------------------------------------------------
+\begin{frame}
+\begin{center}
+Most of the material focuses on\\ \high{BST}s = \high{binary search trees}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{BSTs represent sets}
+\pause
+Any tree represents a set:
+\medskip
+
+\isa{set{\isacharunderscore}tree\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ tree\ {\isasymRightarrow}\ {\isacharprime}a\ set}\\\smallskip\pause
+\high{\isa{set{\isacharunderscore}tree\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ {\isacharbraceleft}{\isacharbraceright}}}\\
+\high{\isa{set{\isacharunderscore}tree\ {\isasymlangle}l{\isacharcomma}\ x{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\ set{\isacharunderscore}tree\ l\ {\isasymunion}\ {\isacharbraceleft}x{\isacharbraceright}\ {\isasymunion}\ set{\isacharunderscore}tree\ r}}
+\bigskip\pause
+
+A BST represents a set that can be searched in time $O(h(t))$
+\pause
+\begin{center}
+Function \isa{set{\isacharunderscore}tree} is called an \high{\emph{abstraction function}}
+because it maps the implementation \\ to the abstract mathematical object
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\isa{bst}}
+\isa{bst\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ tree\ {\isasymRightarrow}\ bool}\pause
+\high{%
+\begin{isabelle}%
+bst\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ True\isasep\isanewline%
+bst\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\isanewline
+{\isacharparenleft}bst\ l\ {\isasymand}\isanewline
+\isaindent{{\isacharparenleft}}bst\ r\ {\isasymand}\isanewline
+\isaindent{{\isacharparenleft}}{\isacharparenleft}{\isasymforall}x{\isasymin}set{\isacharunderscore}tree\ l{\isachardot}\ x\ {\isacharless}\ a{\isacharparenright}\ {\isasymand}\isanewline
+\isaindent{{\isacharparenleft}}{\isacharparenleft}{\isasymforall}x{\isasymin}set{\isacharunderscore}tree\ r{\isachardot}\ a\ {\isacharless}\ x{\isacharparenright}{\isacharparenright}%
+\end{isabelle}}
+\vfill\pause
+Type \isa{{\isacharprime}a} must be in class \isa{linorder} (\isa{{\isachardoublequote}{\isacharprime}a\ {\isacharcolon}{\isacharcolon}\ linorder{\isachardoublequote}})
+where \isa{linorder} are \emph{linear orders} (also called \emph{total orders}).
+\bigskip\pause
+
+Note: \isa{nat}, \isa{int} and \isa{real} are in class \isa{linorder}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Set interface}
+An implementation of sets of elements of type \ \high{\isa{{\isacharprime}a}} \ must provide
+\begin{itemize}
+\pause
+\item An implementation type \ \high{\isa{{\isacharprime}s}}
+\pause
+\item \high{\isa{empty\ {\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}s}}
+\pause
+\item \high{\isa{insert\ {\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}s\ {\isasymRightarrow}\ {\isacharprime}s}}
+\pause
+\item \high{\isa{delete\ {\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}s\ {\isasymRightarrow}\ {\isacharprime}s}}
+\pause
+\item \high{\isa{isin\ {\isacharcolon}{\isacharcolon}} \quad \isa{{\isacharprime}s\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ bool}}
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Map interface}
+Instead of a set, a search tree can also implement a \high{map} from \isa{{\isacharprime}a} to \isa{{\isacharprime}b}:
+\begin{itemize}
+\pause
+\item An implementation type \ \high{\isa{{\isacharprime}m}}
+\pause
+\item \high{\isa{empty\ {\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}m}}
+\pause
+\item \high{\isa{update\ {\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ {\isasymRightarrow}\ {\isacharprime}m\ {\isasymRightarrow}\ {\isacharprime}m}}
+\pause
+\item \high{\isa{delete\ {\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}m\ {\isasymRightarrow}\ {\isacharprime}m}}
+\pause
+\item \high{\isa{lookup\ {\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}m\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ option}}
+\end{itemize}
+\pause\bigskip
+Sets are a special case of maps
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Comparison of elements}
+\pause
+\begin{center}
+\high{We assume that the element type \isa{{\isacharprime}a} is a linear order}
+\end{center}
+\pause
+Instead of using \isa{{\isacharless}} and \isa{{\isasymle}} directly:
+\high{%
+\begin{isabelle}%
+\isacommand{datatype}\ cmp{\isacharunderscore}val\ {\isacharequal}\ LT\ {\isacharbar}\ EQ\ {\isacharbar}\ GT%
+\end{isabelle}}
+\pause
+\high{\isa{cmp\ x\ y\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf if}\ x\ {\isacharless}\ y\ {\sf then}\ LT\ {\sf else}\ {\sf if}\ x\ {\isacharequal}\ y\ {\sf then}\ EQ\ {\sf else}\ GT{\isacharparenright}}}
+\end{frame}
+%-------------------------------------------------------------
+\section{Unbalanced BST}
+%-------------------------------------------------------------
+\begin{frame}{Implementation}
+Implementation type: \isa{{\isacharprime}a\ tree}
+\bigskip\pause
+
+\high{\isa{empty\ {\isacharequal}\ Leaf}}
+\pause
+
+\high{%
+\begin{isabelle}%
+insert\ x\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ {\isasymlangle}{\isasymlangle}{\isasymrangle}{\isacharcomma}\ x{\isacharcomma}\ {\isasymlangle}{\isasymrangle}{\isasymrangle}\isasep\isanewline%
+insert\ x\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\ {\isacharparenleft}{\sf case}\ cmp\ x\ a\ {\sf of}\isanewline
+\isaindent{insert\ x\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\ {\isacharparenleft}}\ \ LT\ {\isasymRightarrow}\ {\isasymlangle}insert\ x\ l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\isanewline
+\isaindent{insert\ x\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\ {\isacharparenleft}}{\isacharbar}\ EQ\ {\isasymRightarrow}\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\isanewline
+\isaindent{insert\ x\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\ {\isacharparenleft}}{\isacharbar}\ GT\ {\isasymRightarrow}\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ insert\ x\ r{\isasymrangle}{\isacharparenright}%
+\end{isabelle}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\mbox{Implementation}}
+\high{%
+\begin{isabelle}%
+isin\ {\isasymlangle}{\isasymrangle}\ x\ {\isacharequal}\ False\isasep\isanewline%
+isin\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\ x\ {\isacharequal}\ {\isacharparenleft}{\sf case}\ cmp\ x\ a\ {\sf of}\isanewline
+\isaindent{isin\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\ x\ {\isacharequal}\ {\isacharparenleft}}\ \ LT\ {\isasymRightarrow}\ isin\ l\ x\isanewline
+\isaindent{isin\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\ x\ {\isacharequal}\ {\isacharparenleft}}{\isacharbar}\ EQ\ {\isasymRightarrow}\ True\isanewline
+\isaindent{isin\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\ x\ {\isacharequal}\ {\isacharparenleft}}{\isacharbar}\ GT\ {\isasymRightarrow}\ isin\ r\ x{\isacharparenright}%
+\end{isabelle}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\mbox{Implementation}}
+\high{\isa{delete\ x\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ {\isasymlangle}{\isasymrangle}}}\\\pause
+\high{\isa{delete\ x\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf case}\ cmp\ x\ a\ {\sf of}\isanewline
+\isaindent{{\isacharparenleft}}\ \ LT\ {\isasymRightarrow}\ {\isasymlangle}delete\ x\ l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\isanewline
+\isaindent{{\isacharparenleft}}{\isacharbar}\ EQ\ {\isasymRightarrow}\ {\sf if}\ r\ {\isacharequal}\ {\isasymlangle}{\isasymrangle}\ {\sf then}\ l\isanewline
+\isaindent{{\isacharparenleft}{\isacharbar}\ EQ\ {\isasymRightarrow}\ }{\sf else}\ {\sf let}\ {\isacharparenleft}a{\isacharprime}{\isacharcomma}\ r{\isacharprime}{\isacharparenright}\ {\isacharequal}\ del{\isacharunderscore}min\ r\ {\sf in}\ {\isasymlangle}l{\isacharcomma}\ a{\isacharprime}{\isacharcomma}\ r{\isacharprime}{\isasymrangle}\isanewline
+\isaindent{{\isacharparenleft}}{\isacharbar}\ GT\ {\isasymRightarrow}\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ delete\ x\ r{\isasymrangle}{\isacharparenright}}}
+\pause
+\high{%
+\begin{isabelle}%
+del{\isacharunderscore}min\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf if}\ l\ {\isacharequal}\ {\isasymlangle}{\isasymrangle}\ {\sf then}\ {\isacharparenleft}a{\isacharcomma}\ r{\isacharparenright}\isanewline
+\isaindent{{\isacharparenleft}}{\sf else}\ {\sf let}\ {\isacharparenleft}x{\isacharcomma}\ l{\isacharprime}{\isacharparenright}\ {\isacharequal}\ del{\isacharunderscore}min\ l\ {\sf in}\ {\isacharparenleft}x{\isacharcomma}\ {\isasymlangle}l{\isacharprime}{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}{\isacharparenright}{\isacharparenright}%
+\end{isabelle}}
+\end{frame}
+%-------------------------------------------------------------
+\subsection{Correctness}
+%-------------------------------------------------------------
+\begin{frame}{Why is this implementation correct?}
+\pause
+\begin{tabular}{@ {}lcccc @ {}}
+Because & \isa{empty} & \isa{insert} & \isa{delete} & \isa{isin}\\
+simulate & \isa{{\isacharbraceleft}{\isacharbraceright}} & \isa{{\isasymunion}\ {\isacharbraceleft}{\isachardot}{\isacharbraceright}} & \isa{{\isacharminus}\ {\isacharbraceleft}{\isachardot}{\isacharbraceright}} & \isa{{\isasymin}}
+\end{tabular}
+\begin{center}
+\begin{tabular}{l}
+\pause
+\high{\isa{set{\isacharunderscore}tree\ empty\ {\isacharequal}\ {\isacharbraceleft}{\isacharbraceright}}}\\\pause
+\high{\isa{{\isachardoublequote}set{\isacharunderscore}tree\ {\isacharparenleft}insert\ x\ t{\isacharparenright}\ {\isacharequal}\ set{\isacharunderscore}tree\ t\ {\isasymunion}\ {\isacharbraceleft}x{\isacharbraceright}{\isachardoublequote}}}\\\pause
+\high{\isa{set{\isacharunderscore}tree\ {\isacharparenleft}delete\ x\ t{\isacharparenright}\ {\isacharequal}\ set{\isacharunderscore}tree\ t\ {\isacharminus}\ {\isacharbraceleft}x{\isacharbraceright}}}\\\pause
+\high{\isa{isin\ t\ x\ {\isacharequal}\ {\isacharparenleft}x\ {\isasymin}\ set{\isacharunderscore}tree\ t{\isacharparenright}}}\\
+\end{tabular}
+\end{center}
+\pause
+
+Under the assumption \pause \high{\isa{bst\ t}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Also: \isa{bst} must be invariant}
+\begin{center}
+\begin{tabular}{l}
+\high{\isa{bst\ empty}}\\
+\high{\isa{bst\ t\ {\isasymLongrightarrow}\ bst\ {\isacharparenleft}insert\ x\ t{\isacharparenright}}}\\
+\high{\isa{bst\ t\ {\isasymLongrightarrow}\ bst\ {\isacharparenleft}delete\ x\ t{\isacharparenright}}}
+\end{tabular}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\subsection{Correctness Proof Method Based on Sorted Lists}
+%-------------------------------------------------------------
+\begin{frame}{}
+\isa{sorted\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ bool}\\\smallskip\pause
+\high{\isa{sorted\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ True}}\\\pause
+\high{\isa{sorted\ {\isacharbrackleft}x{\isacharbrackright}\ {\isacharequal}\ True}}\\\pause
+\high{\isa{sorted\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ zs{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y\ {\isasymand}\ sorted\ {\isacharparenleft}y\ {\isacharhash}\ zs{\isacharparenright}{\isacharparenright}}}\pause
+\begin{center}
+No duplicates!
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+%
+%-------------------------------------------------------------
+\begin{frame}{Structural invariant}
+The proof method works not just for unbalanced trees.
+\pause
+
+We assume that there is some structural invariant on the search tree:
+\high{%
+\begin{isabelle}%
+inv\ {\isacharcolon}\ {\isacharprime}s\ {\isasymRightarrow}\ bool%
+\end{isabelle}}
+e.g. some balance criterion.
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of \isa{insert}}
+\high{%
+\begin{isabelle}%
+inv\ t\ {\isasymand}\ sorted\ {\isacharparenleft}inorder\ t{\isacharparenright}\ {\isasymLongrightarrow}\isanewline
+inorder\ {\isacharparenleft}insert\ x\ t{\isacharparenright}\ {\isacharequal}\ ins{\isacharunderscore}list\ x\ {\isacharparenleft}inorder\ t{\isacharparenright}%
+\end{isabelle}}
+\pause
+where
+\high{%
+\begin{isabelle}%
+ins{\isacharunderscore}list\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list%
+\end{isabelle}}
+inserts an element into a sorted list.
+\bigskip\pause
+\begin{center}
+\high{Also covers preservation of \isa{bst}}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of \isa{delete}}
+\high{%
+\begin{isabelle}%
+inv\ t\ {\isasymand}\ sorted\ {\isacharparenleft}inorder\ t{\isacharparenright}\ {\isasymLongrightarrow}\isanewline
+inorder\ {\isacharparenleft}delete\ x\ t{\isacharparenright}\ {\isacharequal}\ del{\isacharunderscore}list\ x\ {\isacharparenleft}inorder\ t{\isacharparenright}%
+\end{isabelle}}
+\pause
+where
+\high{%
+\begin{isabelle}%
+del{\isacharunderscore}list\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list%
+\end{isabelle}}
+deletes an element from a sorted list.
+\bigskip\pause
+\begin{center}
+\high{Also covers preservation of \isa{bst}}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of \isa{isin}}
+\high{%
+\begin{isabelle}%
+inv\ t\ {\isasymand}\ sorted\ {\isacharparenleft}inorder\ t{\isacharparenright}\ {\isasymLongrightarrow}\isanewline
+isin\ t\ x\ {\isacharequal}\ {\isacharparenleft}x\ {\isasymin}\ elems\ {\isacharparenleft}inorder\ t{\isacharparenright}{\isacharparenright}%
+\end{isabelle}}
+\pause
+where
+\high{%
+\begin{isabelle}%
+elems\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ set%
+\end{isabelle}}
+converts a list into a set.
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness via sorted lists}
+\pause
+\begin{center}
+\high{Correctness proofs of all search trees\\
+ covered in this course\\
+can be automated.}
+\pause\medskip
+
+\bad{Except for the structural invariants.}
+\pause\medskip
+
+\bad{Therefore we concentrate on the latter from now on.}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{isabellebody}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isatagtheory
+%
+\endisatagtheory
+{\isafoldtheory}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\end{isabellebody}%
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: "root"
+%%% End:
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/tex/document/Sorting.tex Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,592 @@
+%
+\begin{isabellebody}%
+\setisabellecontext{Sorting}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isatagtheory
+\isacommand{theory}\isamarkupfalse%
+\ Sorting\isanewline
+\isakeyword{imports}\isanewline
+\ \ Complex{\isacharunderscore}Main\isanewline
+\ \ {\isachardoublequoteopen}{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}HOL{\isacharslash}Library{\isacharslash}Multiset{\isachardoublequoteclose}\isanewline
+\isakeyword{begin}%
+\endisatagtheory
+{\isafoldtheory}%
+%
+\isadelimtheory
+\isanewline
+%
+\endisadelimtheory
+\isanewline
+\isacommand{hide{\isacharunderscore}const}\isamarkupfalse%
+\ List{\isachardot}sorted\ List{\isachardot}insort\isanewline
+\isanewline
+\isacommand{declare}\isamarkupfalse%
+\ Let{\isacharunderscore}def\ {\isacharbrackleft}simp{\isacharbrackright}\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ sorted\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isacharcolon}{\isacharcolon}linorder\ list\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
+{\isachardoublequoteopen}sorted\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ True{\isachardoublequoteclose}\ {\isacharbar}\isanewline
+{\isachardoublequoteopen}sorted\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isacharparenleft}{\isasymforall}y{\isasymin}set\ xs{\isachardot}\ x\ {\isasymle}\ y{\isacharparenright}\ {\isacharampersand}\ sorted\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ sorted{\isacharunderscore}append{\isacharcolon}\isanewline
+\ \ {\isachardoublequoteopen}sorted\ {\isacharparenleft}xs{\isacharat}ys{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}sorted\ xs\ {\isacharampersand}\ sorted\ ys\ {\isacharampersand}\ {\isacharparenleft}{\isasymforall}x\ {\isasymin}\ set\ xs{\isachardot}\ {\isasymforall}y\ {\isasymin}\ set\ ys{\isachardot}\ x{\isasymle}y{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}induct\ xs{\isacharparenright}\ {\isacharparenleft}auto{\isacharparenright}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isamarkupsubsection{Insertion Sort%
+}
+\isamarkuptrue%
+\isacommand{fun}\isamarkupfalse%
+\ insort\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isacharcolon}{\isacharcolon}linorder\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
+{\isachardoublequoteopen}insort\ x\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}x{\isacharbrackright}{\isachardoublequoteclose}\ {\isacharbar}\isanewline
+{\isachardoublequoteopen}insort\ x\ {\isacharparenleft}y{\isacharhash}ys{\isacharparenright}\ {\isacharequal}\isanewline
+\ \ {\isacharparenleft}if\ x\ {\isasymle}\ y\ then\ x{\isacharhash}y{\isacharhash}ys\ else\ y{\isacharhash}{\isacharparenleft}insort\ x\ ys{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ isort\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isacharcolon}{\isacharcolon}linorder\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
+{\isachardoublequoteopen}isort\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ {\isacharbar}\isanewline
+{\isachardoublequoteopen}isort\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ insort\ x\ {\isacharparenleft}isort\ xs{\isacharparenright}{\isachardoublequoteclose}%
+\isamarkupsubsubsection{Functional Correctness%
+}
+\isamarkuptrue%
+\isacommand{lemma}\isamarkupfalse%
+\ mset{\isacharunderscore}insort{\isacharcolon}\ {\isachardoublequoteopen}mset\ {\isacharparenleft}insort\ x\ xs{\isacharparenright}\ {\isacharequal}\ add{\isacharunderscore}mset\ x\ {\isacharparenleft}mset\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}induction\ xs{\isacharparenright}\isanewline
+\isacommand{apply}\isamarkupfalse%
+\ auto\isanewline
+\isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ mset{\isacharunderscore}isort{\isacharcolon}\ {\isachardoublequoteopen}mset\ {\isacharparenleft}isort\ xs{\isacharparenright}\ {\isacharequal}\ mset\ xs{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}induction\ xs{\isacharparenright}\isanewline
+\isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}simp\ add{\isacharcolon}\ mset{\isacharunderscore}insort{\isacharparenright}\isanewline
+\isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ {\isachardoublequoteopen}sorted\ {\isacharparenleft}insort\ a\ xs{\isacharparenright}\ {\isacharequal}\ sorted\ xs{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}induction\ xs{\isacharparenright}\isanewline
+\isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}auto{\isacharparenright}\isanewline
+\isacommand{oops}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ set{\isacharunderscore}insort{\isacharcolon}\ {\isachardoublequoteopen}set\ {\isacharparenleft}insort\ x\ xs{\isacharparenright}\ {\isacharequal}\ insert\ x\ {\isacharparenleft}set\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}metis\ mset{\isacharunderscore}insort\ set{\isacharunderscore}mset{\isacharunderscore}add{\isacharunderscore}mset{\isacharunderscore}insert\ set{\isacharunderscore}mset{\isacharunderscore}mset{\isacharparenright}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ set{\isacharunderscore}isort{\isacharcolon}\ {\isachardoublequoteopen}set\ {\isacharparenleft}isort\ xs{\isacharparenright}\ {\isacharequal}\ set\ xs{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}metis\ mset{\isacharunderscore}isort\ set{\isacharunderscore}mset{\isacharunderscore}mset{\isacharparenright}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ sorted{\isacharunderscore}insort{\isacharcolon}\ {\isachardoublequoteopen}sorted\ {\isacharparenleft}insort\ a\ xs{\isacharparenright}\ {\isacharequal}\ sorted\ xs{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}induction\ xs{\isacharparenright}\isanewline
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}auto\ simp\ add{\isacharcolon}\ set{\isacharunderscore}insort{\isacharparenright}\isanewline
+\isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ {\isachardoublequoteopen}sorted\ {\isacharparenleft}isort\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}induction\ xs{\isacharparenright}\isanewline
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}auto\ simp{\isacharcolon}\ sorted{\isacharunderscore}insort{\isacharparenright}\isanewline
+\isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isamarkupsubsection{Time Complexity%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+We count the number of function calls.%
+\end{isamarkuptext}\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+\isa{insort\ x\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}x{\isacharbrackright}}
+\isa{insort\ x\ {\isacharparenleft}y{\isacharhash}ys{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ x\ {\isasymle}\ y\ then\ x{\isacharhash}y{\isacharhash}ys\ else\ y{\isacharhash}{\isacharparenleft}insort\ x\ ys{\isacharparenright}{\isacharparenright}}%
+\end{isamarkuptext}\isamarkuptrue%
+\isacommand{fun}\isamarkupfalse%
+\ t{\isacharunderscore}insort\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isacharcolon}{\isacharcolon}linorder\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
+{\isachardoublequoteopen}t{\isacharunderscore}insort\ x\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\ {\isacharbar}\isanewline
+{\isachardoublequoteopen}t{\isacharunderscore}insort\ x\ {\isacharparenleft}y{\isacharhash}ys{\isacharparenright}\ {\isacharequal}\isanewline
+\ \ {\isacharparenleft}if\ x\ {\isasymle}\ y\ then\ {\isadigit{0}}\ else\ t{\isacharunderscore}insort\ x\ ys{\isacharparenright}\ {\isacharplus}\ {\isadigit{1}}{\isachardoublequoteclose}%
+\begin{isamarkuptext}%
+\isa{isort\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}}
+\isa{isort\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ insort\ x\ {\isacharparenleft}isort\ xs{\isacharparenright}}%
+\end{isamarkuptext}\isamarkuptrue%
+\isacommand{fun}\isamarkupfalse%
+\ t{\isacharunderscore}isort\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isacharcolon}{\isacharcolon}linorder\ list\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
+{\isachardoublequoteopen}t{\isacharunderscore}isort\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\ {\isacharbar}\isanewline
+{\isachardoublequoteopen}t{\isacharunderscore}isort\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ t{\isacharunderscore}isort\ xs\ {\isacharplus}\ t{\isacharunderscore}insort\ x\ {\isacharparenleft}isort\ xs{\isacharparenright}\ {\isacharplus}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ t{\isacharunderscore}insort{\isacharunderscore}length{\isacharcolon}\ {\isachardoublequoteopen}t{\isacharunderscore}insort\ x\ xs\ {\isasymle}\ length\ xs\ {\isacharplus}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}induction\ xs{\isacharparenright}\isanewline
+\isacommand{apply}\isamarkupfalse%
+\ auto\isanewline
+\isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ length{\isacharunderscore}insort{\isacharcolon}\ {\isachardoublequoteopen}length\ {\isacharparenleft}insort\ x\ xs{\isacharparenright}\ {\isacharequal}\ length\ xs\ {\isacharplus}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}induction\ xs{\isacharparenright}\isanewline
+\isacommand{apply}\isamarkupfalse%
+\ auto\isanewline
+\isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ length{\isacharunderscore}isort{\isacharcolon}\ {\isachardoublequoteopen}length\ {\isacharparenleft}isort\ xs{\isacharparenright}\ {\isacharequal}\ length\ xs{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}induction\ xs{\isacharparenright}\isanewline
+\isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}auto\ simp{\isacharcolon}\ length{\isacharunderscore}insort{\isacharparenright}\isanewline
+\isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ t{\isacharunderscore}isort{\isacharunderscore}length{\isacharcolon}\ {\isachardoublequoteopen}t{\isacharunderscore}isort\ xs\ {\isasymle}\ {\isacharparenleft}length\ xs\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}\ {\isacharcircum}\ {\isadigit{2}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{proof}\isamarkupfalse%
+{\isacharparenleft}induction\ xs{\isacharparenright}\isanewline
+\ \ \isacommand{case}\isamarkupfalse%
+\ Nil\ \isacommand{show}\isamarkupfalse%
+\ {\isacharquery}case\ \isacommand{by}\isamarkupfalse%
+\ simp\isanewline
+\isacommand{next}\isamarkupfalse%
+\isanewline
+\ \ \isacommand{case}\isamarkupfalse%
+\ {\isacharparenleft}Cons\ x\ xs{\isacharparenright}\isanewline
+\ \ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}t{\isacharunderscore}isort\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ t{\isacharunderscore}isort\ xs\ {\isacharplus}\ t{\isacharunderscore}insort\ x\ {\isacharparenleft}isort\ xs{\isacharparenright}\ {\isacharplus}\ {\isadigit{1}}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{also}\isamarkupfalse%
+\ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}{\isasymdots}\ {\isasymle}\ {\isacharparenleft}length\ xs\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}\ {\isacharcircum}\ {\isadigit{2}}\ {\isacharplus}\ t{\isacharunderscore}insort\ x\ {\isacharparenleft}isort\ xs{\isacharparenright}\ {\isacharplus}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{using}\isamarkupfalse%
+\ Cons{\isachardot}IH\ \isacommand{by}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{also}\isamarkupfalse%
+\ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}{\isasymdots}\ {\isasymle}\ {\isacharparenleft}length\ xs\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}\ {\isacharcircum}\ {\isadigit{2}}\ {\isacharplus}\ length\ xs\ {\isacharplus}\ {\isadigit{1}}\ {\isacharplus}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{using}\isamarkupfalse%
+\ t{\isacharunderscore}insort{\isacharunderscore}length{\isacharbrackleft}of\ x\ {\isachardoublequoteopen}isort\ xs{\isachardoublequoteclose}{\isacharbrackright}\ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}simp\ add{\isacharcolon}\ length{\isacharunderscore}isort{\isacharparenright}\isanewline
+\ \ \isacommand{also}\isamarkupfalse%
+\ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}{\isasymdots}\ {\isasymle}\ {\isacharparenleft}length{\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}\ {\isacharcircum}\ {\isadigit{2}}{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}simp\ add{\isacharcolon}\ power{\isadigit{2}}{\isacharunderscore}eq{\isacharunderscore}square{\isacharparenright}\isanewline
+\ \ \isacommand{finally}\isamarkupfalse%
+\ \isacommand{show}\isamarkupfalse%
+\ {\isacharquery}case\ \isacommand{{\isachardot}}\isamarkupfalse%
+\isanewline
+\isacommand{qed}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isamarkupsubsection{Merge Sort%
+}
+\isamarkuptrue%
+\isacommand{fun}\isamarkupfalse%
+\ merge\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isacharcolon}{\isacharcolon}linorder\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
+{\isachardoublequoteopen}merge\ {\isacharbrackleft}{\isacharbrackright}\ ys\ {\isacharequal}\ ys{\isachardoublequoteclose}\ {\isacharbar}\isanewline
+{\isachardoublequoteopen}merge\ xs\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ xs{\isachardoublequoteclose}\ {\isacharbar}\isanewline
+{\isachardoublequoteopen}merge\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharparenleft}y{\isacharhash}ys{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ x\ {\isasymle}\ y\ then\ x\ {\isacharhash}\ merge\ xs\ {\isacharparenleft}y{\isacharhash}ys{\isacharparenright}\ else\ y\ {\isacharhash}\ merge\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ msort\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isacharcolon}{\isacharcolon}linorder\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
+{\isachardoublequoteopen}msort\ xs\ {\isacharequal}\ {\isacharparenleft}let\ n\ {\isacharequal}\ length\ xs\ in\isanewline
+\ \ if\ n\ {\isasymle}\ {\isadigit{1}}\ then\ xs\isanewline
+\ \ else\ merge\ {\isacharparenleft}msort\ {\isacharparenleft}take\ {\isacharparenleft}n\ div\ {\isadigit{2}}{\isacharparenright}\ xs{\isacharparenright}{\isacharparenright}\ {\isacharparenleft}msort\ {\isacharparenleft}drop\ {\isacharparenleft}n\ div\ {\isadigit{2}}{\isacharparenright}\ xs{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isacommand{declare}\isamarkupfalse%
+\ msort{\isachardot}simps\ {\isacharbrackleft}simp\ del{\isacharbrackright}\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ c{\isacharunderscore}merge\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isacharcolon}{\isacharcolon}linorder\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
+{\isachardoublequoteopen}c{\isacharunderscore}merge\ {\isacharbrackleft}{\isacharbrackright}\ ys\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\ {\isacharbar}\isanewline
+{\isachardoublequoteopen}c{\isacharunderscore}merge\ xs\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\ {\isacharbar}\isanewline
+{\isachardoublequoteopen}c{\isacharunderscore}merge\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharparenleft}y{\isacharhash}ys{\isacharparenright}\ {\isacharequal}\ {\isadigit{1}}\ {\isacharplus}\ {\isacharparenleft}if\ x\ {\isasymle}\ y\ then\ c{\isacharunderscore}merge\ xs\ {\isacharparenleft}y{\isacharhash}ys{\isacharparenright}\ else\ c{\isacharunderscore}merge\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ c{\isacharunderscore}merge{\isacharunderscore}ub{\isacharcolon}\ {\isachardoublequoteopen}c{\isacharunderscore}merge\ xs\ ys\ {\isasymle}\ length\ xs\ {\isacharplus}\ length\ ys{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}induction\ xs\ ys\ rule{\isacharcolon}\ c{\isacharunderscore}merge{\isachardot}induct{\isacharparenright}\ auto%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ c{\isacharunderscore}msort\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isacharcolon}{\isacharcolon}linorder\ list\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
+{\isachardoublequoteopen}c{\isacharunderscore}msort\ xs\ {\isacharequal}\isanewline
+\ \ {\isacharparenleft}let\ n\ {\isacharequal}\ length\ xs{\isacharsemicolon}\isanewline
+\ \ \ \ \ \ \ ys\ {\isacharequal}\ take\ {\isacharparenleft}n\ div\ {\isadigit{2}}{\isacharparenright}\ xs{\isacharsemicolon}\isanewline
+\ \ \ \ \ \ \ zs\ {\isacharequal}\ drop\ {\isacharparenleft}n\ div\ {\isadigit{2}}{\isacharparenright}\ xs\isanewline
+\ \ \ in\ if\ n\ {\isasymle}\ {\isadigit{1}}\ then\ {\isadigit{0}}\isanewline
+\ \ \ \ \ \ else\ c{\isacharunderscore}msort\ ys\ {\isacharplus}\ c{\isacharunderscore}msort\ zs\ {\isacharplus}\ c{\isacharunderscore}merge\ {\isacharparenleft}msort\ ys{\isacharparenright}\ {\isacharparenleft}msort\ zs{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isacommand{declare}\isamarkupfalse%
+\ c{\isacharunderscore}msort{\isachardot}simps\ {\isacharbrackleft}simp\ del{\isacharbrackright}\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ length{\isacharunderscore}merge{\isacharcolon}\ {\isachardoublequoteopen}length{\isacharparenleft}merge\ xs\ ys{\isacharparenright}\ {\isacharequal}\ length\ xs\ {\isacharplus}\ length\ ys{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}induction\ xs\ ys\ rule{\isacharcolon}\ merge{\isachardot}induct{\isacharparenright}\isanewline
+\isacommand{apply}\isamarkupfalse%
+\ auto\isanewline
+\isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ length{\isacharunderscore}msort{\isacharcolon}\ {\isachardoublequoteopen}length{\isacharparenleft}msort\ xs{\isacharparenright}\ {\isacharequal}\ length\ xs{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{proof}\isamarkupfalse%
+\ {\isacharparenleft}induction\ xs\ rule{\isacharcolon}\ msort{\isachardot}induct{\isacharparenright}\isanewline
+\ \ \isacommand{case}\isamarkupfalse%
+\ {\isacharparenleft}{\isadigit{1}}\ xs{\isacharparenright}\isanewline
+\ \ \isacommand{thus}\isamarkupfalse%
+\ {\isacharquery}case\ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}auto\ simp{\isacharcolon}\ msort{\isachardot}simps{\isacharbrackleft}of\ xs{\isacharbrackright}\ length{\isacharunderscore}merge{\isacharparenright}\isanewline
+\isacommand{qed}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\begin{isamarkuptext}%
+Why structured proof?
+ To have the name "xs" to specialize msort.simps with xs
+ to ensure that msort.simps cannot be used recursively.
+Also works without this precaution, but that is just luck.%
+\end{isamarkuptext}\isamarkuptrue%
+\isacommand{lemma}\isamarkupfalse%
+\ c{\isacharunderscore}msort{\isacharunderscore}le{\isacharcolon}\ {\isachardoublequoteopen}length\ xs\ {\isacharequal}\ {\isadigit{2}}{\isacharcircum}k\ {\isasymLongrightarrow}\ c{\isacharunderscore}msort\ xs\ {\isasymle}\ k\ {\isacharasterisk}\ {\isadigit{2}}{\isacharcircum}k{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{proof}\isamarkupfalse%
+{\isacharparenleft}induction\ k\ arbitrary{\isacharcolon}\ xs{\isacharparenright}\isanewline
+\ \ \isacommand{case}\isamarkupfalse%
+\ {\isadigit{0}}\ \isacommand{thus}\isamarkupfalse%
+\ {\isacharquery}case\ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}simp\ add{\isacharcolon}\ c{\isacharunderscore}msort{\isachardot}simps{\isacharparenright}\isanewline
+\isacommand{next}\isamarkupfalse%
+\isanewline
+\ \ \isacommand{case}\isamarkupfalse%
+\ {\isacharparenleft}Suc\ k{\isacharparenright}\isanewline
+\ \ \isacommand{let}\isamarkupfalse%
+\ {\isacharquery}n\ {\isacharequal}\ {\isachardoublequoteopen}length\ xs{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{let}\isamarkupfalse%
+\ {\isacharquery}ys\ {\isacharequal}\ {\isachardoublequoteopen}take\ {\isacharparenleft}{\isacharquery}n\ div\ {\isadigit{2}}{\isacharparenright}\ xs{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{let}\isamarkupfalse%
+\ {\isacharquery}zs\ {\isacharequal}\ {\isachardoublequoteopen}drop\ {\isacharparenleft}{\isacharquery}n\ div\ {\isadigit{2}}{\isacharparenright}\ xs{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{show}\isamarkupfalse%
+\ {\isacharquery}case\isanewline
+\ \ \isacommand{proof}\isamarkupfalse%
+\ {\isacharparenleft}cases\ {\isachardoublequoteopen}{\isacharquery}n\ {\isasymle}\ {\isadigit{1}}{\isachardoublequoteclose}{\isacharparenright}\isanewline
+\ \ \ \ \isacommand{case}\isamarkupfalse%
+\ True\isanewline
+\ \ \ \ \isacommand{thus}\isamarkupfalse%
+\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
+{\isacharparenleft}simp\ add{\isacharcolon}\ c{\isacharunderscore}msort{\isachardot}simps{\isacharparenright}\isanewline
+\ \ \isacommand{next}\isamarkupfalse%
+\isanewline
+\ \ \ \ \isacommand{case}\isamarkupfalse%
+\ False\isanewline
+\ \ \ \ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}c{\isacharunderscore}msort{\isacharparenleft}xs{\isacharparenright}\ {\isacharequal}\isanewline
+\ \ \ \ \ \ c{\isacharunderscore}msort\ {\isacharquery}ys\ {\isacharplus}\ c{\isacharunderscore}msort\ {\isacharquery}zs\ {\isacharplus}\ c{\isacharunderscore}merge\ {\isacharparenleft}msort\ {\isacharquery}ys{\isacharparenright}\ {\isacharparenleft}msort\ {\isacharquery}zs{\isacharparenright}{\isachardoublequoteclose}\isanewline
+\ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}simp\ add{\isacharcolon}\ c{\isacharunderscore}msort{\isachardot}simps\ msort{\isachardot}simps{\isacharparenright}\isanewline
+\ \ \ \ \isacommand{also}\isamarkupfalse%
+\ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}{\isasymdots}\ {\isasymle}\ c{\isacharunderscore}msort\ {\isacharquery}ys\ {\isacharplus}\ c{\isacharunderscore}msort\ {\isacharquery}zs\ {\isacharplus}\ length\ {\isacharquery}ys\ {\isacharplus}\ length\ {\isacharquery}zs{\isachardoublequoteclose}\isanewline
+\ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
+\ c{\isacharunderscore}merge{\isacharunderscore}ub{\isacharbrackleft}of\ {\isachardoublequoteopen}msort\ {\isacharquery}ys{\isachardoublequoteclose}\ {\isachardoublequoteopen}msort\ {\isacharquery}zs{\isachardoublequoteclose}{\isacharbrackright}\ length{\isacharunderscore}msort{\isacharbrackleft}of\ {\isacharquery}ys{\isacharbrackright}\ length{\isacharunderscore}msort{\isacharbrackleft}of\ {\isacharquery}zs{\isacharbrackright}\isanewline
+\ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
+\ arith\isanewline
+\ \ \ \ \isacommand{also}\isamarkupfalse%
+\ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}{\isasymdots}\ {\isasymle}\ k\ {\isacharasterisk}\ {\isadigit{2}}{\isacharcircum}k\ {\isacharplus}\ c{\isacharunderscore}msort\ {\isacharquery}zs\ {\isacharplus}\ length\ {\isacharquery}ys\ {\isacharplus}\ length\ {\isacharquery}zs{\isachardoublequoteclose}\isanewline
+\ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
+\ Suc{\isachardot}IH{\isacharbrackleft}of\ {\isacharquery}ys{\isacharbrackright}\ Suc{\isachardot}prems\ \isacommand{by}\isamarkupfalse%
+\ simp\isanewline
+\ \ \ \ \isacommand{also}\isamarkupfalse%
+\ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}{\isasymdots}\ {\isasymle}\ k\ {\isacharasterisk}\ {\isadigit{2}}{\isacharcircum}k\ {\isacharplus}\ k\ {\isacharasterisk}\ {\isadigit{2}}{\isacharcircum}k\ {\isacharplus}\ length\ {\isacharquery}ys\ {\isacharplus}\ length\ {\isacharquery}zs{\isachardoublequoteclose}\isanewline
+\ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
+\ Suc{\isachardot}IH{\isacharbrackleft}of\ {\isacharquery}zs{\isacharbrackright}\ Suc{\isachardot}prems\ \isacommand{by}\isamarkupfalse%
+\ simp\isanewline
+\ \ \ \ \isacommand{also}\isamarkupfalse%
+\ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}{\isasymdots}\ {\isacharequal}\ {\isadigit{2}}\ {\isacharasterisk}\ k\ {\isacharasterisk}\ {\isadigit{2}}{\isacharcircum}k\ {\isacharplus}\ {\isadigit{2}}\ {\isacharasterisk}\ {\isadigit{2}}\ {\isacharcircum}\ k{\isachardoublequoteclose}\isanewline
+\ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
+\ Suc{\isachardot}prems\ \isacommand{by}\isamarkupfalse%
+\ simp\isanewline
+\ \ \ \ \isacommand{finally}\isamarkupfalse%
+\ \isacommand{show}\isamarkupfalse%
+\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{qed}\isamarkupfalse%
+\isanewline
+\isacommand{qed}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ {\isachardoublequoteopen}length\ xs\ {\isacharequal}\ {\isadigit{2}}{\isacharcircum}k\ {\isasymLongrightarrow}\ c{\isacharunderscore}msort\ xs\ {\isasymle}\ length\ xs\ {\isacharasterisk}\ log\ {\isadigit{2}}\ {\isacharparenleft}length\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ c{\isacharunderscore}msort{\isacharunderscore}le{\isacharbrackleft}of\ xs\ k{\isacharbrackright}\ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}simp\ add{\isacharcolon}\ log{\isacharunderscore}nat{\isacharunderscore}power\ algebra{\isacharunderscore}simps{\isacharparenright}\isanewline
+\isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}metis\ {\isacharparenleft}mono{\isacharunderscore}tags{\isacharparenright}\ numeral{\isacharunderscore}power{\isacharunderscore}eq{\isacharunderscore}real{\isacharunderscore}of{\isacharunderscore}nat{\isacharunderscore}cancel{\isacharunderscore}iff\ of{\isacharunderscore}nat{\isacharunderscore}le{\isacharunderscore}iff\ of{\isacharunderscore}nat{\isacharunderscore}mult{\isacharparenright}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+%
+\isadelimtheory
+\isanewline
+%
+\endisadelimtheory
+%
+\isatagtheory
+\isacommand{end}\isamarkupfalse%
+%
+\endisatagtheory
+{\isafoldtheory}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\end{isabellebody}%
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: "root"
+%%% End:
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/tex/document/Sorting_Slides.tex Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,252 @@
+%
+\begin{isabellebody}%
+\setisabellecontext{Sorting{\isacharunderscore}Slides}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isatagtheory
+%
+\endisatagtheory
+{\isafoldtheory}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\end{isabellebody}%
+%-------------------------------------------------------------
+\section{Correctness}
+%-------------------------------------------------------------
+\begin{frame}
+\high{\isa{sorted} \isa{{\isacharcolon}{\isacharcolon}} \isa{{\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcolon}{\isacharcolon}linorder{\isacharparenright}\ list\ {\isasymRightarrow}\ bool{\isachardoublequote}}}
+\pause
+\high{%
+\begin{isabelle}%
+sorted\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ True\isasep\isanewline%
+sorted\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isacharparenleft}{\isasymforall}y{\isasymin}set\ xs{\isachardot}\ x\ {\isasymle}\ y{\isacharparenright}\ {\isasymand}\ sorted\ xs{\isacharparenright}%
+\end{isabelle}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of sorting}
+Specification of \isa{sort\ {\isacharcolon}{\isacharcolon}} \isa{{\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcolon}{\isacharcolon}linorder{\isacharparenright}\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequote}}:
+\begin{center}
+\high{\isa{sorted\ {\isacharparenleft}sort\ xs{\isacharparenright}}}
+\end{center}
+\pause
+Is that it? \pause How about
+\begin{center}
+\isa{set\ {\isacharparenleft}sort\ xs{\isacharparenright}\ {\isacharequal}\ set\ xs}
+\end{center}
+\pause
+Better: every \isa{x} occurs as often in \isa{sort\ xs} as in \isa{xs}.
+\pause\medskip
+
+More succinctly:
+\begin{center}
+\high{\isa{mset\ {\isacharparenleft}sort\ xs{\isacharparenright}\ {\isacharequal}\ mset\ xs}}
+\end{center}
+where \high{\isa{mset\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ multiset}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{What are multisets?}
+\pause
+\begin{center}
+\high{Sets with (possibly) repeated elements}
+\end{center}
+\pause
+Some operations:\medskip
+
+\begin{tabular}{rcl}
+\high{\isa{{\isacharbraceleft}{\isacharhash}{\isacharbraceright}}} &\isa{{\isacharcolon}{\isacharcolon}}& \isa{{\isacharprime}a\ multiset}\\
+\high{\isa{add{\isacharunderscore}mset}} &\isa{{\isacharcolon}{\isacharcolon}}& \isa{{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ multiset\ {\isasymRightarrow}\ {\isacharprime}a\ multiset}\\
+\high{\isa{{\isacharplus}}} &\isa{{\isacharcolon}{\isacharcolon}}& \isa{{\isacharprime}a\ multiset\ {\isasymRightarrow}\ {\isacharprime}a\ multiset\ {\isasymRightarrow}\ {\isacharprime}a\ multiset}\\\pause
+\high{\isa{mset}} &\isa{{\isacharcolon}{\isacharcolon}}& \isa{{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ multiset}\\
+\high{\isa{set{\isacharunderscore}mset}} &\isa{{\isacharcolon}{\isacharcolon}}& \isa{{\isacharprime}a\ multiset\ {\isasymRightarrow}\ {\isacharprime}a\ set}\\
+\end{tabular}
+\end{frame}
+%-------------------------------------------------------------
+\section{Insertion Sort}
+%-------------------------------------------------------------
+\thyfile{Thys/Sorting}{Insertion sort correctness}
+%-------------------------------------------------------------
+\section{Time}
+%-------------------------------------------------------------
+\begin{frame}{Principle: Count function calls}
+\pause
+\begin{tabular}{@ {}ll@ {}}
+For every function & \ \isa{f\ {\isacharcolon}{\isacharcolon}\ {\isasymtau}\isactrlsub {\isadigit{1}}\ {\isasymRightarrow}\ {\isachardot}{\isachardot}{\isachardot}\ {\isasymRightarrow}\ {\isasymtau}\isactrlsub n\ {\isasymRightarrow}\ {\isasymtau}}\\
+define a \concept{timing function} & \high{\isa{t{\isacharunderscore}f\ {\isacharcolon}{\isacharcolon}\ {\isasymtau}\isactrlsub {\isadigit{1}}\ {\isasymRightarrow}\ {\isachardot}{\isachardot}{\isachardot}\ {\isasymRightarrow}\ {\isasymtau}\isactrlsub n\ {\isasymRightarrow}\ nat}}:
+\end{tabular}
+\smallskip\pause
+Translation of defining equations:
+\[
+\frac{e \leadsto \high{e'}}{f\,p_1 \dots p_n = e ~\leadsto~ \high{t\_f\,p_1 \dots p_n = e' + 1}}
+\]
+\pause
+
+Translation of expressions:
+\[ \frac{s_1 \leadsto \high{t_1} \quad \dots \quad s_k \leadsto \high{t_k}}
+ {g\,s_1 \dots s_k \leadsto \high{t_1 + \dots + t_k + t\_g\, s_1 \dots s_k}}
+\]
+\pause
+\begin{itemize}
+\item
+Variable $\leadsto$ 0, Constant $\leadsto$ 0
+\pause
+\item
+Constructor calls and primitive operations on \isa{bool} and numbers cost 1
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Example}
+\isa{app\ {\isacharbrackleft}{\isacharbrackright}\ ys\ {\isacharequal}\ ys}\\
+\pause$\leadsto$\\
+\high{\isa{t{\isacharunderscore}app\ {\isacharbrackleft}{\isacharbrackright}\ ys\ {\isacharequal}\ {\isadigit{0}}\ {\isacharplus}\ {\isadigit{1}}}}\\
+\pause\bigskip
+
+\isa{app\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ ys\ {\isacharequal}\ x\ {\isacharhash}\ app\ xs\ ys}\\
+\pause$\leadsto$\\
+\high{\isa{t{\isacharunderscore}app\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ ys\ {\isacharequal}\ {\isadigit{0}}\ {\isacharplus}\ {\isacharparenleft}{\isadigit{0}}\ {\isacharplus}\ {\isadigit{0}}\ {\isacharplus}\ t{\isacharunderscore}app\ xs\ ys{\isacharparenright}\ {\isacharplus}\ {\isadigit{1}}\ {\isacharplus}\ {\isadigit{1}}}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{A compact formulation of $e \leadsto t$}
+\pause
+\begin{center}
+\begin{tabular}{l}
+\high{\isa{t}} \ is the sum of all \high{\isa{t{\isacharunderscore}g\ s\isactrlsub {\isadigit{1}}\ {\isachardot}{\isachardot}{\isachardot}\ s\isactrlsub k}}\\
+such that \isa{g\ s\isactrlsub {\isadigit{1}}\ {\isachardot}{\isachardot}{\isachardot}\ s\isactrlsub n} is a subterm of \isa{e}
+\end{tabular}
+\end{center}
+\pause\smallskip
+
+If $g$ is
+\begin{itemize}
+\item a constructor or
+\item a predefined function on \isa{bool} or numbers
+\end{itemize}
+then \isa{t{\isacharunderscore}g\ {\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isadigit{1}}}.
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\emph{if} and \emph{case}}
+\pause
+\begin{center}
+So far we model a call-by-value semantics
+\end{center}
+\pause
+Conditionals and case expressions are evaluated \bad{lazily}.
+\smallskip\pause
+
+Translation:
+\[ \frac{b \leadsto \high{t} \quad s_1 \leadsto \high{t_1} \quad s_2 \leadsto \high{t_2}}
+ {\textit{if } b \textit{ then } s_1 \textit{ else } s_2 \leadsto \high{t + (\textit{if } b \textit{ then } t_1 \textit{ else } t_2)}}
+\]
+\pause
+Similarly for \emph{case}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{$O(.)$ is enough}
+\pause
+$\Longrightarrow$ Reduce all additive constants to 1
+\pause
+\begin{example}
+\high{\isa{t{\isacharunderscore}app\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ ys\ {\isacharequal}\ t{\isacharunderscore}app\ xs\ ys\ {\isacharplus}\ {\isadigit{1}}}}
+\end{example}
+%\bigskip\pause
+
+%$\Longrightarrow$ May start with 0 instead of 1 in base case
+%\pause
+%\begin{example}
+%\high{\isa{t{\isacharunderscore}app\ {\isacharbrackleft}{\isacharbrackright}\ ys\ {\isacharequal}\ {\isadigit{0}}}}
+%\end{example}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Discussion}
+\begin{itemize}
+\pause
+\item
+The definition of \isa{t{\isacharunderscore}f} from \isa{f} can be automated.
+\pause
+\item
+The correctness of \isa{t{\isacharunderscore}f} could be proved w.r.t.\\ a semantics
+that counts computation steps.
+\pause
+\item
+Precise complexity bounds (as opposed to $O(.)$)
+would require a formal model of (at least) the compiler and the hardware.
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\thyfile{Thys/Sorting}{Insertion sort complexity}
+%-------------------------------------------------------------
+\section{Merge Sort}
+%-------------------------------------------------------------
+\begin{frame}
+\isa{merge\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list}\smallskip\\
+\pause
+\high{\isa{merge\ {\isacharbrackleft}{\isacharbrackright}\ ys\ {\isacharequal}\ ys}}\\
+\high{\isa{merge\ xs\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ xs}}\\
+\pause
+\high{\isa{merge\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf if}\ x\ {\isasymle}\ y\ {\sf then}\ x\ {\isacharhash}\ merge\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}\isanewline
+\isaindent{{\isacharparenleft}}{\sf else}\ y\ {\isacharhash}\ merge\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}\ ys{\isacharparenright}}}
+\bigskip\pause
+
+\isa{msort\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list}\smallskip\\
+\pause
+\high{\isa{msort\ xs\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf let}\ n\ {\isacharequal}\ length\ xs\isanewline
+\isaindent{{\isacharparenleft}}{\sf in}\ {\sf if}\ n\ {\isasymle}\ {\isadigit{1}}\ {\sf then}\ xs\isanewline
+\isaindent{{\isacharparenleft}{\sf in}\ }{\sf else}\ merge\ {\isacharparenleft}msort\ {\isacharparenleft}take\ {\isacharparenleft}n\ div\ {\isadigit{2}}{\isacharparenright}\ xs{\isacharparenright}{\isacharparenright}\isanewline
+\isaindent{{\isacharparenleft}{\sf in}\ {\sf else}\ \ }{\isacharparenleft}msort\ {\isacharparenleft}drop\ {\isacharparenleft}n\ div\ {\isadigit{2}}{\isacharparenright}\ xs{\isacharparenright}{\isacharparenright}{\isacharparenright}}}
+\end{frame}
+%-------------------------------------------------------------
+\thyfile{Thys/Sorting}{Merge sort}
+%-------------------------------------------------------------
+\begin{isabellebody}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isatagtheory
+%
+\endisatagtheory
+{\isafoldtheory}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\end{isabellebody}%
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: "root"
+%%% End:
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/tex/document/Tree23_Slides.tex Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,314 @@
+%
+\begin{isabellebody}%
+\setisabellecontext{Tree{\isadigit{2}}{\isadigit{3}}{\isacharunderscore}Slides}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isatagtheory
+%
+\endisatagtheory
+{\isafoldtheory}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\end{isabellebody}%
+%-------------------------------------------------------------
+\section{2-3 Trees}
+%-------------------------------------------------------------
+\thyfile{HOL/Data\char`_Structures/Tree23\char`_Set}{}
+%-------------------------------------------------------------
+\begin{frame}{2-3 Trees}
+\high{%
+\begin{isabelle}%
+\isacommand{datatype}\ {\isacharprime}a\ tree{\isadigit{2}}{\isadigit{3}}\ {\isacharequal}\ {\isasymlangle}{\isasymrangle}\isanewline
+\isaindent{\ \ }{\isacharbar}\ Node{\isadigit{2}}\ {\isacharparenleft}{\isacharprime}a\ tree{\isadigit{2}}{\isadigit{3}}{\isacharparenright}\ {\isacharprime}a\ {\isacharparenleft}{\isacharprime}a\ tree{\isadigit{2}}{\isadigit{3}}{\isacharparenright}\isanewline
+\isaindent{\ \ }{\isacharbar}\ Node{\isadigit{3}}\ {\isacharparenleft}{\isacharprime}a\ tree{\isadigit{2}}{\isadigit{3}}{\isacharparenright}\ {\isacharprime}a\ {\isacharparenleft}{\isacharprime}a\ tree{\isadigit{2}}{\isadigit{3}}{\isacharparenright}\ {\isacharprime}a\ {\isacharparenleft}{\isacharprime}a\ tree{\isadigit{2}}{\isadigit{3}}{\isacharparenright}%
+\end{isabelle}}
+\pause
+
+Abbreviations:
+\smallskip
+\begin{center}
+\begin{tabular}{rcl}
+\high{\isa{{\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}}} &\isa{{\isasymequiv}}& \isa{Node{\isadigit{2}}\ l\ a\ r}\\\pause
+\high{\isa{{\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ m{\isacharcomma}\ b{\isacharcomma}\ r{\isasymrangle}}} &\isa{{\isasymequiv}}& \isa{Node{\isadigit{3}}\ l\ a\ m\ b\ r}
+\end{tabular}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\isa{isin}}
+\high{%
+\begin{isabelle}%
+isin\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ m{\isacharcomma}\ b{\isacharcomma}\ r{\isasymrangle}\ x\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf case}\ cmp\ x\ a\ {\sf of}\isanewline
+\isaindent{{\isacharparenleft}}\ \ LT\ {\isasymRightarrow}\ isin\ l\ x\isanewline
+\isaindent{{\isacharparenleft}}{\isacharbar}\ EQ\ {\isasymRightarrow}\ True\isanewline
+\isaindent{{\isacharparenleft}}{\isacharbar}\ GT\ {\isasymRightarrow}\ {\sf case}\ cmp\ x\ b\ {\sf of}\isanewline
+\isaindent{{\isacharparenleft}{\isacharbar}\ GT\ {\isasymRightarrow}\ }\ \ LT\ {\isasymRightarrow}\ isin\ m\ x\isanewline
+\isaindent{{\isacharparenleft}{\isacharbar}\ GT\ {\isasymRightarrow}\ }{\isacharbar}\ EQ\ {\isasymRightarrow}\ True\isanewline
+\isaindent{{\isacharparenleft}{\isacharbar}\ GT\ {\isasymRightarrow}\ }{\isacharbar}\ GT\ {\isasymRightarrow}\ isin\ r\ x{\isacharparenright}%
+\end{isabelle}}
+\pause
+
+Assumes the usual ordering invariant
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Invariant \isa{bal}}
+All leaves are at the same level:
+\pause
+\high{%
+\begin{isabelle}%
+bal\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ True%
+\end{isabelle}}
+\pause
+\high{\isa{bal\ {\isasymlangle}l{\isacharcomma}\ \_{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\ {\isacharparenleft}bal\ l\ {\isasymand}\ bal\ r\ {\isasymand}\ h{\isacharparenleft}l{\isacharparenright}\ {\isacharequal}\ h{\isacharparenleft}r{\isacharparenright}{\isacharparenright}}}
+\pause
+\high{%
+\begin{isabelle}%
+bal\ {\isasymlangle}l{\isacharcomma}\ \_{\isacharcomma}\ m{\isacharcomma}\ \_{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\isanewline
+{\isacharparenleft}bal\ l\ {\isasymand}\ bal\ m\ {\isasymand}\ bal\ r\ {\isasymand}\ h{\isacharparenleft}l{\isacharparenright}\ {\isacharequal}\ h{\isacharparenleft}m{\isacharparenright}\ {\isasymand}\ h{\isacharparenleft}m{\isacharparenright}\ {\isacharequal}\ h{\isacharparenleft}r{\isacharparenright}{\isacharparenright}%
+\end{isabelle}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Insertion}
+The idea:
+\begin{center}
+\begin{tabular}{rcl}
+\pause
+\isa{Leaf} & $\leadsto$ & \isa{Node{\isadigit{2}}} \\
+\pause
+\isa{Node{\isadigit{2}}} & $\leadsto$ & \isa{Node{\isadigit{3}}} \\
+\pause
+\isa{Node{\isadigit{3}}} & $\leadsto$ & \bad{overflow}, pass 1 element back up \\
+\end{tabular}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Insertion}
+Two possible return values:
+\pause
+\begin{itemize}
+\item tree accommodates new element\\ without increasing height:
+ \high{\isa{T\isactrlsub i\ t}}
+\pause
+\item tree overflows: \bad{\isa{Up\isactrlsub i\ l\ x\ r}}
+\end{itemize}
+\pause
+%
+\begin{isabelle}%
+\isacommand{datatype}\ {\isacharprime}a\ up\isactrlsub i\ {\isacharequal}\ T\isactrlsub i\ {\isacharparenleft}{\isacharprime}a\ tree{\isadigit{2}}{\isadigit{3}}{\isacharparenright}\isanewline
+\isaindent{\ \ }{\isacharbar}\ Up\isactrlsub i\ {\isacharparenleft}{\isacharprime}a\ tree{\isadigit{2}}{\isadigit{3}}{\isacharparenright}\ {\isacharprime}a\ {\isacharparenleft}{\isacharprime}a\ tree{\isadigit{2}}{\isadigit{3}}{\isacharparenright}%
+\end{isabelle}
+\pause
+\isa{tree\isactrlsub i\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ up\isactrlsub i\ {\isasymRightarrow}\ {\isacharprime}a\ tree{\isadigit{2}}{\isadigit{3}}}\\
+\pause\smallskip
+\isa{tree\isactrlsub i\ {\isacharparenleft}T\isactrlsub i\ t{\isacharparenright}\ {\isacharequal}\ t\isasep\isanewline%
+tree\isactrlsub i\ {\isacharparenleft}Up\isactrlsub i\ l\ a\ r{\isacharparenright}\ {\isacharequal}\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Insertion}
+\isa{insert\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ tree{\isadigit{2}}{\isadigit{3}}\ {\isasymRightarrow}\ {\isacharprime}a\ tree{\isadigit{2}}{\isadigit{3}}}\\\smallskip\pause
+\high{\isa{insert\ x\ t\ {\isacharequal}\ tree\isactrlsub i\ {\isacharparenleft}ins\ x\ t{\isacharparenright}}}
+\bigskip\pause
+
+\isa{ins\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ tree{\isadigit{2}}{\isadigit{3}}\ {\isasymRightarrow}\ {\isacharprime}a\ up\isactrlsub i}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Insertion}
+\high{\isa{ins\ x\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ Up\isactrlsub i\ {\isasymlangle}{\isasymrangle}\ x\ {\isasymlangle}{\isasymrangle}}}
+\smallskip\pause
+
+\high{\isa{ins\ x\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}} \isa{{\isacharequal}}}\\\pause
+\high{\isa{{\sf case}\ cmp\ x\ a\ {\sf of}\isanewline
+\ \ LT\ {\isasymRightarrow}\ {\sf case}\ ins\ x\ l\ {\sf of}\isanewline
+\isaindent{\ \ LT\ {\isasymRightarrow}\ }\ \ T\isactrlsub i\ l{\isacharprime}\ {\isasymRightarrow}\ T\isactrlsub i\ {\isasymlangle}l{\isacharprime}{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\isanewline
+\isaindent{\ \ LT\ {\isasymRightarrow}\ }{\isacharbar}\ Up\isactrlsub i\ l\isactrlsub {\isadigit{1}}\ b\ l\isactrlsub {\isadigit{2}}\ {\isasymRightarrow}\ T\isactrlsub i\ {\isasymlangle}l\isactrlsub {\isadigit{1}}{\isacharcomma}\ b{\isacharcomma}\ l\isactrlsub {\isadigit{2}}{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\isanewline
+{\isacharbar}\ EQ\ {\isasymRightarrow}\ T\isactrlsub i\ {\isasymlangle}l{\isacharcomma}\ x{\isacharcomma}\ r{\isasymrangle}\isanewline
+{\isacharbar}\ GT\ {\isasymRightarrow}\ {\sf case}\ ins\ x\ r\ {\sf of}\isanewline
+\isaindent{{\isacharbar}\ GT\ {\isasymRightarrow}\ }\ \ T\isactrlsub i\ r{\isacharprime}\ {\isasymRightarrow}\ T\isactrlsub i\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ r{\isacharprime}{\isasymrangle}\isanewline
+\isaindent{{\isacharbar}\ GT\ {\isasymRightarrow}\ }{\isacharbar}\ Up\isactrlsub i\ r\isactrlsub {\isadigit{1}}\ b\ r\isactrlsub {\isadigit{2}}\ {\isasymRightarrow}\ T\isactrlsub i\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ r\isactrlsub {\isadigit{1}}{\isacharcomma}\ b{\isacharcomma}\ r\isactrlsub {\isadigit{2}}{\isasymrangle}}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Insertion}
+\high{\isa{ins\ x\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ m{\isacharcomma}\ b{\isacharcomma}\ r{\isasymrangle}} \isa{{\isacharequal}}}\\\pause
+\high{\isa{{\sf case}\ cmp\ x\ a\ {\sf of}\isanewline
+\ \ LT\ {\isasymRightarrow}\ {\sf case}\ ins\ x\ l\ {\sf of}\isanewline
+\isaindent{\ \ LT\ {\isasymRightarrow}\ }\ \ T\isactrlsub i\ l{\isacharprime}\ {\isasymRightarrow}\ T\isactrlsub i\ {\isasymlangle}l{\isacharprime}{\isacharcomma}\ a{\isacharcomma}\ m{\isacharcomma}\ b{\isacharcomma}\ r{\isasymrangle}\isanewline
+\isaindent{\ \ LT\ {\isasymRightarrow}\ }{\isacharbar}\ Up\isactrlsub i\ l\isactrlsub {\isadigit{1}}\ c\ l\isactrlsub {\isadigit{2}}\ {\isasymRightarrow}\ Up\isactrlsub i\ {\isasymlangle}l\isactrlsub {\isadigit{1}}{\isacharcomma}\ c{\isacharcomma}\ l\isactrlsub {\isadigit{2}}{\isasymrangle}\ a\ {\isasymlangle}m{\isacharcomma}\ b{\isacharcomma}\ r{\isasymrangle}\isanewline
+{\isacharbar}\ EQ\ {\isasymRightarrow}\ T\isactrlsub i\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ m{\isacharcomma}\ b{\isacharcomma}\ r{\isasymrangle}\isanewline
+{\isacharbar}\ GT\ {\isasymRightarrow}\isanewline
+\isaindent{{\isacharbar}\ \ \ }{\sf case}\ cmp\ x\ b\ {\sf of}\isanewline
+\isaindent{{\isacharbar}\ \ \ }\ \ LT\ {\isasymRightarrow}\isanewline
+\isaindent{{\isacharbar}\ \ \ \ \ \ \ }{\sf case}\ ins\ x\ m\ {\sf of}\isanewline
+\isaindent{{\isacharbar}\ \ \ \ \ \ \ }\ \ T\isactrlsub i\ m{\isacharprime}\ {\isasymRightarrow}\ T\isactrlsub i\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ m{\isacharprime}{\isacharcomma}\ b{\isacharcomma}\ r{\isasymrangle}\isanewline
+\isaindent{{\isacharbar}\ \ \ \ \ \ \ }{\isacharbar}\ Up\isactrlsub i\ m\isactrlsub {\isadigit{1}}\ c\ m\isactrlsub {\isadigit{2}}\ {\isasymRightarrow}\ Up\isactrlsub i\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ m\isactrlsub {\isadigit{1}}{\isasymrangle}\ c\ {\isasymlangle}m\isactrlsub {\isadigit{2}}{\isacharcomma}\ b{\isacharcomma}\ r{\isasymrangle}\isanewline
+\isaindent{{\isacharbar}\ \ \ }{\isacharbar}\ EQ\ {\isasymRightarrow}\ T\isactrlsub i\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ m{\isacharcomma}\ b{\isacharcomma}\ r{\isasymrangle}\isanewline
+\isaindent{{\isacharbar}\ \ \ }{\isacharbar}\ GT\ {\isasymRightarrow}\isanewline
+\isaindent{{\isacharbar}\ \ \ {\isacharbar}\ \ \ }{\sf case}\ ins\ x\ r\ {\sf of}\isanewline
+\isaindent{{\isacharbar}\ \ \ {\isacharbar}\ \ \ }\ \ T\isactrlsub i\ r{\isacharprime}\ {\isasymRightarrow}\ T\isactrlsub i\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ m{\isacharcomma}\ b{\isacharcomma}\ r{\isacharprime}{\isasymrangle}\isanewline
+\isaindent{{\isacharbar}\ \ \ {\isacharbar}\ \ \ }{\isacharbar}\ Up\isactrlsub i\ r\isactrlsub {\isadigit{1}}\ c\ r\isactrlsub {\isadigit{2}}\ {\isasymRightarrow}\ Up\isactrlsub i\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ m{\isasymrangle}\ b\ {\isasymlangle}r\isactrlsub {\isadigit{1}}{\isacharcomma}\ c{\isacharcomma}\ r\isactrlsub {\isadigit{2}}{\isasymrangle}}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Insertion preserves \isa{bal}}
+\begin{lemma}
+\high{\alt<1-4>{\isa{bal\ t\ {\isasymLongrightarrow}\ bal\ {\isacharparenleft}tree\isactrlsub i\ {\isacharparenleft}ins\ a\ t{\isacharparenright}{\isacharparenright}}}{\isa{bal\ t\ {\isasymLongrightarrow}\ bal\ {\isacharparenleft}tree\isactrlsub i\ {\isacharparenleft}ins\ a\ t{\isacharparenright}{\isacharparenright}\ {\isasymand}\ h{\isacharparenleft}ins\ a\ t{\isacharparenright}\ {\isacharequal}\ h{\isacharparenleft}t{\isacharparenright}}}}
+\smallskip
+
+\onslide<3->
+where \isa{h\ {\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}a\ up\isactrlsub i\ {\isasymRightarrow}\ nat}\\
+\onslide<4->
+\isa{h{\isacharparenleft}T\isactrlsub i\ t{\isacharparenright}\ {\isacharequal}\ h{\isacharparenleft}t{\isacharparenright}\isasep\isanewline%
+h{\isacharparenleft}Up\isactrlsub i\ l\ a\ r{\isacharparenright}\ {\isacharequal}\ h{\isacharparenleft}l{\isacharparenright}}
+\end{lemma}
+\onslide<2->
+\textbf{Proof} by induction on \isa{t}
+\bigskip\onslide<6->
+\begin{corollary}
+\high{\isa{bal\ t\ {\isasymLongrightarrow}\ bal\ {\isacharparenleft}insert\ a\ t{\isacharparenright}}}
+\end{corollary}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Deletion}
+The idea:
+\begin{center}
+\begin{tabular}{rcl}
+\pause
+\isa{Node{\isadigit{3}}} & $\leadsto$ & \isa{Node{\isadigit{2}}} \\
+\pause
+\isa{Node{\isadigit{2}}} & $\leadsto$ & \bad{underflow}, height -1
+\end{tabular}
+\end{center}
+\pause
+Underflow: merge with siblings on the way up
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Deletion}
+Two possible return values:
+\begin{itemize}
+\pause
+\item height unchanged: \high{\isa{T\isactrlsub d\ t}}
+\pause
+\item height decreased by 1: \bad{\isa{Up\isactrlsub d\ t}}
+\end{itemize}
+\pause
+%
+\begin{isabelle}%
+\isacommand{datatype}\ {\isacharprime}a\ up\isactrlsub d\ {\isacharequal}\ T\isactrlsub d\ {\isacharparenleft}{\isacharprime}a\ tree{\isadigit{2}}{\isadigit{3}}{\isacharparenright}\ {\isacharbar}\ Up\isactrlsub d\ {\isacharparenleft}{\isacharprime}a\ tree{\isadigit{2}}{\isadigit{3}}{\isacharparenright}%
+\end{isabelle}
+\pause
+%
+\begin{isabelle}%
+tree\isactrlsub d\ {\isacharparenleft}T\isactrlsub d\ t{\isacharparenright}\ {\isacharequal}\ t\isasep\isanewline%
+tree\isactrlsub d\ {\isacharparenleft}Up\isactrlsub d\ t{\isacharparenright}\ {\isacharequal}\ t%
+\end{isabelle}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Deletion}
+\isa{delete\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ tree{\isadigit{2}}{\isadigit{3}}\ {\isasymRightarrow}\ {\isacharprime}a\ tree{\isadigit{2}}{\isadigit{3}}}\\\smallskip\pause
+\high{\isa{delete\ x\ t\ {\isacharequal}\ tree\isactrlsub d\ {\isacharparenleft}del\ x\ t{\isacharparenright}}}
+\bigskip\pause
+
+\isa{del\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ tree{\isadigit{2}}{\isadigit{3}}\ {\isasymRightarrow}\ {\isacharprime}a\ up\isactrlsub d}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Deletion}
+\isa{del\ x\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ T\isactrlsub d\ {\isasymlangle}{\isasymrangle}}
+\smallskip\pause
+
+\isa{del\ x\ {\isasymlangle}{\isasymlangle}{\isasymrangle}{\isacharcomma}\ a{\isacharcomma}\ {\isasymlangle}{\isasymrangle}{\isasymrangle}\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf if}\ x\ {\isacharequal}\ a\ {\sf then}\ Up\isactrlsub d\ {\isasymlangle}{\isasymrangle}\ {\sf else}\ T\isactrlsub d\ {\isasymlangle}{\isasymlangle}{\isasymrangle}{\isacharcomma}\ a{\isacharcomma}\ {\isasymlangle}{\isasymrangle}{\isasymrangle}{\isacharparenright}}
+\smallskip\pause
+
+\isa{del\ x\ {\isasymlangle}{\isasymlangle}{\isasymrangle}{\isacharcomma}\ a{\isacharcomma}\ {\isasymlangle}{\isasymrangle}{\isacharcomma}\ b{\isacharcomma}\ {\isasymlangle}{\isasymrangle}{\isasymrangle}} \isa{{\isacharequal}\ {\isachardot}{\isachardot}{\isachardot}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}
+\high{%
+\begin{isabelle}%
+del\ x\ {\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf case}\ cmp\ x\ a\ {\sf of}\isanewline
+\isaindent{{\isacharparenleft}}\ \ LT\ {\isasymRightarrow}\ node{\isadigit{2}}{\isadigit{1}}\ {\isacharparenleft}del\ x\ l{\isacharparenright}\ a\ r\isanewline
+\isaindent{{\isacharparenleft}}{\isacharbar}\ EQ\ {\isasymRightarrow}\ {\sf let}\ {\isacharparenleft}a{\isacharprime}{\isacharcomma}\ t{\isacharparenright}\ {\isacharequal}\ del{\isacharunderscore}min\ r\ {\sf in}\ node{\isadigit{2}}{\isadigit{2}}\ l\ a{\isacharprime}\ t\isanewline
+\isaindent{{\isacharparenleft}}{\isacharbar}\ GT\ {\isasymRightarrow}\ node{\isadigit{2}}{\isadigit{2}}\ l\ a\ {\isacharparenleft}del\ x\ r{\isacharparenright}{\isacharparenright}%
+\end{isabelle}}
+\pause
+%
+\begin{isabelle}%
+node{\isadigit{2}}{\isadigit{1}}\ {\isacharparenleft}T\isactrlsub d\ t\isactrlsub {\isadigit{1}}{\isacharparenright}\ a\ t\isactrlsub {\isadigit{2}}\ {\isacharequal}\ T\isactrlsub d\ {\isasymlangle}t\isactrlsub {\isadigit{1}}{\isacharcomma}\ a{\isacharcomma}\ t\isactrlsub {\isadigit{2}}{\isasymrangle}\isasep\isanewline%
+node{\isadigit{2}}{\isadigit{1}}\ {\isacharparenleft}Up\isactrlsub d\ t\isactrlsub {\isadigit{1}}{\isacharparenright}\ a\ {\isasymlangle}t\isactrlsub {\isadigit{2}}{\isacharcomma}\ b{\isacharcomma}\ t\isactrlsub {\isadigit{3}}{\isasymrangle}\ {\isacharequal}\ Up\isactrlsub d\ {\isasymlangle}t\isactrlsub {\isadigit{1}}{\isacharcomma}\ a{\isacharcomma}\ t\isactrlsub {\isadigit{2}}{\isacharcomma}\ b{\isacharcomma}\ t\isactrlsub {\isadigit{3}}{\isasymrangle}\isasep\isanewline%
+node{\isadigit{2}}{\isadigit{1}}\ {\isacharparenleft}Up\isactrlsub d\ t\isactrlsub {\isadigit{1}}{\isacharparenright}\ a\ {\isasymlangle}t\isactrlsub {\isadigit{2}}{\isacharcomma}\ b{\isacharcomma}\ t\isactrlsub {\isadigit{3}}{\isacharcomma}\ c{\isacharcomma}\ t\isactrlsub {\isadigit{4}}{\isasymrangle}\ {\isacharequal}\isanewline
+T\isactrlsub d\ {\isasymlangle}{\isasymlangle}t\isactrlsub {\isadigit{1}}{\isacharcomma}\ a{\isacharcomma}\ t\isactrlsub {\isadigit{2}}{\isasymrangle}{\isacharcomma}\ b{\isacharcomma}\ {\isasymlangle}t\isactrlsub {\isadigit{3}}{\isacharcomma}\ c{\isacharcomma}\ t\isactrlsub {\isadigit{4}}{\isasymrangle}{\isasymrangle}%
+\end{isabelle}
+\pause
+%
+\begin{isabelle}%
+node{\isadigit{2}}{\isadigit{2}}\ t\isactrlsub {\isadigit{1}}\ a\ {\isacharparenleft}T\isactrlsub d\ t\isactrlsub {\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ T\isactrlsub d\ {\isasymlangle}t\isactrlsub {\isadigit{1}}{\isacharcomma}\ a{\isacharcomma}\ t\isactrlsub {\isadigit{2}}{\isasymrangle}\isasep\isanewline%
+node{\isadigit{2}}{\isadigit{2}}\ {\isasymlangle}t\isactrlsub {\isadigit{1}}{\isacharcomma}\ b{\isacharcomma}\ t\isactrlsub {\isadigit{2}}{\isasymrangle}\ a\ {\isacharparenleft}Up\isactrlsub d\ t\isactrlsub {\isadigit{3}}{\isacharparenright}\ {\isacharequal}\ Up\isactrlsub d\ {\isasymlangle}t\isactrlsub {\isadigit{1}}{\isacharcomma}\ b{\isacharcomma}\ t\isactrlsub {\isadigit{2}}{\isacharcomma}\ a{\isacharcomma}\ t\isactrlsub {\isadigit{3}}{\isasymrangle}\isasep\isanewline%
+node{\isadigit{2}}{\isadigit{2}}\ {\isasymlangle}t\isactrlsub {\isadigit{1}}{\isacharcomma}\ b{\isacharcomma}\ t\isactrlsub {\isadigit{2}}{\isacharcomma}\ c{\isacharcomma}\ t\isactrlsub {\isadigit{3}}{\isasymrangle}\ a\ {\isacharparenleft}Up\isactrlsub d\ t\isactrlsub {\isadigit{4}}{\isacharparenright}\ {\isacharequal}\isanewline
+T\isactrlsub d\ {\isasymlangle}{\isasymlangle}t\isactrlsub {\isadigit{1}}{\isacharcomma}\ b{\isacharcomma}\ t\isactrlsub {\isadigit{2}}{\isasymrangle}{\isacharcomma}\ c{\isacharcomma}\ {\isasymlangle}t\isactrlsub {\isadigit{3}}{\isacharcomma}\ a{\isacharcomma}\ t\isactrlsub {\isadigit{4}}{\isasymrangle}{\isasymrangle}%
+\end{isabelle}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Deletion preserves \isa{bal}}
+\begin{lemma}
+\high{\isa{bal\ t\ {\isasymLongrightarrow}\ bal\ {\isacharparenleft}tree\isactrlsub d\ {\isacharparenleft}del\ x\ t{\isacharparenright}{\isacharparenright}}}
+\end{lemma}
+
+\begin{corollary}
+\high{\isa{bal\ t\ {\isasymLongrightarrow}\ bal\ {\isacharparenleft}delete\ x\ t{\isacharparenright}}}
+\end{corollary}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Beyond 2-3 trees}
+\pause
+\high{\isakeyword{datatype} \isa{{\isacharprime}a\ tree{\isadigit{2}}{\isadigit{3}}{\isadigit{4}}\ {\isacharequal}}\\
+\quad \isa{Leaf\ {\isacharbar}\ Node{\isadigit{2}}\ {\isachardot}{\isachardot}{\isachardot}\ {\isacharbar}\ Node{\isadigit{3}}\ {\isachardot}{\isachardot}{\isachardot}\ {\isacharbar}\ Node{\isadigit{4}}\ {\isachardot}{\isachardot}{\isachardot}}}
+\pause
+
+\begin{center}
+Like 2-3 tress, but with many more cases
+\end{center}
+\pause
+The general case:
+\begin{center}
+ B-trees and (a,b) trees
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{isabellebody}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isatagtheory
+%
+\endisatagtheory
+{\isafoldtheory}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\end{isabellebody}%
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: "root"
+%%% End:
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/tex/document/Tree_Slides.tex Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,259 @@
+%
+\begin{isabellebody}%
+\setisabellecontext{Tree{\isacharunderscore}Slides}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isatagtheory
+%
+\endisatagtheory
+{\isafoldtheory}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\end{isabellebody}%
+\section{Binary Trees}
+%-------------------------------------------------------------
+\ttfilepage{HOL/Library/Tree.thy\\ Thys/Tree\char`\_Additions.thy}{}
+%-------------------------------------------------------------
+\begin{frame}{Binary trees}
+\begin{center}
+\isakeyword{datatype} \high{\isa{{\isacharprime}a\ tree\ {\isacharequal}\ Leaf\ {\isacharbar}\ Node\ {\isacharparenleft}{\isacharprime}a\ tree{\isacharparenright}\ {\isacharprime}a\ {\isacharparenleft}{\isacharprime}a\ tree{\isacharparenright}}}
+\end{center}
+\pause
+
+Abbreviations:
+\begin{tabular}{rcl}
+\high{\isa{{\isasymlangle}{\isasymrangle}}} &\isa{{\isasymequiv}}& \isa{Leaf}\\\pause
+\high{\isa{{\isasymlangle}l{\isacharcomma}\ a{\isacharcomma}\ r{\isasymrangle}}} &\isa{{\isasymequiv}}& \isa{Node\ l\ a\ r}
+\end{tabular}
+%\pause\medskip
+
+%Terminology:
+%\begin{tabular}{rcl}
+%\high{leaf} & means & @ {text Leaf}\\\pause
+%\high{node} & means & @ {const Node}\\\pause
+%\high{leaf node} & means & @ {term"Node Leaf a Leaf"}
+%\end{tabular}
+\pause\vfill
+\begin{center}
+Most of the time: \high{tree = binary tree}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\section{Basic Functions}
+%-------------------------------------------------------------
+\begin{frame}{Tree traversal}
+\isa{inorder\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ tree\ {\isasymRightarrow}\ {\isacharprime}a\ list}\\\smallskip\pause
+\high{\isa{inorder\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}}}\\
+\high{\isa{inorder\ {\isasymlangle}l{\isacharcomma}\ x{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\ inorder\ l\ {\isacharat}\ {\isacharbrackleft}x{\isacharbrackright}\ {\isacharat}\ inorder\ r}}
+\bigskip\pause
+
+\isa{preorder\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ tree\ {\isasymRightarrow}\ {\isacharprime}a\ list}\\\smallskip\pause
+\high{\isa{preorder\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}}}\\
+\high{\isa{preorder\ {\isasymlangle}l{\isacharcomma}\ x{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\ x\ {\isacharhash}\ preorder\ l\ {\isacharat}\ preorder\ r}}
+\bigskip\pause
+
+\isa{postorder\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ tree\ {\isasymRightarrow}\ {\isacharprime}a\ list}\\\smallskip\pause
+\high{\isa{postorder\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}}}\\
+\high{\isa{postorder\ {\isasymlangle}l{\isacharcomma}\ x{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\ postorder\ l\ {\isacharat}\ postorder\ r\ {\isacharat}\ {\isacharbrackleft}x{\isacharbrackright}}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Size}
+
+\isa{size} \isa{{\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}a\ tree\ {\isasymRightarrow}\ nat}\\\smallskip\pause
+\high{\isa{{\isacharbar}{\isasymlangle}{\isasymrangle}{\isacharbar}\ {\isacharequal}\ {\isadigit{0}}}}\\
+\high{\isa{{\isacharbar}{\isasymlangle}l{\isacharcomma}\ \_{\isacharcomma}\ r{\isasymrangle}{\isacharbar}\ {\isacharequal}\ {\isacharbar}l{\isacharbar}\ {\isacharplus}\ {\isacharbar}r{\isacharbar}\ {\isacharplus}\ {\isadigit{1}}}}
+\bigskip\pause
+
+\isa{size{\isadigit{1}}\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ tree\ {\isasymRightarrow}\ nat}\\\pause
+\high{\isa{{\isacharbar}t{\isacharbar}\isactrlsub {\isadigit{1}}\ {\isacharequal}\ {\isacharbar}t{\isacharbar}\ {\isacharplus}\ {\isadigit{1}}}}\\
+\bigskip\pause
+
+\isa{{\isasymLongrightarrow}}\\
+\isa{{\isacharbar}{\isasymlangle}{\isasymrangle}{\isacharbar}\isactrlsub {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}}\\
+\isa{{\isacharbar}{\isasymlangle}l{\isacharcomma}\ x{\isacharcomma}\ r{\isasymrangle}{\isacharbar}\isactrlsub {\isadigit{1}}\ {\isacharequal}\ {\isacharbar}l{\isacharbar}\isactrlsub {\isadigit{1}}\ {\isacharplus}\ {\isacharbar}r{\isacharbar}\isactrlsub {\isadigit{1}}}
+\bigskip\pause
+
+\textbf{Lemma} The number of leaves in \isa{t} is \isa{{\isacharbar}t{\isacharbar}\isactrlsub {\isadigit{1}}}.
+\bigskip\pause
+
+\bad{Warning: \isa{{\isacharbar}{\isachardot}{\isacharbar}} and \isa{{\isacharbar}{\isachardot}{\isacharbar}\isactrlsub {\isadigit{1}}} only on slides}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Height}
+
+\isa{height} \isa{{\isacharcolon}{\isacharcolon}} \isa{{\isacharprime}a\ tree\ {\isasymRightarrow}\ nat}\\\smallskip\pause
+\high{\isa{h{\isacharparenleft}{\isasymlangle}{\isasymrangle}{\isacharparenright}\ {\isacharequal}\ {\isadigit{0}}}}\\
+\high{\isa{h{\isacharparenleft}{\isasymlangle}l{\isacharcomma}\ \_{\isacharcomma}\ r{\isasymrangle}{\isacharparenright}\ {\isacharequal}\ max\ {\isacharparenleft}h{\isacharparenleft}l{\isacharparenright}{\isacharparenright}\ {\isacharparenleft}h{\isacharparenleft}r{\isacharparenright}{\isacharparenright}\ {\isacharplus}\ {\isadigit{1}}}}
+\pause
+\begin{center}
+\bad{Warning: \isa{h{\isacharparenleft}{\isachardot}{\isacharparenright}} only on slides}
+\end{center}
+\pause
+
+\textbf{Lemma} \isa{h{\isacharparenleft}t{\isacharparenright}\ {\isasymle}\ {\isacharbar}t{\isacharbar}}
+\bigskip\pause
+
+\textbf{Lemma} \isa{{\isacharbar}t{\isacharbar}\isactrlsub {\isadigit{1}}\ {\isasymle}\ {\isadigit{2}}\isactrlbsup \isa{h{\isacharparenleft}t{\isacharparenright}}\isactrlesup }
+\end{frame}
+%-------------------------------------------------------------
+%\thyfile{Thys/Tree\char`\_Additions}{\isa{{\isacharbar}t{\isacharbar}\isactrlsub {\isadigit{1}}\ {\isasymle}\ {\isadigit{2}}\isactrlbsup \isa{h{\isacharparenleft}t{\isacharparenright}}\isactrlesup }}
+%-------------------------------------------------------------
+\begin{frame}{Minimal height}
+\isa{min{\isacharunderscore}height\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ tree\ {\isasymRightarrow}\ nat}\\\smallskip\pause
+\high{\isa{mh{\isacharparenleft}{\isasymlangle}{\isasymrangle}{\isacharparenright}\ {\isacharequal}\ {\isadigit{0}}}}\\
+\high{\isa{mh{\isacharparenleft}{\isasymlangle}l{\isacharcomma}\ \_{\isacharcomma}\ r{\isasymrangle}{\isacharparenright}\ {\isacharequal}\ min\ {\isacharparenleft}mh{\isacharparenleft}l{\isacharparenright}{\isacharparenright}\ {\isacharparenleft}mh{\isacharparenleft}r{\isacharparenright}{\isacharparenright}\ {\isacharplus}\ {\isadigit{1}}}}
+\pause
+\begin{center}
+\bad{Warning: \isa{mh{\isacharparenleft}{\isachardot}{\isacharparenright}} only on slides}
+\end{center}
+\pause
+
+\textbf{Lemma} \isa{mh{\isacharparenleft}t{\isacharparenright}\ {\isasymle}\ h{\isacharparenleft}t{\isacharparenright}}
+\bigskip\pause
+
+\textbf{Lemma} \isa{{\isadigit{2}}\isactrlbsup \isa{mh{\isacharparenleft}t{\isacharparenright}}\isactrlesup \ {\isasymle}\ {\isacharbar}t{\isacharbar}\isactrlsub {\isadigit{1}}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Internal path length}
+\pause
+\begin{center}
+The sum of the lengths of the (simpl) paths\\ from the root to any other node
+\end{center}
+\pause
+
+\isa{ipl\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ tree\ {\isasymRightarrow}\ nat}\\\smallskip\pause
+\high{\isa{ipl\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ {\isadigit{0}}}}\\
+\high{\isa{ipl\ {\isasymlangle}l{\isacharcomma}\ \_{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\ ipl\ l\ {\isacharplus}\ {\isacharbar}l{\isacharbar}\ {\isacharplus}\ ipl\ r\ {\isacharplus}\ {\isacharbar}r{\isacharbar}}}
+\vfill\pause
+\begin{center}
+Why relevant?
+\pause\bigskip
+
+Upper bound?
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\section{Complete and Balanced Trees}
+%-------------------------------------------------------------
+\begin{frame}{Complete tree}
+\isa{complete\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ tree\ {\isasymRightarrow}\ bool}\\\pause
+\high{\isa{complete\ {\isasymlangle}{\isasymrangle}\ {\isacharequal}\ True\isasep\isanewline%
+complete\ {\isasymlangle}l{\isacharcomma}\ \_{\isacharcomma}\ r{\isasymrangle}\ {\isacharequal}\isanewline
+{\isacharparenleft}complete\ l\ {\isasymand}\ complete\ r\ {\isasymand}\ h{\isacharparenleft}l{\isacharparenright}\ {\isacharequal}\ h{\isacharparenleft}r{\isacharparenright}{\isacharparenright}}}
+\bigskip\pause
+
+\textbf{Lemma} \isa{complete\ t\ {\isacharequal}\ {\isacharparenleft}mh{\isacharparenleft}t{\isacharparenright}\ {\isacharequal}\ h{\isacharparenleft}t{\isacharparenright}{\isacharparenright}}
+\bigskip\pause
+
+\textbf{Lemma} \isa{complete\ t\ {\isasymLongrightarrow}\ {\isacharbar}t{\isacharbar}\isactrlsub {\isadigit{1}}\ {\isacharequal}\ {\isadigit{2}}\isactrlbsup \isa{h{\isacharparenleft}t{\isacharparenright}}\isactrlesup }
+%Proof by induction; induction-free proof: 2^mh(t) <= size1 t <= 2^h(t)
+\bigskip\pause
+
+\textbf{Lemma} \isa{{\isacharbar}t{\isacharbar}\isactrlsub {\isadigit{1}}\ {\isacharequal}\ {\isadigit{2}}\isactrlbsup \isa{h{\isacharparenleft}t{\isacharparenright}}\isactrlesup \ {\isasymLongrightarrow}\ complete\ t}
+\pause
+
+\textbf{Lemma} \isa{{\isacharbar}t{\isacharbar}\isactrlsub {\isadigit{1}}\ {\isacharequal}\ {\isadigit{2}}\isactrlbsup \isa{mh{\isacharparenleft}t{\isacharparenright}}\isactrlesup \ {\isasymLongrightarrow}\ complete\ t}
+\bigskip\pause
+
+\textbf{Corollary} \isa{{\isasymnot}\ complete\ t\ {\isasymLongrightarrow}\ {\isacharbar}t{\isacharbar}\isactrlsub {\isadigit{1}}\ {\isacharless}\ {\isadigit{2}}\isactrlbsup \isa{h{\isacharparenleft}t{\isacharparenright}}\isactrlesup }
+\pause
+
+\textbf{Corollary} \isa{{\isasymnot}\ complete\ t\ {\isasymLongrightarrow}\ {\isadigit{2}}\isactrlbsup \isa{mh{\isacharparenleft}t{\isacharparenright}}\isactrlesup \ {\isacharless}\ {\isacharbar}t{\isacharbar}\isactrlsub {\isadigit{1}}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Complete tree: \isa{ipl}}
+\textbf{Lemma} A complete tree of height \isa{h} has internal path length
+\high{\isa{{\isacharparenleft}h\ {\isacharminus}\ {\isadigit{2}}{\isacharparenright}\ {\isacharasterisk}\ {\isadigit{2}}\isactrlbsup \isa{h}\isactrlesup \ {\isacharplus}\ {\isadigit{2}}}}.
+\bigskip\pause
+
+In a search tree, finding the node labelled \isa{x}
+takes as many steps as the path from the root to \isa{x} is long.
+\pause
+
+Thus the average time to find an element that is in the tree is
+\high{\isa{ipl\ t} \isa{{\isacharslash}} \isa{{\isacharbar}t{\isacharbar}}}.
+\bigskip\pause
+
+\textbf{Lemma} Let \isa{t} be a complete search tree of height \isa{h}.
+\pause
+The average time to find a random element that is in the tree
+is asymptotically \high{\isa{h\ {\isacharminus}\ {\isadigit{2}}}} (as \isa{h} approaches \isa{{\isasyminfinity}}):\pause
+\begin{center}
+\high{\isa{ipl\ t} \isa{{\isacharslash}} \isa{{\isacharbar}t{\isacharbar}} \isa{{\isasymsim}} \isa{h\ {\isacharminus}\ {\isadigit{2}}}}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Complete tree: \isa{ipl}}
+A problem: \high{\isa{{\isacharparenleft}h\ {\isacharminus}\ {\isadigit{2}}{\isacharparenright}\ {\isacharasterisk}\ {\isadigit{2}}\isactrlbsup \isa{h}\isactrlesup \ {\isacharplus}\ {\isadigit{2}}}} is only correct if interpreted
+over type \high{\isa{int}}, not \bad{\isa{nat}}.
+\bigskip\pause
+
+Correct version:
+
+\textbf{Lemma} \isa{complete\ t} \isa{{\isasymLongrightarrow}}
+\mbox{\isa{int\ {\isacharparenleft}ipl\ t{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}int\ {\isacharparenleft}h{\isacharparenleft}t{\isacharparenright}{\isacharparenright}\ {\isacharminus}\ {\isadigit{2}}{\isacharparenright}\ {\isacharasterisk}\ {\isadigit{2}}\isactrlbsup \isa{h{\isacharparenleft}t{\isacharparenright}}\isactrlesup \ {\isacharplus}\ {\isadigit{2}}}}
+\vfill\pause
+
+We do not cover the Isabelle formalization of limits.
+
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Balanced tree}
+\isa{balanced\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ tree\ {\isasymRightarrow}\ bool}\\\smallskip\pause
+\high{\isa{balanced\ t\ {\isacharequal}\ {\isacharparenleft}h{\isacharparenleft}t{\isacharparenright}\ {\isacharminus}\ mh{\isacharparenleft}t{\isacharparenright}\ {\isasymle}\ {\isadigit{1}}{\isacharparenright}}}
+\bigskip\pause
+
+Balanced trees have optimal height:\\
+\textbf{Lemma} \isa{{\sf If\,}\ \mbox{balanced\ t}\ {\sf \,and\,}\ \mbox{{\isacharbar}t{\isacharbar}\ {\isasymle}\ {\isacharbar}t{\isacharprime}{\isacharbar}}\ {\sf \,then\,}\ h{\isacharparenleft}t{\isacharparenright}\ {\isasymle}\ h{\isacharparenleft}t{\isacharprime}{\isacharparenright}{\isachardot}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Warning}
+\begin{itemize}
+\item The terms \bad{\emph{complete}} and \bad{\emph{balanced}}\\
+ are not defined uniquely in the literature.
+\bigskip\pause
+
+\item For example,\\
+Knuth calls \bad{\emph{complete}} what we call \bad{\emph{balanced}}.
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{isabellebody}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isatagtheory
+%
+\endisatagtheory
+{\isafoldtheory}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\end{isabellebody}%
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: "root"
+%%% End:
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/tex/document/Trie_Slides.tex Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,337 @@
+%
+\begin{isabellebody}%
+\setisabellecontext{Trie{\isacharunderscore}Slides}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isatagtheory
+%
+\endisatagtheory
+{\isafoldtheory}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\end{isabellebody}%
+%-------------------------------------------------------------
+\section{Tries and Patricia Tries}
+%-------------------------------------------------------------
+\ttfilepage{Thys/Trie1\\Thys/Trie2}{}
+%-------------------------------------------------------------
+\begin{frame}{Trie}
+\framesubtitle{[Fredkin, CACM 1960]}
+Name: \emph{reTRIEval}
+\pause
+\begin{itemize}
+\item
+Tries are search trees indexed by lists.
+\pause
+\item
+Each node maps the next element in the list\\ to a subtrie.
+\pause
+\item
+For simplicity we consider only binary tries:
+\high{%
+\begin{isabelle}%
+\isacommand{datatype}\ trie\ {\isacharequal}\ Leaf\ {\isacharbar}\ Node\ bool\ {\isacharparenleft}trie\ {\isasymtimes}\ trie{\isacharparenright}%
+\end{isabelle}}
+\end{itemize}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Patricia trie}
+\framesubtitle{[Morrison, JACM 1968]}
+Sequences of edges without branching are compressed
+\pause
+\high{%
+\begin{isabelle}%
+\isacommand{datatype}\ ptrie\ {\isacharequal}\ LeafP\isanewline
+\isaindent{\ \ }{\isacharbar}\ NodeP\ {\isacharparenleft}bool\ list{\isacharparenright}\ bool\ {\isacharparenleft}ptrie\ {\isasymtimes}\ ptrie{\isacharparenright}%
+\end{isabelle}}
+\pause
+
+Name: PATRICIA = \emph{Practical Algorithm To Retrieve Information Coded in Alphanumeric}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{From trie to Patricia trie}
+\framesubtitle{and beyond}
+\pause
+\begin{center}
+\Large An exercise in data refinement
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\subsection{Data Refinement Basics}
+%-------------------------------------------------------------
+\begin{frame}{Data refinement}
+How to implement an abstract type \isa{A} by a concrete (``refined'') type \isa{C}
+with \high{invariant \isa{inv\isactrlsub C\ {\isacharcolon}{\isacharcolon}\ C\ {\isasymRightarrow}\ bool}}:
+\begin{itemize}
+\pause
+\item
+Define an \high{abstraction function \isa{abs\isactrlsub C\ {\isacharcolon}{\isacharcolon}\ C\ {\isasymRightarrow}\ A}}
+\pause
+\item Implement each interface function \isa{f\isactrlsub A} on \isa{A}\\ by some \isa{f\isactrlsub C} on \isa{C}
+\pause
+\item Prove that each \isa{f\isactrlsub C} implements \isa{f\isactrlsub A}:\\
+\high{\isa{inv\isactrlsub C\ t\ {\isasymLongrightarrow}\ abs\isactrlsub C\ {\isacharparenleft}f\isactrlsub C\ t{\isacharparenright}\ {\isacharequal}\ f\isactrlsub A\ {\isacharparenleft}abs\isactrlsub C\ t{\isacharparenright}}}\\\pause
+and preserves the invariant:\\
+\high{\isa{inv\isactrlsub C\ t\ {\isasymLongrightarrow}\ inv\isactrlsub C\ {\isacharparenleft}f\isactrlsub C\ t{\isacharparenright}}}
+\end{itemize}
+\pause
+The \isa{f}\,s may take more parameters\\\pause
+ and may return values not of type \isa{A}/\isa{C}.
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Data refinement}
+\begin{center}
+\high{Multiple refinement steps can help\\
+to reduce the complexity of correctness proofs}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\subsection{Trie}
+%-------------------------------------------------------------
+\begin{frame}{Trie}
+\high{%
+\begin{isabelle}%
+\isacommand{datatype}\ trie\ {\isacharequal}\ Leaf\ {\isacharbar}\ Node\ bool\ {\isacharparenleft}trie\ {\isasymtimes}\ trie{\isacharparenright}%
+\end{isabelle}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Trie \isa{isin}}
+\isa{isin\ Leaf\ ks\ {\isacharequal}\ False}\pause
+%
+\begin{isabelle}%
+isin\ {\isacharparenleft}Node\ b\ {\isacharparenleft}l{\isacharcomma}\ r{\isacharparenright}{\isacharparenright}\ ks\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf case}\ ks\ {\sf of}\isanewline
+\isaindent{{\isacharparenleft}}\ \ {\isacharbrackleft}{\isacharbrackright}\ {\isasymRightarrow}\ b\isanewline
+\isaindent{{\isacharparenleft}}{\isacharbar}\ k\ {\isacharhash}\ x\ {\isasymRightarrow}\ isin\ {\isacharparenleft}{\sf if}\ k\ {\sf then}\ r\ {\sf else}\ l{\isacharparenright}\ x{\isacharparenright}%
+\end{isabelle}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Trie \isa{insert}}
+\isa{insert\ {\isacharbrackleft}{\isacharbrackright}\ Leaf\ {\isacharequal}\ Node\ True\ {\isacharparenleft}Leaf{\isacharcomma}\ Leaf{\isacharparenright}}\\\pause
+\isa{insert\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharparenleft}Node\ b\ lr{\isacharparenright}\ {\isacharequal}\ Node\ True\ lr}\\\pause
+%
+\begin{isabelle}%
+insert\ {\isacharparenleft}k\ {\isacharhash}\ ks{\isacharparenright}\ Leaf\ {\isacharequal}\isanewline
+Node\ False\isanewline
+\isaindent{\ }{\isacharparenleft}{\sf if}\ k\ {\sf then}\ {\isacharparenleft}Leaf{\isacharcomma}\ insert\ ks\ Leaf{\isacharparenright}\isanewline
+\isaindent{\ {\isacharparenleft}}{\sf else}\ {\isacharparenleft}insert\ ks\ Leaf{\isacharcomma}\ Leaf{\isacharparenright}{\isacharparenright}%
+\end{isabelle}\pause
+\isa{insert\ {\isacharparenleft}k\ {\isacharhash}\ ks{\isacharparenright}\ {\isacharparenleft}Node\ b\ {\isacharparenleft}l{\isacharcomma}\ r{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline
+Node\ b\ {\isacharparenleft}{\sf if}\ k\ {\sf then}\ {\isacharparenleft}l{\isacharcomma}\ insert\ ks\ r{\isacharparenright}\ {\sf else}\ {\isacharparenleft}insert\ ks\ l{\isacharcomma}\ r{\isacharparenright}{\isacharparenright}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of \isa{trie} w.r.t. \isa{set}}
+\begin{center}
+\high{\isa{isin\ {\isacharparenleft}insert\ as\ t{\isacharparenright}\ bs\ {\isacharequal}\ {\isacharparenleft}as\ {\isacharequal}\ bs\ {\isasymor}\ isin\ t\ bs{\isacharparenright}}}
+\end{center}
+\pause
+Note: \isa{isin} doubles as abstraction function
+\end{frame}
+%-------------------------------------------------------------
+\subsection{Patricia Trie}
+%-------------------------------------------------------------
+\begin{frame}{Patricia trie}
+\high{%
+\begin{isabelle}%
+\isacommand{datatype}\ ptrie\ {\isacharequal}\ LeafP\isanewline
+\isaindent{\ \ }{\isacharbar}\ NodeP\ {\isacharparenleft}bool\ list{\isacharparenright}\ bool\ {\isacharparenleft}ptrie\ {\isasymtimes}\ ptrie{\isacharparenright}%
+\end{isabelle}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\isa{isinP}}
+\isa{isinP\ LeafP\ ks\ {\isacharequal}\ False}\pause
+%
+\begin{isabelle}%
+isinP\ {\isacharparenleft}NodeP\ ps\ b\ {\isacharparenleft}l{\isacharcomma}\ r{\isacharparenright}{\isacharparenright}\ ks\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf let}\ n\ {\isacharequal}\ length\ ps\isanewline
+\isaindent{{\isacharparenleft}}{\sf in}\ {\sf if}\ ps\ {\isacharequal}\ take\ n\ ks\isanewline
+\isaindent{{\isacharparenleft}{\sf in}\ }{\sf then}\ {\sf case}\ drop\ n\ ks\ {\sf of}\isanewline
+\isaindent{{\isacharparenleft}{\sf in}\ {\sf then}\ }\ \ {\isacharbrackleft}{\isacharbrackright}\ {\isasymRightarrow}\ b\isanewline
+\isaindent{{\isacharparenleft}{\sf in}\ {\sf then}\ }{\isacharbar}\ k\ {\isacharhash}\ x\ {\isasymRightarrow}\ isinP\ {\isacharparenleft}{\sf if}\ k\ {\sf then}\ r\ {\sf else}\ l{\isacharparenright}\ x\isanewline
+\isaindent{{\isacharparenleft}{\sf in}\ }{\sf else}\ False{\isacharparenright}%
+\end{isabelle}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Splitting lists}
+\high{\isa{split\ xs\ ys\ {\isacharequal}\ {\isacharparenleft}zs{\isacharcomma}\ xs{\isacharprime}{\isacharcomma}\ ys{\isacharprime}{\isacharparenright}}}\\\pause
+iff \isa{zs} is the longest common prefix of \isa{xs} and \isa{ys}\\\pause
+and \isa{xs{\isacharprime}}/\isa{ys{\isacharprime}} is the remainder of \isa{xs}/\isa{ys}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\isa{insertP}}
+\pause
+\small
+\isa{insertP\ ks\ {\isacharparenleft}NodeP\ ps\ b\ {\isacharparenleft}l{\isacharcomma}\ r{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf case}\ split\ ks\ ps\ {\sf of}\isanewline
+\isaindent{{\isacharparenleft}}\ \ {\isacharparenleft}qs{\isacharcomma}\ {\isacharbrackleft}{\isacharbrackright}{\isacharcomma}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isasymRightarrow}\ NodeP\ ps\ True\ {\isacharparenleft}l{\isacharcomma}\ r{\isacharparenright}\isanewline
+\isaindent{{\isacharparenleft}}{\isacharbar}\ {\isacharparenleft}qs{\isacharcomma}\ {\isacharbrackleft}{\isacharbrackright}{\isacharcomma}\ p\ {\isacharhash}\ ps{\isacharprime}{\isacharparenright}\ {\isasymRightarrow}\isanewline
+\isaindent{{\isacharparenleft}{\isacharbar}\ \ \ }{\sf let}\ t\ {\isacharequal}\ NodeP\ ps{\isacharprime}\ b\ {\isacharparenleft}l{\isacharcomma}\ r{\isacharparenright}\isanewline
+\isaindent{{\isacharparenleft}{\isacharbar}\ \ \ }{\sf in}\ NodeP\ qs\ True\ {\isacharparenleft}{\sf if}\ p\ {\sf then}\ {\isacharparenleft}LeafP{\isacharcomma}\ t{\isacharparenright}\ {\sf else}\ {\isacharparenleft}t{\isacharcomma}\ LeafP{\isacharparenright}{\isacharparenright}\isanewline
+\isaindent{{\isacharparenleft}}{\isacharbar}\ {\isacharparenleft}qs{\isacharcomma}\ k\ {\isacharhash}\ ks{\isacharprime}{\isacharcomma}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isasymRightarrow}\isanewline
+\isaindent{{\isacharparenleft}{\isacharbar}\ \ \ }NodeP\ ps\ b\isanewline
+\isaindent{{\isacharparenleft}{\isacharbar}\ \ \ \ }{\isacharparenleft}{\sf if}\ k\ {\sf then}\ {\isacharparenleft}l{\isacharcomma}\ insertP\ ks{\isacharprime}\ r{\isacharparenright}\ {\sf else}\ {\isacharparenleft}insertP\ ks{\isacharprime}\ l{\isacharcomma}\ r{\isacharparenright}{\isacharparenright}\isanewline
+\isaindent{{\isacharparenleft}}{\isacharbar}\ {\isacharparenleft}qs{\isacharcomma}\ k\ {\isacharhash}\ ks{\isacharprime}{\isacharcomma}\ p\ {\isacharhash}\ ps{\isacharprime}{\isacharparenright}\ {\isasymRightarrow}\isanewline
+\isaindent{{\isacharparenleft}{\isacharbar}\ \ \ }{\sf let}\ tp\ {\isacharequal}\ NodeP\ ps{\isacharprime}\ b\ {\isacharparenleft}l{\isacharcomma}\ r{\isacharparenright}{\isacharsemicolon}\isanewline
+\isaindent{{\isacharparenleft}{\isacharbar}\ \ \ {\sf let}\ }tk\ {\isacharequal}\ NodeP\ ks{\isacharprime}\ True\ {\isacharparenleft}LeafP{\isacharcomma}\ LeafP{\isacharparenright}\isanewline
+\isaindent{{\isacharparenleft}{\isacharbar}\ \ \ }{\sf in}\ NodeP\ qs\ False\ {\isacharparenleft}{\sf if}\ k\ {\sf then}\ {\isacharparenleft}tp{\isacharcomma}\ tk{\isacharparenright}\ {\sf else}\ {\isacharparenleft}tk{\isacharcomma}\ tp{\isacharparenright}{\isacharparenright}{\isacharparenright}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\isa{abs{\isacharunderscore}ptrie\ {\isacharcolon}{\isacharcolon}\ ptrie\ {\isasymRightarrow}\ trie}}
+\pause
+\high{\isa{abs{\isacharunderscore}ptrie\ LeafP\ {\isacharequal}\ Leaf}}
+\pause
+\high{%
+\begin{isabelle}%
+abs{\isacharunderscore}ptrie\ {\isacharparenleft}NodeP\ ps\ b\ {\isacharparenleft}l{\isacharcomma}\ r{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline
+prefix{\isacharunderscore}trie\ ps\ {\isacharparenleft}Node\ b\ {\isacharparenleft}abs{\isacharunderscore}ptrie\ l{\isacharcomma}\ abs{\isacharunderscore}ptrie\ r{\isacharparenright}{\isacharparenright}%
+\end{isabelle}}
+\pause
+\isa{prefix{\isacharunderscore}trie\ {\isacharcolon}{\isacharcolon}\ bool\ list\ {\isasymRightarrow}\ trie\ {\isasymRightarrow}\ trie}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of \isa{ptrie} w.r.t. \isa{trie}}
+\begin{center}
+\high{\isa{isinP\ t\ ks\ {\isacharequal}\ isin\ {\isacharparenleft}abs{\isacharunderscore}ptrie\ t{\isacharparenright}\ ks}}
+\bigskip\pause
+
+\high{\isa{abs{\isacharunderscore}ptrie\ {\isacharparenleft}insertP\ ks\ t{\isacharparenright}\ {\isacharequal}\ insert\ ks\ {\isacharparenleft}abs{\isacharunderscore}ptrie\ t{\isacharparenright}}}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\subsection{Patricia Trie for Words}
+%-------------------------------------------------------------
+\begin{frame}{Patricia trie for words}
+\framesubtitle{[Okasaki \& Gill, 1998]}
+\high{%
+\begin{isabelle}%
+\isacommand{datatype}\ ptrieW\ {\isacharequal}\ LeafW\isanewline
+\isaindent{\ \ }{\isacharbar}\ NodeW\ {\isacharparenleft}{\isadigit{3}}{\isadigit{2}}\ word{\isacharparenright}\ {\isacharparenleft}{\isadigit{3}}{\isadigit{2}}\ word{\isacharparenright}\ bool\isanewline
+\isaindent{\ \ {\isacharbar}\ \ \ }{\isacharparenleft}ptrieW\ {\isasymtimes}\ ptrieW{\isacharparenright}%
+\end{isabelle}}
+\bigskip\pause
+
+\isa{to{\isacharunderscore}bl\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ word\ {\isasymRightarrow}\ bool\ list}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\isa{isinW}}
+\isa{isinW\ LeafW\ k\ {\isacharequal}\ False}\pause
+%
+\begin{isabelle}%
+isinW\ {\isacharparenleft}NodeW\ p\ m\ b\ {\isacharparenleft}l{\isacharcomma}\ r{\isacharparenright}{\isacharparenright}\ k\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf if}\ b\ {\sf then}\ k\ {\isacharequal}\ p\isanewline
+\isaindent{{\isacharparenleft}}{\sf else}\ {\sf if}\ match{\isacharunderscore}prefix\ p\ m\ k\isanewline
+\isaindent{{\isacharparenleft}{\sf else}\ }{\sf then}\ isinW\ {\isacharparenleft}{\sf if}\ go{\isacharunderscore}left\ m\ k\ {\sf then}\ l\ {\sf else}\ r{\isacharparenright}\ k\isanewline
+\isaindent{{\isacharparenleft}{\sf else}\ }{\sf else}\ False{\isacharparenright}%
+\end{isabelle}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Correctness of \isa{ptrie} w.r.t. \isa{trie}}
+\begin{center}
+\high{%
+\begin{isabelle}%
+inv{\isacharunderscore}ptrieW\ t\ {\isasymLongrightarrow}\isanewline
+isinW\ t\ k\ {\isacharequal}\ isinL\ {\isacharparenleft}abs{\isacharunderscore}ptrieW\ t{\isacharparenright}\ {\isacharparenleft}rev\ {\isacharparenleft}to{\isacharunderscore}bl\ k{\isacharparenright}{\isacharparenright}%
+\end{isabelle}}
+\end{center}
+\pause
+Proof uses intermedidate version of Patricia trie for lists
+where nodes contain complete prefixes (\isa{isinL}).
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Refinement hierarchy}
+\begin{center}
+\begin{tabular}{cl}
+\isa{trie}\\
+$\uparrow$ & \isa{abs{\isacharunderscore}ptrie}\\
+\isa{ptrie}\\
+$\uparrow$ & \isa{abs{\isacharunderscore}ptrieL}\\
+\isa{ptrieL}\\
+$\uparrow$ & \isa{abs{\isacharunderscore}ptrieW}\\
+\isa{ptrieW}\\[1ex]
+\multicolumn{2}{l}{\pause where \isa{ptrieL} \isa{{\isacharequal}} \isa{ptrie}}
+\end{tabular}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\mbox{\isa{inv{\isacharunderscore}ptrieW\ {\isacharcolon}{\isacharcolon}\ ptrieW\ {\isasymRightarrow}\ bool}}}
+\pause
+\high{\isa{inv{\isacharunderscore}ptrieW\ LeafW\ {\isacharequal}\ True}}
+\pause
+\high{%
+\begin{isabelle}%
+inv{\isacharunderscore}ptrieW\ {\isacharparenleft}NodeW\ p\ m\ b\ {\isacharparenleft}l{\isacharcomma}\ r{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline
+{\isacharparenleft}inv{\isacharunderscore}pref\ p\ m\ {\isasymand}\ inv{\isacharunderscore}mask\ m\ {\isasymand}\ inv{\isacharunderscore}ptrieW\ l\ {\isasymand}\ inv{\isacharunderscore}ptrieW\ r{\isacharparenright}%
+\end{isabelle}}
+\pause
+%
+\begin{isabelle}%
+inv{\isacharunderscore}pref\ p\ m\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf let}\ ps\ {\isacharequal}\ to{\isacharunderscore}bl\ p{\isacharsemicolon}\ ms\ {\isacharequal}\ to{\isacharunderscore}bl\ m{\isacharsemicolon}\isanewline
+\isaindent{{\isacharparenleft}{\sf let}\ }n\ {\isacharequal}\ length\ {\isacharparenleft}dropWhile\ Not\ ms{\isacharparenright}{\isacharsemicolon}\isanewline
+\isaindent{{\isacharparenleft}{\sf let}\ }ps{\isacharprime}\ {\isacharequal}\ take\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}\ ps\isanewline
+\isaindent{{\isacharparenleft}}{\sf in}\ set\ ps{\isacharprime}\ {\isacharequal}\ {\isacharbraceleft}False{\isacharbraceright}{\isacharparenright}%
+\end{isabelle}
+\pause
+\isa{inv{\isacharunderscore}mask\ m\ {\isacharequal}\ {\isacharparenleft}{\isacharbrackleft}b{\isasymleftarrow}to{\isacharunderscore}bl\ m\ {\isachardot}\ b{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}True{\isacharbrackright}{\isacharparenright}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\mbox{\isa{abs{\isacharunderscore}ptrieW\ {\isacharcolon}{\isacharcolon}\ ptrieW\ {\isasymRightarrow}\ ptrie}}}
+\pause
+\high{\isa{abs{\isacharunderscore}ptrieW\ LeafW\ {\isacharequal}\ LeafP}}
+\pause
+\high{%
+\begin{isabelle}%
+abs{\isacharunderscore}ptrieW\ {\isacharparenleft}NodeW\ p\ m\ b\ {\isacharparenleft}l{\isacharcomma}\ r{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline
+NodeP\ {\isacharparenleft}abs{\isacharunderscore}pmb\ p\ m\ b{\isacharparenright}\ b\ {\isacharparenleft}abs{\isacharunderscore}ptrieW\ l{\isacharcomma}\ abs{\isacharunderscore}ptrieW\ r{\isacharparenright}%
+\end{isabelle}}
+\pause
+%
+\begin{isabelle}%
+abs{\isacharunderscore}pmb\ p\ m\ b\ {\isacharequal}\isanewline
+{\isacharparenleft}{\sf if}\ b\ {\sf then}\ rev\ {\isacharparenleft}to{\isacharunderscore}bl\ p{\isacharparenright}\isanewline
+\isaindent{{\isacharparenleft}}{\sf else}\ {\sf let}\ ps\ {\isacharequal}\ to{\isacharunderscore}bl\ p{\isacharsemicolon}\ ms\ {\isacharequal}\ to{\isacharunderscore}bl\ m{\isacharsemicolon}\isanewline
+\isaindent{{\isacharparenleft}{\sf else}\ {\sf let}\ }n\ {\isacharequal}\ length\ {\isacharparenleft}dropWhile\ Not\ ms{\isacharparenright}\isanewline
+\isaindent{{\isacharparenleft}{\sf else}\ }{\sf in}\ rev\ {\isacharparenleft}drop\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}\ ps{\isacharparenright}{\isacharparenright}%
+\end{isabelle}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{\mbox{\isa{insertW}}}
+\begin{center}
+\Large\bad{?}
+\end{center}
+\end{frame}
+%-------------------------------------------------------------
+\begin{frame}{Further space optimization?}
+\pause
+\textbf{datatype} \isa{ptrieW{\isadigit{2}}\ {\isacharequal}}\\
+\quad \isa{LeafW{\isadigit{2}}} ~$|$\\
+\quad \isa{SingW{\isadigit{2}}\ {\isacharparenleft}{\isadigit{3}}{\isadigit{2}}\ word{\isacharparenright}} ~$|$\\
+\quad \isa{NodeW{\isadigit{2}}\ {\isacharparenleft}{\isadigit{3}}{\isadigit{2}}\ word{\isacharparenright}\ {\isacharparenleft}{\isadigit{3}}{\isadigit{2}}\ word{\isacharparenright}\ {\isacharparenleft}ptrieW{\isadigit{2}}\ {\isasymtimes}\ ptrieW{\isadigit{2}}{\isacharparenright}}
+\end{frame}
+%-------------------------------------------------------------
+\begin{isabellebody}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isatagtheory
+%
+\endisatagtheory
+{\isafoldtheory}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\end{isabellebody}%
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: "root"
+%%% End:
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/tex/document/comment.sty Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,278 @@
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+% Comment.sty version 3.6, October 1999
+%
+% Purpose:
+% selectively in/exclude pieces of text: the user can define new
+% comment versions, and each is controlled separately.
+% Special comments can be defined where the user specifies the
+% action that is to be taken with each comment line.
+%
+% Author
+% Victor Eijkhout
+% Department of Computer Science
+% University of Tennessee
+% 107 Ayres Hall
+% Knoxville TN 37996
+% USA
+%
+% victor@eijkhout.net
+%
+% This program is free software; you can redistribute it and/or
+% modify it under the terms of the GNU General Public License
+% as published by the Free Software Foundation; either version 2
+% of the License, or (at your option) any later version.
+%
+% This program is distributed in the hope that it will be useful,
+% but WITHOUT ANY WARRANTY; without even the implied warranty of
+% MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the
+% GNU General Public License for more details.
+%
+% For a copy of the GNU General Public License, write to the
+% Free Software Foundation, Inc.,
+% 59 Temple Place - Suite 330, Boston, MA 02111-1307, USA,
+% or find it on the net, for instance at
+% http://www.gnu.org/copyleft/gpl.html
+%
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+% This style can be used with plain TeX or LaTeX, and probably
+% most other packages too.
+%
+% Usage: all text included between
+% \comment ... \endcomment
+% or \begin{comment} ... \end{comment}
+% is discarded.
+%
+% The opening and closing commands should appear on a line
+% of their own. No starting spaces, nothing after it.
+% This environment should work with arbitrary amounts
+% of comment, and the comment can be arbitrary text.
+%
+% Other `comment' environments are defined by
+% and are selected/deselected with
+% \includecomment{versiona}
+% \excludecoment{versionb}
+%
+% These environments are used as
+% \versiona ... \endversiona
+% or \begin{versiona} ... \end{versiona}
+% with the opening and closing commands again on a line of
+% their own.
+%
+% LaTeX users note: for an included comment, the
+% \begin and \end lines act as if they don't exist.
+% In particular, they don't imply grouping, so assignments
+% &c are not local.
+%
+% Special comments are defined as
+% \specialcomment{name}{before commands}{after commands}
+% where the second and third arguments are executed before
+% and after each comment block. You can use this for global
+% formatting commands.
+% To keep definitions &c local, you can include \begingroup
+% in the `before commands' and \endgroup in the `after commands'.
+% ex:
+% \specialcomment{smalltt}
+% {\begingroup\ttfamily\footnotesize}{\endgroup}
+% You do *not* have to do an additional
+% \includecomment{smalltt}
+% To remove 'smalltt' blocks, give \excludecomment{smalltt}
+% after the definition.
+%
+% Processing comments can apply processing to each line.
+% \processcomment{name}{each-line commands}%
+% {before commands}{after commands}
+% By defining a control sequence
+% \def\Thiscomment##1{...} in the before commands the user can
+% specify what is to be done with each comment line.
+% BUG this does not work quite yet BUG
+%
+% Trick for short in/exclude macros (such as \maybe{this snippet}):
+%\includecomment{cond}
+%\newcommand{\maybe}[1]{}
+%\begin{cond}
+%\renewcommand{\maybe}[1]{#1}
+%\end{cond}
+%
+% Basic approach of the implementation:
+% to comment something out, scoop up every line in verbatim mode
+% as macro argument, then throw it away.
+% For inclusions, in LaTeX the block is written out to
+% a file \CommentCutFile (default "comment.cut"), which is
+% then included.
+% In plain TeX (and other formats) both the opening and
+% closing comands are defined as noop.
+%
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+% Changes in version 3.1
+% - updated author's address
+% - cleaned up some code
+% - trailing contents on \begin{env} line is always discarded
+% even if you've done \includecomment{env}
+% - comments no longer define grouping!! you can even
+% \includecomment{env}
+% \begin{env}
+% \begin{itemize}
+% \end{env}
+% Isn't that something ...
+% - included comments are written to file and input again.
+% Changes in 3.2
+% - \specialcomment brought up to date (thanks to Ivo Welch).
+% Changes in 3.3
+% - updated author's address again
+% - parametrised \CommentCutFile
+% Changes in 3.4
+% - added GNU public license
+% - added \processcomment, because Ivo's fix (above) brought an
+% inconsistency to light.
+% Changes in 3.5
+% - corrected typo in header.
+% - changed author email
+% - corrected \specialcomment yet again.
+% - fixed excludecomment of an earlier defined environment.
+% Changes in 3.6
+% - The 'cut' file is now written more verbatim, using \meaning;
+% some people reported having trouble with ISO latin 1, or umlaute.sty.
+% - removed some \newif statements.
+% Has this suddenly become \outer again?
+%
+% Known bugs:
+% - excludecomment leads to one superfluous space
+% - processcomment leads to a superfluous line break
+%
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+
+\def\makeinnocent#1{\catcode`#1=12 }
+\def\csarg#1#2{\expandafter#1\csname#2\endcsname}
+\def\latexname{lplain}\def\latexename{LaTeX2e}
+\newwrite\CommentStream
+\def\CommentCutFile{comment.cut}
+
+\def\ProcessComment#1% start it all of
+ {\begingroup
+ \def\CurrentComment{#1}%
+ \let\do\makeinnocent \dospecials
+ \makeinnocent\^^L% and whatever other special cases
+ \endlinechar`\^^M \catcode`\^^M=12 \xComment}
+%\def\ProcessCommentWithArg#1#2% to be used in \leveledcomment
+% {\begingroup
+% \def\CurrentComment{#1}%
+% \let\do\makeinnocent \dospecials
+% \makeinnocent\^^L% and whatever other special cases
+% \endlinechar`\^^M \catcode`\^^M=12 \xComment}
+{\catcode`\^^M=12 \endlinechar=-1 %
+ \gdef\xComment#1^^M{%
+ \expandafter\ProcessCommentLine}
+ \gdef\ProcessCommentLine#1^^M{\def\test{#1}
+ \csarg\ifx{End\CurrentComment Test}\test
+ \edef\next{\noexpand\EndOfComment{\CurrentComment}}%
+ \else \ThisComment{#1}\let\next\ProcessCommentLine
+ \fi \next}
+}
+
+\def\CSstringmeaning#1{\expandafter\CSgobblearrow\meaning#1}
+\def\CSstringcsnoescape#1{\expandafter\CSgobbleescape\string#1}
+{\escapechar-1
+\expandafter\expandafter\expandafter\gdef
+ \expandafter\expandafter\expandafter\CSgobblearrow
+ \expandafter\string\csname macro:->\endcsname{}
+}
+\def\CSgobbleescape#1{\ifnum`\\=`#1 \else #1\fi}
+\def\WriteCommentLine#1{\def\CStmp{#1}%
+ \immediate\write\CommentStream{\CSstringmeaning\CStmp}}
+
+% 3.1 change: in LaTeX and LaTeX2e prevent grouping
+\if 0%
+\ifx\fmtname\latexename
+ 0%
+\else \ifx\fmtname\latexname
+ 0%
+ \else
+ 1%
+\fi \fi
+%%%%
+%%%% definitions for LaTeX
+%%%%
+\def\AfterIncludedComment
+ {\immediate\closeout\CommentStream
+ \input{\CommentCutFile}\relax
+ }%
+\def\TossComment{\immediate\closeout\CommentStream}
+\def\BeforeIncludedComment
+ {\immediate\openout\CommentStream=\CommentCutFile
+ \let\ThisComment\WriteCommentLine}
+\def\includecomment
+ #1{\message{Include comment '#1'}%
+ \csarg\let{After#1Comment}\AfterIncludedComment
+ \csarg\def{#1}{\BeforeIncludedComment
+ \ProcessComment{#1}}%
+ \CommentEndDef{#1}}
+\long\def\specialcomment
+ #1#2#3{\message{Special comment '#1'}%
+ % note: \AfterIncludedComment does \input, so #2 goes here!
+ \csarg\def{After#1Comment}{#2\AfterIncludedComment#3}%
+ \csarg\def{#1}{\BeforeIncludedComment\relax
+ \ProcessComment{#1}}%
+ \CommentEndDef{#1}}
+\long\def\processcomment
+ #1#2#3#4{\message{Lines-Processing comment '#1'}%
+ \csarg\def{After#1Comment}{#3\AfterIncludedComment#4}%
+ \csarg\def{#1}{\BeforeIncludedComment#2\relax
+ \ProcessComment{#1}}%
+ \CommentEndDef{#1}}
+\def\leveledcomment
+ #1#2{\message{Include comment '#1' up to level '#2'}%
+ %\csname #1IsLeveledCommenttrue\endcsname
+ \csarg\let{After#1Comment}\AfterIncludedComment
+ \csarg\def{#1}{\BeforeIncludedComment
+ \ProcessCommentWithArg{#1}}%
+ \CommentEndDef{#1}}
+\else
+%%%%
+%%%%plain TeX and other formats
+%%%%
+\def\includecomment
+ #1{\message{Including comment '#1'}%
+ \csarg\def{#1}{}%
+ \csarg\def{end#1}{}}
+\long\def\specialcomment
+ #1#2#3{\message{Special comment '#1'}%
+ \csarg\def{#1}{\def\ThisComment{}\def\AfterComment{#3}#2%
+ \ProcessComment{#1}}%
+ \CommentEndDef{#1}}
+\fi
+
+%%%%
+%%%% general definition of skipped comment
+%%%%
+\def\excludecomment
+ #1{\message{Excluding comment '#1'}%
+ \csarg\def{#1}{\let\AfterComment\relax
+ \def\ThisComment####1{}\ProcessComment{#1}}%
+ \csarg\let{After#1Comment}\TossComment
+ \CommentEndDef{#1}}
+
+\if 0%
+\ifx\fmtname\latexename
+ 0%
+\else \ifx\fmtname\latexname
+ 0%
+ \else
+ 1%
+\fi \fi
+% latex & latex2e:
+\def\EndOfComment#1{\endgroup\end{#1}%
+ \csname After#1Comment\endcsname}
+\def\CommentEndDef#1{{\escapechar=-1\relax
+ \csarg\xdef{End#1Test}{\string\\end\string\{#1\string\}}%
+ }}
+\else
+% plain & other
+\def\EndOfComment#1{\endgroup\AfterComment}
+\def\CommentEndDef#1{{\escapechar=-1\relax
+ \csarg\xdef{End#1Test}{\string\\end#1}%
+ }}
+\fi
+
+\excludecomment{comment}
+
+\endinput
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/tex/document/isabelle.sty Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,259 @@
+%%
+%% macros for Isabelle generated LaTeX output
+%%
+
+%%% Simple document preparation (based on theory token language and symbols)
+
+% isabelle environments
+
+\newcommand{\isabellecontext}{UNKNOWN}
+\newcommand{\setisabellecontext}[1]{\def\isabellecontext{#1}}
+
+\newcommand{\isastyle}{\UNDEF}
+\newcommand{\isastylett}{\UNDEF}
+\newcommand{\isastyleminor}{\UNDEF}
+\newcommand{\isastyleminortt}{\UNDEF}
+\newcommand{\isastylescript}{\UNDEF}
+\newcommand{\isastyletext}{\normalsize\rm}
+\newcommand{\isastyletxt}{\rm}
+\newcommand{\isastylecmt}{\rm}
+
+\newcommand{\isaspacing}{%
+ \sfcode 42 1000 % .
+ \sfcode 63 1000 % ?
+ \sfcode 33 1000 % !
+ \sfcode 58 1000 % :
+ \sfcode 59 1000 % ;
+ \sfcode 44 1000 % ,
+}
+
+%symbol markup -- \emph achieves decent spacing via italic corrections
+\newcommand{\isamath}[1]{\emph{$#1$}}
+\newcommand{\isatext}[1]{\emph{#1}}
+\DeclareRobustCommand{\isascriptstyle}{\def\isamath##1{##1}\def\isatext##1{\mbox{\isaspacing\isastylescript##1}}}
+\newcommand{\isactrlsub}[1]{\emph{\isascriptstyle${}\sb{#1}$}}
+\newcommand{\isactrlsup}[1]{\emph{\isascriptstyle${}\sp{#1}$}}
+\DeclareRobustCommand{\isactrlbsub}{\emph\bgroup\math{}\sb\bgroup\mbox\bgroup\isaspacing\isastylescript}
+\DeclareRobustCommand{\isactrlesub}{\egroup\egroup\endmath\egroup}
+\DeclareRobustCommand{\isactrlbsup}{\emph\bgroup\math{}\sp\bgroup\mbox\bgroup\isaspacing\isastylescript}
+\DeclareRobustCommand{\isactrlesup}{\egroup\egroup\endmath\egroup}
+\newcommand{\isactrlbold}[1]{{\bfseries\upshape\boldmath#1}}
+
+\newcommand{\isaantiqcontrol}[1]{\isatt{{\char`\\}{\char`\<}{\char`\^}#1{\char`\>}}}
+\newenvironment{isaantiq}{{\isacharat\isacharbraceleft}}{{\isacharbraceright}}
+
+\newdimen\isa@parindent\newdimen\isa@parskip
+
+\newenvironment{isabellebody}{%
+\isamarkuptrue\par%
+\isa@parindent\parindent\parindent0pt%
+\isa@parskip\parskip\parskip0pt%
+\isaspacing\isastyle}{\par}
+
+\newenvironment{isabellebodytt}{%
+\isamarkuptrue\par%
+\isa@parindent\parindent\parindent0pt%
+\isa@parskip\parskip\parskip0pt%
+\isaspacing\isastylett}{\par}
+
+\newenvironment{isabelle}
+{\begin{trivlist}\begin{isabellebody}\item\relax}
+{\end{isabellebody}\end{trivlist}}
+
+\newenvironment{isabellett}
+{\begin{trivlist}\begin{isabellebodytt}\item\relax}
+{\end{isabellebodytt}\end{trivlist}}
+
+\newcommand{\isa}[1]{\emph{\isaspacing\isastyleminor #1}}
+\newcommand{\isatt}[1]{\emph{\isaspacing\isastyleminortt #1}}
+
+\newcommand{\isaindent}[1]{\hphantom{#1}}
+\newcommand{\isanewline}{\mbox{}\par\mbox{}}
+\newcommand{\isasep}{}
+\newcommand{\isadigit}[1]{#1}
+
+\newcommand{\isachardefaults}{%
+\def\isacharbell{\isamath{\bigbox}}%requires stmaryrd
+\chardef\isacharbang=`\!%
+\chardef\isachardoublequote=`\"%
+\chardef\isachardoublequoteopen=`\"%
+\chardef\isachardoublequoteclose=`\"%
+\chardef\isacharhash=`\#%
+\chardef\isachardollar=`\$%
+\chardef\isacharpercent=`\%%
+\chardef\isacharampersand=`\&%
+\chardef\isacharprime=`\'%
+\chardef\isacharparenleft=`\(%
+\chardef\isacharparenright=`\)%
+\chardef\isacharasterisk=`\*%
+\chardef\isacharplus=`\+%
+\chardef\isacharcomma=`\,%
+\chardef\isacharminus=`\-%
+\chardef\isachardot=`\.%
+\chardef\isacharslash=`\/%
+\chardef\isacharcolon=`\:%
+\chardef\isacharsemicolon=`\;%
+\chardef\isacharless=`\<%
+\chardef\isacharequal=`\=%
+\chardef\isachargreater=`\>%
+\chardef\isacharquery=`\?%
+\chardef\isacharat=`\@%
+\chardef\isacharbrackleft=`\[%
+\chardef\isacharbackslash=`\\%
+\chardef\isacharbrackright=`\]%
+\chardef\isacharcircum=`\^%
+\chardef\isacharunderscore=`\_%
+\def\isacharunderscorekeyword{\_}%
+\chardef\isacharbackquote=`\`%
+\chardef\isacharbackquoteopen=`\`%
+\chardef\isacharbackquoteclose=`\`%
+\chardef\isacharbraceleft=`\{%
+\chardef\isacharbar=`\|%
+\chardef\isacharbraceright=`\}%
+\chardef\isachartilde=`\~%
+\def\isacharverbatimopen{\isacharbraceleft\isacharasterisk}%
+\def\isacharverbatimclose{\isacharasterisk\isacharbraceright}%
+\def\isacartoucheopen{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\langle$}}}%
+\def\isacartoucheclose{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\rangle$}}}%
+}
+
+
+% keyword and section markup
+
+\newcommand{\isakeyword}[1]
+{\emph{\bf\def\isachardot{.}\def\isacharunderscore{\isacharunderscorekeyword}%
+\def\isacharbraceleft{\{}\def\isacharbraceright{\}}#1}}
+\newcommand{\isacommand}[1]{\isakeyword{#1}}
+
+\newcommand{\isamarkupheader}[1]{\section{#1}}
+\newcommand{\isamarkupchapter}[1]{\chapter{#1}}
+\newcommand{\isamarkupsection}[1]{\section{#1}}
+\newcommand{\isamarkupsubsection}[1]{\subsection{#1}}
+\newcommand{\isamarkupsubsubsection}[1]{\subsubsection{#1}}
+\newcommand{\isamarkupparagraph}[1]{\paragraph{#1}}
+\newcommand{\isamarkupsubparagraph}[1]{\subparagraph{#1}}
+
+\newif\ifisamarkup
+\newcommand{\isabeginpar}{\par\ifisamarkup\relax\else\medskip\fi}
+\newcommand{\isaendpar}{\par\medskip}
+\newenvironment{isapar}{\parindent\isa@parindent\parskip\isa@parskip\isabeginpar}{\isaendpar}
+\newenvironment{isamarkuptext}{\par\isastyletext\begin{isapar}}{\end{isapar}}
+\newenvironment{isamarkuptxt}{\par\isastyletxt\begin{isapar}}{\end{isapar}}
+\newcommand{\isamarkupcmt}[1]{{\isastylecmt--- #1}}
+
+
+% styles
+
+\def\isabellestyle#1{\csname isabellestyle#1\endcsname}
+
+\newcommand{\isabellestyledefault}{%
+\def\isastyle{\small\tt\slshape}%
+\def\isastylett{\small\tt}%
+\def\isastyleminor{\small\tt\slshape}%
+\def\isastyleminortt{\small\tt}%
+\def\isastylescript{\footnotesize\tt\slshape}%
+\isachardefaults%
+}
+\isabellestyledefault
+
+\newcommand{\isabellestylett}{%
+\def\isastyle{\small\tt}%
+\def\isastylett{\small\tt}%
+\def\isastyleminor{\small\tt}%
+\def\isastyleminortt{\small\tt}%
+\def\isastylescript{\footnotesize\tt}%
+\isachardefaults%
+}
+
+\newcommand{\isabellestyleit}{%
+\def\isastyle{\small\it}%
+\def\isastylett{\small\tt}%
+\def\isastyleminor{\it}%
+\def\isastyleminortt{\tt}%
+\def\isastylescript{\footnotesize\it}%
+\isachardefaults%
+\def\isacharunderscorekeyword{\mbox{-}}%
+\def\isacharbang{\isamath{!}}%
+\def\isachardoublequote{}%
+\def\isachardoublequoteopen{}%
+\def\isachardoublequoteclose{}%
+\def\isacharhash{\isamath{\#}}%
+\def\isachardollar{\isamath{\$}}%
+\def\isacharpercent{\isamath{\%}}%
+\def\isacharampersand{\isamath{\&}}%
+\def\isacharprime{\isamath{\mskip2mu{'}\mskip-2mu}}%
+\def\isacharparenleft{\isamath{(}}%
+\def\isacharparenright{\isamath{)}}%
+\def\isacharasterisk{\isamath{*}}%
+\def\isacharplus{\isamath{+}}%
+\def\isacharcomma{\isamath{\mathord,}}%
+\def\isacharminus{\isamath{-}}%
+\def\isachardot{\isamath{\mathord.}}%
+\def\isacharslash{\isamath{/}}%
+\def\isacharcolon{\isamath{\mathord:}}%
+\def\isacharsemicolon{\isamath{\mathord;}}%
+\def\isacharless{\isamath{<}}%
+\def\isacharequal{\isamath{=}}%
+\def\isachargreater{\isamath{>}}%
+\def\isacharat{\isamath{@}}%
+\def\isacharbrackleft{\isamath{[}}%
+\def\isacharbackslash{\isamath{\backslash}}%
+\def\isacharbrackright{\isamath{]}}%
+\def\isacharunderscore{\mbox{-}}%
+\def\isacharbraceleft{\isamath{\{}}%
+\def\isacharbar{\isamath{\mid}}%
+\def\isacharbraceright{\isamath{\}}}%
+\def\isachartilde{\isamath{{}\sp{\sim}}}%
+\def\isacharbackquoteopen{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\langle$}}}%
+\def\isacharbackquoteclose{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\rangle$}}}%
+\def\isacharverbatimopen{\isamath{\langle\!\langle}}%
+\def\isacharverbatimclose{\isamath{\rangle\!\rangle}}%
+}
+
+\newcommand{\isabellestyleliteral}{%
+\isabellestyleit%
+\def\isacharunderscore{\_}%
+\def\isacharunderscorekeyword{\_}%
+\chardef\isacharbackquoteopen=`\`%
+\chardef\isacharbackquoteclose=`\`%
+}
+
+\newcommand{\isabellestyleliteralunderscore}{%
+\isabellestyleliteral%
+\def\isacharunderscore{\textunderscore}%
+\def\isacharunderscorekeyword{\textunderscore}%
+}
+
+\newcommand{\isabellestylesl}{%
+\isabellestyleit%
+\def\isastyle{\small\sl}%
+\def\isastylett{\small\tt}%
+\def\isastyleminor{\sl}%
+\def\isastyleminortt{\tt}%
+\def\isastylescript{\footnotesize\sl}%
+}
+
+
+% tagged regions
+
+%plain TeX version of comment package -- much faster!
+\let\isafmtname\fmtname\def\fmtname{plain}
+\usepackage{comment}
+\let\fmtname\isafmtname
+
+\newcommand{\isafold}[1]{\emph{$\langle\mathord{\mathit{#1}}\rangle$}}
+
+\newcommand{\isakeeptag}[1]%
+{\includecomment{isadelim#1}\includecomment{isatag#1}\csarg\def{isafold#1}{}}
+\newcommand{\isadroptag}[1]%
+{\excludecomment{isadelim#1}\excludecomment{isatag#1}\csarg\def{isafold#1}{}}
+\newcommand{\isafoldtag}[1]%
+{\includecomment{isadelim#1}\excludecomment{isatag#1}\csarg\def{isafold#1}{\isafold{#1}}}
+
+\isakeeptag{theory}
+\isakeeptag{proof}
+\isakeeptag{ML}
+\isakeeptag{visible}
+\isadroptag{invisible}
+
+\IfFileExists{isabelletags.sty}{\usepackage{isabelletags}}{}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/tex/document/isabellesym.sty Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,372 @@
+%%
+%% definitions of standard Isabelle symbols
+%%
+
+\newcommand{\isasymzero}{\isamath{\mathbf{0}}} %requires amssymb
+\newcommand{\isasymone}{\isamath{\mathbf{1}}} %requires amssymb
+\newcommand{\isasymtwo}{\isamath{\mathbf{2}}} %requires amssymb
+\newcommand{\isasymthree}{\isamath{\mathbf{3}}} %requires amssymb
+\newcommand{\isasymfour}{\isamath{\mathbf{4}}} %requires amssymb
+\newcommand{\isasymfive}{\isamath{\mathbf{5}}} %requires amssymb
+\newcommand{\isasymsix}{\isamath{\mathbf{6}}} %requires amssymb
+\newcommand{\isasymseven}{\isamath{\mathbf{7}}} %requires amssymb
+\newcommand{\isasymeight}{\isamath{\mathbf{8}}} %requires amssymb
+\newcommand{\isasymnine}{\isamath{\mathbf{9}}} %requires amssymb
+\newcommand{\isasymA}{\isamath{\mathcal{A}}}
+\newcommand{\isasymB}{\isamath{\mathcal{B}}}
+\newcommand{\isasymC}{\isamath{\mathcal{C}}}
+\newcommand{\isasymD}{\isamath{\mathcal{D}}}
+\newcommand{\isasymE}{\isamath{\mathcal{E}}}
+\newcommand{\isasymF}{\isamath{\mathcal{F}}}
+\newcommand{\isasymG}{\isamath{\mathcal{G}}}
+\newcommand{\isasymH}{\isamath{\mathcal{H}}}
+\newcommand{\isasymI}{\isamath{\mathcal{I}}}
+\newcommand{\isasymJ}{\isamath{\mathcal{J}}}
+\newcommand{\isasymK}{\isamath{\mathcal{K}}}
+\newcommand{\isasymL}{\isamath{\mathcal{L}}}
+\newcommand{\isasymM}{\isamath{\mathcal{M}}}
+\newcommand{\isasymN}{\isamath{\mathcal{N}}}
+\newcommand{\isasymO}{\isamath{\mathcal{O}}}
+\newcommand{\isasymP}{\isamath{\mathcal{P}}}
+\newcommand{\isasymQ}{\isamath{\mathcal{Q}}}
+\newcommand{\isasymR}{\isamath{\mathcal{R}}}
+\newcommand{\isasymS}{\isamath{\mathcal{S}}}
+\newcommand{\isasymT}{\isamath{\mathcal{T}}}
+\newcommand{\isasymU}{\isamath{\mathcal{U}}}
+\newcommand{\isasymV}{\isamath{\mathcal{V}}}
+\newcommand{\isasymW}{\isamath{\mathcal{W}}}
+\newcommand{\isasymX}{\isamath{\mathcal{X}}}
+\newcommand{\isasymY}{\isamath{\mathcal{Y}}}
+\newcommand{\isasymZ}{\isamath{\mathcal{Z}}}
+\newcommand{\isasyma}{\isamath{\mathrm{a}}}
+\newcommand{\isasymb}{\isamath{\mathrm{b}}}
+\newcommand{\isasymc}{\isamath{\mathrm{c}}}
+\newcommand{\isasymd}{\isamath{\mathrm{d}}}
+\newcommand{\isasyme}{\isamath{\mathrm{e}}}
+\newcommand{\isasymf}{\isamath{\mathrm{f}}}
+\newcommand{\isasymg}{\isamath{\mathrm{g}}}
+\newcommand{\isasymh}{\isamath{\mathrm{h}}}
+\newcommand{\isasymi}{\isamath{\mathrm{i}}}
+\newcommand{\isasymj}{\isamath{\mathrm{j}}}
+\newcommand{\isasymk}{\isamath{\mathrm{k}}}
+\newcommand{\isasyml}{\isamath{\mathrm{l}}}
+\newcommand{\isasymm}{\isamath{\mathrm{m}}}
+\newcommand{\isasymn}{\isamath{\mathrm{n}}}
+\newcommand{\isasymo}{\isamath{\mathrm{o}}}
+\newcommand{\isasymp}{\isamath{\mathrm{p}}}
+\newcommand{\isasymq}{\isamath{\mathrm{q}}}
+\newcommand{\isasymr}{\isamath{\mathrm{r}}}
+\newcommand{\isasyms}{\isamath{\mathrm{s}}}
+\newcommand{\isasymt}{\isamath{\mathrm{t}}}
+\newcommand{\isasymu}{\isamath{\mathrm{u}}}
+\newcommand{\isasymv}{\isamath{\mathrm{v}}}
+\newcommand{\isasymw}{\isamath{\mathrm{w}}}
+\newcommand{\isasymx}{\isamath{\mathrm{x}}}
+\newcommand{\isasymy}{\isamath{\mathrm{y}}}
+\newcommand{\isasymz}{\isamath{\mathrm{z}}}
+\newcommand{\isasymAA}{\isamath{\mathfrak{A}}} %requires eufrak
+\newcommand{\isasymBB}{\isamath{\mathfrak{B}}} %requires eufrak
+\newcommand{\isasymCC}{\isamath{\mathfrak{C}}} %requires eufrak
+\newcommand{\isasymDD}{\isamath{\mathfrak{D}}} %requires eufrak
+\newcommand{\isasymEE}{\isamath{\mathfrak{E}}} %requires eufrak
+\newcommand{\isasymFF}{\isamath{\mathfrak{F}}} %requires eufrak
+\newcommand{\isasymGG}{\isamath{\mathfrak{G}}} %requires eufrak
+\newcommand{\isasymHH}{\isamath{\mathfrak{H}}} %requires eufrak
+\newcommand{\isasymII}{\isamath{\mathfrak{I}}} %requires eufrak
+\newcommand{\isasymJJ}{\isamath{\mathfrak{J}}} %requires eufrak
+\newcommand{\isasymKK}{\isamath{\mathfrak{K}}} %requires eufrak
+\newcommand{\isasymLL}{\isamath{\mathfrak{L}}} %requires eufrak
+\newcommand{\isasymMM}{\isamath{\mathfrak{M}}} %requires eufrak
+\newcommand{\isasymNN}{\isamath{\mathfrak{N}}} %requires eufrak
+\newcommand{\isasymOO}{\isamath{\mathfrak{O}}} %requires eufrak
+\newcommand{\isasymPP}{\isamath{\mathfrak{P}}} %requires eufrak
+\newcommand{\isasymQQ}{\isamath{\mathfrak{Q}}} %requires eufrak
+\newcommand{\isasymRR}{\isamath{\mathfrak{R}}} %requires eufrak
+\newcommand{\isasymSS}{\isamath{\mathfrak{S}}} %requires eufrak
+\newcommand{\isasymTT}{\isamath{\mathfrak{T}}} %requires eufrak
+\newcommand{\isasymUU}{\isamath{\mathfrak{U}}} %requires eufrak
+\newcommand{\isasymVV}{\isamath{\mathfrak{V}}} %requires eufrak
+\newcommand{\isasymWW}{\isamath{\mathfrak{W}}} %requires eufrak
+\newcommand{\isasymXX}{\isamath{\mathfrak{X}}} %requires eufrak
+\newcommand{\isasymYY}{\isamath{\mathfrak{Y}}} %requires eufrak
+\newcommand{\isasymZZ}{\isamath{\mathfrak{Z}}} %requires eufrak
+\newcommand{\isasymaa}{\isamath{\mathfrak{a}}} %requires eufrak
+\newcommand{\isasymbb}{\isamath{\mathfrak{b}}} %requires eufrak
+\newcommand{\isasymcc}{\isamath{\mathfrak{c}}} %requires eufrak
+\newcommand{\isasymdd}{\isamath{\mathfrak{d}}} %requires eufrak
+\newcommand{\isasymee}{\isamath{\mathfrak{e}}} %requires eufrak
+\newcommand{\isasymff}{\isamath{\mathfrak{f}}} %requires eufrak
+\newcommand{\isasymgg}{\isamath{\mathfrak{g}}} %requires eufrak
+\newcommand{\isasymhh}{\isamath{\mathfrak{h}}} %requires eufrak
+\newcommand{\isasymii}{\isamath{\mathfrak{i}}} %requires eufrak
+\newcommand{\isasymjj}{\isamath{\mathfrak{j}}} %requires eufrak
+\newcommand{\isasymkk}{\isamath{\mathfrak{k}}} %requires eufrak
+\newcommand{\isasymll}{\isamath{\mathfrak{l}}} %requires eufrak
+\newcommand{\isasymmm}{\isamath{\mathfrak{m}}} %requires eufrak
+\newcommand{\isasymnn}{\isamath{\mathfrak{n}}} %requires eufrak
+\newcommand{\isasymoo}{\isamath{\mathfrak{o}}} %requires eufrak
+\newcommand{\isasympp}{\isamath{\mathfrak{p}}} %requires eufrak
+\newcommand{\isasymqq}{\isamath{\mathfrak{q}}} %requires eufrak
+\newcommand{\isasymrr}{\isamath{\mathfrak{r}}} %requires eufrak
+\newcommand{\isasymss}{\isamath{\mathfrak{s}}} %requires eufrak
+\newcommand{\isasymtt}{\isamath{\mathfrak{t}}} %requires eufrak
+\newcommand{\isasymuu}{\isamath{\mathfrak{u}}} %requires eufrak
+\newcommand{\isasymvv}{\isamath{\mathfrak{v}}} %requires eufrak
+\newcommand{\isasymww}{\isamath{\mathfrak{w}}} %requires eufrak
+\newcommand{\isasymxx}{\isamath{\mathfrak{x}}} %requires eufrak
+\newcommand{\isasymyy}{\isamath{\mathfrak{y}}} %requires eufrak
+\newcommand{\isasymzz}{\isamath{\mathfrak{z}}} %requires eufrak
+\newcommand{\isasymalpha}{\isamath{\alpha}}
+\newcommand{\isasymbeta}{\isamath{\beta}}
+\newcommand{\isasymgamma}{\isamath{\gamma}}
+\newcommand{\isasymdelta}{\isamath{\delta}}
+\newcommand{\isasymepsilon}{\isamath{\varepsilon}}
+\newcommand{\isasymzeta}{\isamath{\zeta}}
+\newcommand{\isasymeta}{\isamath{\eta}}
+\newcommand{\isasymtheta}{\isamath{\vartheta}}
+\newcommand{\isasymiota}{\isamath{\iota}}
+\newcommand{\isasymkappa}{\isamath{\kappa}}
+\newcommand{\isasymlambda}{\isamath{\lambda}}
+\newcommand{\isasymmu}{\isamath{\mu}}
+\newcommand{\isasymnu}{\isamath{\nu}}
+\newcommand{\isasymxi}{\isamath{\xi}}
+\newcommand{\isasympi}{\isamath{\pi}}
+\newcommand{\isasymrho}{\isamath{\varrho}}
+\newcommand{\isasymsigma}{\isamath{\sigma}}
+\newcommand{\isasymtau}{\isamath{\tau}}
+\newcommand{\isasymupsilon}{\isamath{\upsilon}}
+\newcommand{\isasymphi}{\isamath{\varphi}}
+\newcommand{\isasymchi}{\isamath{\chi}}
+\newcommand{\isasympsi}{\isamath{\psi}}
+\newcommand{\isasymomega}{\isamath{\omega}}
+\newcommand{\isasymGamma}{\isamath{\Gamma}}
+\newcommand{\isasymDelta}{\isamath{\Delta}}
+\newcommand{\isasymTheta}{\isamath{\Theta}}
+\newcommand{\isasymLambda}{\isamath{\Lambda}}
+\newcommand{\isasymXi}{\isamath{\Xi}}
+\newcommand{\isasymPi}{\isamath{\Pi}}
+\newcommand{\isasymSigma}{\isamath{\Sigma}}
+\newcommand{\isasymUpsilon}{\isamath{\Upsilon}}
+\newcommand{\isasymPhi}{\isamath{\Phi}}
+\newcommand{\isasymPsi}{\isamath{\Psi}}
+\newcommand{\isasymOmega}{\isamath{\Omega}}
+\newcommand{\isasymbool}{\isamath{\mathrm{I}\mkern-3.8mu\mathrm{B}}}
+\newcommand{\isasymcomplex}{\isamath{\mathrm{C}\mkern-15mu{\phantom{\mathrm{t}}\vrule}\mkern9mu}}
+\newcommand{\isasymnat}{\isamath{\mathrm{I}\mkern-3.8mu\mathrm{N}}}
+\newcommand{\isasymrat}{\isamath{\mathrm{Q}\mkern-16mu{\phantom{\mathrm{t}}\vrule}\mkern10mu}}
+\newcommand{\isasymreal}{\isamath{\mathrm{I}\mkern-3.8mu\mathrm{R}}}
+\newcommand{\isasymint}{\isamath{\mathsf{Z}\mkern-7.5mu\mathsf{Z}}}
+\newcommand{\isasymleftarrow}{\isamath{\leftarrow}}
+\newcommand{\isasymrightarrow}{\isamath{\rightarrow}}
+\newcommand{\isasymlongleftarrow}{\isamath{\longleftarrow}}
+\newcommand{\isasymlongrightarrow}{\isamath{\longrightarrow}}
+\newcommand{\isasymlonglongleftarrow}{\isamath{\xleftarrow{\hphantom{AAA}}}} %requires amsmath
+\newcommand{\isasymlonglongrightarrow}{\isamath{\xrightarrow{\hphantom{AAA}}}} %requires amsmath
+\newcommand{\isasymlonglonglongleftarrow}{\isamath{\xleftarrow{\hphantom{AAAA}}}} %requires amsmath
+\newcommand{\isasymlonglonglongrightarrow}{\isamath{\xrightarrow{\hphantom{AAAA}}}} %requires amsmath
+\newcommand{\isasymLeftarrow}{\isamath{\Leftarrow}}
+\newcommand{\isasymRightarrow}{\isamath{\Rightarrow}}
+\newcommand{\isasymLongleftarrow}{\isamath{\Longleftarrow}}
+\newcommand{\isasymLongrightarrow}{\isamath{\Longrightarrow}}
+\newcommand{\isasymLleftarrow}{\isamath{\Lleftarrow}} %requires amssymb
+\newcommand{\isasymRrightarrow}{\isamath{\Rrightarrow}} %requires amssymb
+\newcommand{\isasymleftrightarrow}{\isamath{\leftrightarrow}}
+\newcommand{\isasymLeftrightarrow}{\isamath{\Leftrightarrow}}
+\newcommand{\isasymlongleftrightarrow}{\isamath{\longleftrightarrow}}
+\newcommand{\isasymLongleftrightarrow}{\isamath{\Longleftrightarrow}}
+\newcommand{\isasymmapsto}{\isamath{\mapsto}}
+\newcommand{\isasymlongmapsto}{\isamath{\longmapsto}}
+\newcommand{\isasymmidarrow}{\isamath{\relbar}}
+\newcommand{\isasymMidarrow}{\isamath{\Relbar}}
+\newcommand{\isasymhookleftarrow}{\isamath{\hookleftarrow}}
+\newcommand{\isasymhookrightarrow}{\isamath{\hookrightarrow}}
+\newcommand{\isasymleftharpoondown}{\isamath{\leftharpoondown}}
+\newcommand{\isasymrightharpoondown}{\isamath{\rightharpoondown}}
+\newcommand{\isasymleftharpoonup}{\isamath{\leftharpoonup}}
+\newcommand{\isasymrightharpoonup}{\isamath{\rightharpoonup}}
+\newcommand{\isasymrightleftharpoons}{\isamath{\rightleftharpoons}}
+\newcommand{\isasymleadsto}{\isamath{\leadsto}} %requires amssymb
+\newcommand{\isasymdownharpoonleft}{\isamath{\downharpoonleft}} %requires amssymb
+\newcommand{\isasymdownharpoonright}{\isamath{\downharpoonright}} %requires amssymb
+\newcommand{\isasymupharpoonleft}{\isamath{\upharpoonleft}} %requires amssymb
+\newcommand{\isasymupharpoonright}{\isamath{\upharpoonright}} %requires amssymb
+\newcommand{\isasymrestriction}{\isamath{\restriction}} %requires amssymb
+\newcommand{\isasymColon}{\isamath{\mathrel{::}}}
+\newcommand{\isasymup}{\isamath{\uparrow}}
+\newcommand{\isasymUp}{\isamath{\Uparrow}}
+\newcommand{\isasymdown}{\isamath{\downarrow}}
+\newcommand{\isasymDown}{\isamath{\Downarrow}}
+\newcommand{\isasymupdown}{\isamath{\updownarrow}}
+\newcommand{\isasymUpdown}{\isamath{\Updownarrow}}
+\newcommand{\isasymlangle}{\isamath{\langle}}
+\newcommand{\isasymrangle}{\isamath{\rangle}}
+\newcommand{\isasymlceil}{\isamath{\lceil}}
+\newcommand{\isasymrceil}{\isamath{\rceil}}
+\newcommand{\isasymlfloor}{\isamath{\lfloor}}
+\newcommand{\isasymrfloor}{\isamath{\rfloor}}
+\newcommand{\isasymlparr}{\isamath{\mathopen{(\mkern-3mu\mid}}}
+\newcommand{\isasymrparr}{\isamath{\mathclose{\mid\mkern-3mu)}}}
+\newcommand{\isasymlbrakk}{\isamath{\mathopen{\lbrack\mkern-3mu\lbrack}}}
+\newcommand{\isasymrbrakk}{\isamath{\mathclose{\rbrack\mkern-3mu\rbrack}}}
+\newcommand{\isasymlbrace}{\isamath{\mathopen{\lbrace\mkern-4.5mu\mid}}}
+\newcommand{\isasymrbrace}{\isamath{\mathclose{\mid\mkern-4.5mu\rbrace}}}
+\newcommand{\isasymguillemotleft}{\isatext{\flqq}} %requires babel
+\newcommand{\isasymguillemotright}{\isatext{\frqq}} %requires babel
+\newcommand{\isasymbottom}{\isamath{\bot}}
+\newcommand{\isasymtop}{\isamath{\top}}
+\newcommand{\isasymand}{\isamath{\wedge}}
+\newcommand{\isasymAnd}{\isamath{\bigwedge}}
+\newcommand{\isasymor}{\isamath{\vee}}
+\newcommand{\isasymOr}{\isamath{\bigvee}}
+\newcommand{\isasymforall}{\isamath{\forall\,}}
+\newcommand{\isasymexists}{\isamath{\exists\,}}
+\newcommand{\isasymnot}{\isamath{\neg}}
+\newcommand{\isasymnexists}{\isamath{\nexists\,}} %requires amssymb
+\newcommand{\isasymcircle}{\isamath{\ocircle}} %requires wasysym
+\newcommand{\isasymbox}{\isamath{\Box}} %requires amssymb
+\newcommand{\isasymdiamond}{\isamath{\Diamond}} %requires amssymb
+\newcommand{\isasymdiamondop}{\isamath{\diamond}}
+\newcommand{\isasymsurd}{\isamath{\surd}}
+\newcommand{\isasymturnstile}{\isamath{\vdash}}
+\newcommand{\isasymTurnstile}{\isamath{\models}}
+\newcommand{\isasymtturnstile}{\isamath{\vdash\!\!\!\vdash}}
+\newcommand{\isasymTTurnstile}{\isamath{\mid\!\models}}
+\newcommand{\isasymstileturn}{\isamath{\dashv}}
+\newcommand{\isasymle}{\isamath{\le}}
+\newcommand{\isasymge}{\isamath{\ge}}
+\newcommand{\isasymlless}{\isamath{\ll}}
+\newcommand{\isasymggreater}{\isamath{\gg}}
+\newcommand{\isasymlesssim}{\isamath{\lesssim}} %requires amssymb
+\newcommand{\isasymgreatersim}{\isamath{\gtrsim}} %requires amssymb
+\newcommand{\isasymlessapprox}{\isamath{\lessapprox}} %requires amssymb
+\newcommand{\isasymgreaterapprox}{\isamath{\gtrapprox}} %requires amssymb
+\newcommand{\isasymin}{\isamath{\in}}
+\newcommand{\isasymnotin}{\isamath{\notin}}
+\newcommand{\isasymsubset}{\isamath{\subset}}
+\newcommand{\isasymsupset}{\isamath{\supset}}
+\newcommand{\isasymsubseteq}{\isamath{\subseteq}}
+\newcommand{\isasymsupseteq}{\isamath{\supseteq}}
+\newcommand{\isasymsqsubset}{\isamath{\sqsubset}} %requires amssymb
+\newcommand{\isasymsqsupset}{\isamath{\sqsupset}} %requires amssymb
+\newcommand{\isasymsqsubseteq}{\isamath{\sqsubseteq}}
+\newcommand{\isasymsqsupseteq}{\isamath{\sqsupseteq}}
+\newcommand{\isasyminter}{\isamath{\cap}}
+\newcommand{\isasymInter}{\isamath{\bigcap\,}}
+\newcommand{\isasymunion}{\isamath{\cup}}
+\newcommand{\isasymUnion}{\isamath{\bigcup\,}}
+\newcommand{\isasymsqunion}{\isamath{\sqcup}}
+\newcommand{\isasymSqunion}{\isamath{\bigsqcup\,}}
+\newcommand{\isasymsqinter}{\isamath{\sqcap}}
+\newcommand{\isasymSqinter}{\isamath{\bigsqcap\,}} %requires stmaryrd
+\newcommand{\isasymsetminus}{\isamath{\setminus}}
+\newcommand{\isasympropto}{\isamath{\propto}}
+\newcommand{\isasymuplus}{\isamath{\uplus}}
+\newcommand{\isasymUplus}{\isamath{\biguplus\,}}
+\newcommand{\isasymnoteq}{\isamath{\not=}}
+\newcommand{\isasymsim}{\isamath{\sim}}
+\newcommand{\isasymdoteq}{\isamath{\doteq}}
+\newcommand{\isasymsimeq}{\isamath{\simeq}}
+\newcommand{\isasymapprox}{\isamath{\approx}}
+\newcommand{\isasymasymp}{\isamath{\asymp}}
+\newcommand{\isasymcong}{\isamath{\cong}}
+\newcommand{\isasymsmile}{\isamath{\smile}}
+\newcommand{\isasymequiv}{\isamath{\equiv}}
+\newcommand{\isasymfrown}{\isamath{\frown}}
+\newcommand{\isasymJoin}{\isamath{\Join}} %requires amssymb
+\newcommand{\isasymbowtie}{\isamath{\bowtie}}
+\newcommand{\isasymprec}{\isamath{\prec}}
+\newcommand{\isasymsucc}{\isamath{\succ}}
+\newcommand{\isasympreceq}{\isamath{\preceq}}
+\newcommand{\isasymsucceq}{\isamath{\succeq}}
+\newcommand{\isasymparallel}{\isamath{\parallel}}
+\newcommand{\isasymbar}{\isamath{\mid}}
+\newcommand{\isasymplusminus}{\isamath{\pm}}
+\newcommand{\isasymminusplus}{\isamath{\mp}}
+\newcommand{\isasymtimes}{\isamath{\times}}
+\newcommand{\isasymdiv}{\isamath{\div}}
+\newcommand{\isasymcdot}{\isamath{\cdot}}
+\newcommand{\isasymstar}{\isamath{\star}}
+\newcommand{\isasymbullet}{\boldmath\isamath{\mathchoice{\displaystyle{\cdot}}{\textstyle{\cdot}}{\scriptstyle{\bullet}}{\scriptscriptstyle{\bullet}}}}
+\newcommand{\isasymcirc}{\isamath{\circ}}
+\newcommand{\isasymdagger}{\isamath{\dagger}}
+\newcommand{\isasymddagger}{\isamath{\ddagger}}
+\newcommand{\isasymlhd}{\isamath{\lhd}} %requires amssymb
+\newcommand{\isasymrhd}{\isamath{\rhd}} %requires amssymb
+\newcommand{\isasymunlhd}{\isamath{\unlhd}} %requires amssymb
+\newcommand{\isasymunrhd}{\isamath{\unrhd}} %requires amssymb
+\newcommand{\isasymtriangleleft}{\isamath{\triangleleft}}
+\newcommand{\isasymtriangleright}{\isamath{\triangleright}}
+\newcommand{\isasymtriangle}{\isamath{\triangle}}
+\newcommand{\isasymtriangleq}{\isamath{\triangleq}} %requires amssymb
+\newcommand{\isasymoplus}{\isamath{\oplus}}
+\newcommand{\isasymOplus}{\isamath{\bigoplus\,}}
+\newcommand{\isasymotimes}{\isamath{\otimes}}
+\newcommand{\isasymOtimes}{\isamath{\bigotimes\,}}
+\newcommand{\isasymodot}{\isamath{\odot}}
+\newcommand{\isasymOdot}{\isamath{\bigodot\,}}
+\newcommand{\isasymominus}{\isamath{\ominus}}
+\newcommand{\isasymoslash}{\isamath{\oslash}}
+\newcommand{\isasymdots}{\isamath{\dots}}
+\newcommand{\isasymcdots}{\isamath{\cdots}}
+\newcommand{\isasymSum}{\isamath{\sum\,}}
+\newcommand{\isasymProd}{\isamath{\prod\,}}
+\newcommand{\isasymCoprod}{\isamath{\coprod\,}}
+\newcommand{\isasyminfinity}{\isamath{\infty}}
+\newcommand{\isasymintegral}{\isamath{\int\,}}
+\newcommand{\isasymointegral}{\isamath{\oint\,}}
+\newcommand{\isasymclubsuit}{\isamath{\clubsuit}}
+\newcommand{\isasymdiamondsuit}{\isamath{\diamondsuit}}
+\newcommand{\isasymheartsuit}{\isamath{\heartsuit}}
+\newcommand{\isasymspadesuit}{\isamath{\spadesuit}}
+\newcommand{\isasymaleph}{\isamath{\aleph}}
+\newcommand{\isasymemptyset}{\isamath{\emptyset}}
+\newcommand{\isasymnabla}{\isamath{\nabla}}
+\newcommand{\isasympartial}{\isamath{\partial}}
+\newcommand{\isasymRe}{\isamath{\Re}}
+\newcommand{\isasymIm}{\isamath{\Im}}
+\newcommand{\isasymflat}{\isamath{\flat}}
+\newcommand{\isasymnatural}{\isamath{\natural}}
+\newcommand{\isasymsharp}{\isamath{\sharp}}
+\newcommand{\isasymangle}{\isamath{\angle}}
+\newcommand{\isasymcopyright}{\isatext{\rm\copyright}}
+\newcommand{\isasymregistered}{\isatext{\rm\textregistered}}
+\newcommand{\isasyminverse}{\isamath{{}^{-1}}}
+\newcommand{\isasymonequarter}{\isatext{\rm\textonequarter}} %requires textcomp
+\newcommand{\isasymonehalf}{\isatext{\rm\textonehalf}} %requires textcomp
+\newcommand{\isasymthreequarters}{\isatext{\rm\textthreequarters}} %requires textcomp
+\newcommand{\isasymordfeminine}{\isatext{\rm\textordfeminine}}
+\newcommand{\isasymordmasculine}{\isatext{\rm\textordmasculine}}
+\newcommand{\isasymsection}{\isatext{\rm\S}}
+\newcommand{\isasymparagraph}{\isatext{\rm\P}}
+\newcommand{\isasymexclamdown}{\isatext{\rm\textexclamdown}}
+\newcommand{\isasymquestiondown}{\isatext{\rm\textquestiondown}}
+\newcommand{\isasymeuro}{\isatext{\euro}} %requires eurosym
+\newcommand{\isasympounds}{\isamath{\pounds}}
+\newcommand{\isasymyen}{\isatext{\yen}} %requires amssymb
+\newcommand{\isasymcent}{\isatext{\textcent}} %requires textcomp
+\newcommand{\isasymcurrency}{\isatext{\textcurrency}} %requires textcomp
+\newcommand{\isasymdegree}{\isatext{\rm\textdegree}} %requires textcomp
+\newcommand{\isasymhyphen}{\isatext{\rm-}}
+\newcommand{\isasymamalg}{\isamath{\amalg}}
+\newcommand{\isasymmho}{\isamath{\mho}} %requires amssymb
+\newcommand{\isasymlozenge}{\isamath{\lozenge}} %requires amssymb
+\newcommand{\isasymwp}{\isamath{\wp}}
+\newcommand{\isasymwrong}{\isamath{\wr}}
+\newcommand{\isasymacute}{\isatext{\'\relax}}
+\newcommand{\isasymindex}{\isatext{\i}}
+\newcommand{\isasymdieresis}{\isatext{\"\relax}}
+\newcommand{\isasymcedilla}{\isatext{\c\relax}}
+\newcommand{\isasymhungarumlaut}{\isatext{\H\relax}}
+\newcommand{\isasymmodule}{\isamath{\langle}\isakeyword{module}\isamath{\rangle}}
+\newcommand{\isasymsome}{\isamath{\epsilon\,}}
+\newcommand{\isasymhole}{\isatext{\rm\wasylozenge}} %requires wasysym
+\newcommand{\isasymbind}{\isamath{\mathbin{>\!\!\!>\mkern-6.7mu=}}}
+\newcommand{\isasymthen}{\isamath{\mathbin{>\!\!\!>}}}
+\newcommand{\isasymopen}{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\langle$}}}
+\newcommand{\isasymclose}{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\rangle$}}}
+\newcommand{\isasymnewline}{\isatext{\fbox{$\hookleftarrow$}}}
+\newcommand{\isasymcomment}{\isatext{---}}
+\newcommand{\isasymproof}{\isamath{\,\langle\mathit{proof}\rangle}}
+\newcommand{\isactrlundefined}{\isakeyword{undefined}\ }
+\newcommand{\isactrlfile}{\isakeyword{file}\ }
+\newcommand{\isactrldir}{\isakeyword{dir}\ }
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/tex/document/isabelletags.sty Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,1 @@
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/tex/document/pdfsetup.sty Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,7 @@
+%%
+%% default hyperref setup (both for pdf and dvi output)
+%%
+
+\usepackage{color}
+\definecolor{linkcolor}{rgb}{0,0,0.5}
+\usepackage[colorlinks=true,linkcolor=linkcolor,citecolor=linkcolor,filecolor=linkcolor,urlcolor=linkcolor,pdfpagelabels]{hyperref}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/tex/document/prelude.tex Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,104 @@
+%Beamer configuration
+
+% Include page numbers
+\setbeamertemplate{footline}{\hfill\insertframenumber\hspace*{2mm}\vspace*{2mm}}
+
+\usetheme{Pittsburgh}
+
+% switch off the fancy navigation symbols
+\setbeamertemplate{navigation symbols}{}
+
+\usefonttheme[onlymath]{serif}
+\usepackage{isabelle,isabellesym}
+\usepackage{tikz}
+\usetikzlibrary{matrix}
+
+\isabellestyle{it}
+
+% for uniform font size
+\renewcommand{\isastyleminor}{\rmfamily\itshape}
+\renewcommand{\isastyle}{\isastyleminor}
+%\renewcommand{\isastyletext}{\normalsize\sf}
+
+% typeset underscore
+\renewcommand{\isacharunderscore}{\_}%
+% normal _ has no bold version, use tt _
+% both for theory snippets and for \isakeyword uses
+\renewcommand{\isacharunderscorekeyword}{\texttt{\char`_}}%
+\let\isakeywordold=\isakeyword
+\renewcommand{\isakeyword}[1]{\isakeywordold{\def\_{{\texttt{\char`_}}}#1}}
+
+\renewcommand{\isacharbackquoteopen}{\isacharbackquote}
+\renewcommand{\isacharbackquoteclose}{\isacharbackquote}
+\renewcommand{\isasymnabla}{\raisebox{2pt}{\isamath{\normalbigtriangledown}}}
+
+\let\isacharprimeR=\isacharprime
+\newcommand{\isactrlisubprime}[1]{{$\sb{#1}\sp{\mskip3mu\prime}$}}
+
+\newcommand{\xsymbol}[1]{\texttt{\char`\\\char`<#1>}}
+
+\definecolor{darkblue}{rgb}{0.0,0.0,0.5}
+\newcommand<>{\high}[1]{{\color#2{blue}{#1}}}
+\newcommand<>{\bad}[1]{{\color#2{red}{#1}}}
+\newcommand{\gray}[1]{\textcolor{gray}{#1}}
+
+\newcommand{\concept}[1]{\high{\emph{#1}}}
+\newcommand{\important}{{\Large \textbf{!}}}
+
+\newcommand{\filepage}[2]
+{\title{\textcolor{darkblue}{#1}}
+\author{#2}\institute{}\date{}
+\maketitle}
+
+\newcommand{\ttfilepage}[2]
+{\filepage{\texttt{#1}}{#2}}
+
+\newcommand{\thyfile}[2]
+{\ttfilepage{#1.thy}{#2}}
+
+\newcommand{\demofile}[2]
+{\thyfile{#1\char`\_Demo}{#2}}
+
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+\setbeamersize{text margin left=0.5cm,text margin right=0.5cm}
+\renewcommand\partname{Chapter}
+
+\setbeamertemplate{part page}
+{
+ \begin{centering}
+ {\usebeamerfont{part name}\usebeamercolor[fg]{part name}\partname~\insertpartnumber}
+ \vskip1em\par
+ \begin{beamercolorbox}[sep=16pt,center]{part title}
+ \usebeamerfont{part title}\insertpart\par
+ \end{beamercolorbox}
+ \end{centering}
+}
+
+\newcommand{\mypartpage}[2]{
+\begin{frame}
+ \begin{centering}
+ {\usebeamerfont{part name}\usebeamercolor[fg]{part name}Part #1}
+ \vskip1em\par
+ \begin{beamercolorbox}[sep=16pt,center]{part title}
+ \usebeamerfont{part title}#2\par
+ \end{beamercolorbox}
+ \end{centering}
+\end{frame}}
+
+\AtBeginPart{\frame{\partpage}\frame{\setcounter{tocdepth}{1}\tableofcontents}}
+
+\AtBeginSection[]
+{
+\begin{frame}
+\setcounter{tocdepth}{1}
+\tableofcontents[currentsection]
+\end{frame}
+}
+
+\AtBeginSubsection[]
+{
+\begin{frame}
+\setcounter{tocdepth}{2}
+\tableofcontents[sectionstyle=show/hide,subsectionstyle=show/shaded/hide]
+\end{frame}
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/tex/document/railsetup.sty Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,1202 @@
+% rail.sty - style file to support railroad diagrams
+%
+% 09-Jul-90 L. Rooijakkers
+% 08-Oct-90 L. Rooijakkers fixed centering bug when \rail@tmpc<0.
+% 07-Feb-91 L. Rooijakkers added \railoptions command, indexing
+% 08-Feb-91 L. Rooijakkers minor fixes
+% 28-Jun-94 K. Barthelmann turned into LaTeX2e package
+% 08-Dec-96 K. Barthelmann replaced \@writefile
+% 13-Dec-96 K. Barthelmann cleanup
+% 22-Feb-98 K. Barthelmann fixed catcodes of special characters
+% 18-Apr-98 K. Barthelmann fixed \par handling
+% 19-May-98 J. Olsson Added new macros to support arrow heads.
+% 26-Jul-98 K. Barthelmann changed \par to output newlines
+% 02-May-11 M. Wenzel default setup for Isabelle
+%
+% This style file needs to be used in conjunction with the 'rail'
+% program. Running LaTeX as 'latex file' produces file.rai, which should be
+% processed by Rail with 'rail file'. This produces file.rao, which will
+% be picked up by LaTeX on the next 'latex file' run.
+%
+% LaTeX will warn if there is no file.rao or it's out of date.
+%
+% The macros in this file thus consist of two parts: those that read and
+% write the .rai and .rao files, and those that do the actual formatting
+% of the railroad diagrams.
+
+\NeedsTeXFormat{LaTeX2e}
+\ProvidesPackage{rail}[1998/05/19]
+
+% railroad diagram formatting parameters (user level)
+% all of these are copied into their internal versions by \railinit
+%
+% \railunit : \unitlength within railroad diagrams
+% \railextra : extra length at outside of diagram
+% \railboxheight : height of ovals and frames
+% \railboxskip : vertical space between lines
+% \railboxleft : space to the left of a box
+% \railboxright : space to the right of a box
+% \railovalspace : extra space around contents of oval
+% \railframespace : extra space around contents of frame
+% \railtextleft : space to the left of text
+% \railtextright : space to the right of text
+% \railtextup : space to lift text up
+% \railjoinsize : circle size of join/split arcs
+% \railjoinadjust : space to adjust join
+%
+% \railnamesep : separator between name and rule body
+
+\newlength\railunit
+\newlength\railextra
+\newlength\railboxheight
+\newlength\railboxskip
+\newlength\railboxleft
+\newlength\railboxright
+\newlength\railovalspace
+\newlength\railframespace
+\newlength\railtextleft
+\newlength\railtextright
+\newlength\railtextup
+\newlength\railjoinsize
+\newlength\railjoinadjust
+\newlength\railnamesep
+
+% initialize the parameters
+
+\setlength\railunit{1sp}
+\setlength\railextra{4ex}
+\setlength\railboxleft{1ex}
+\setlength\railboxright{1ex}
+\setlength\railovalspace{2ex}
+\setlength\railframespace{2ex}
+\setlength\railtextleft{1ex}
+\setlength\railtextright{1ex}
+\setlength\railjoinadjust{0pt}
+\setlength\railnamesep{1ex}
+
+\DeclareOption{10pt}{
+ \setlength\railboxheight{16pt}
+ \setlength\railboxskip{24pt}
+ \setlength\railtextup{5pt}
+ \setlength\railjoinsize{16pt}
+}
+\DeclareOption{11pt}{
+ \setlength\railboxheight{16pt}
+ \setlength\railboxskip{24pt}
+ \setlength\railtextup{5pt}
+ \setlength\railjoinsize{16pt}
+}
+\DeclareOption{12pt}{
+ \setlength\railboxheight{20pt}
+ \setlength\railboxskip{28pt}
+ \setlength\railtextup{6pt}
+ \setlength\railjoinsize{20pt}
+}
+
+\ExecuteOptions{10pt}
+\ProcessOptions
+
+% internal versions of the formatting parameters
+%
+% \rail@extra : \railextra
+% \rail@boxht : \railboxheight
+% \rail@boxsp : \railboxskip
+% \rail@boxlf : \railboxleft
+% \rail@boxrt : \railboxright
+% \rail@boxhht : \railboxheight / 2
+% \rail@ovalsp : \railovalspace
+% \rail@framesp : \railframespace
+% \rail@textlf : \railtextleft
+% \rail@textrt : \railtextright
+% \rail@textup : \railtextup
+% \rail@joinsz : \railjoinsize
+% \rail@joinhsz : \railjoinsize / 2
+% \rail@joinadj : \railjoinadjust
+%
+% \railinit : internalize all of the parameters.
+
+\newcount\rail@extra
+\newcount\rail@boxht
+\newcount\rail@boxsp
+\newcount\rail@boxlf
+\newcount\rail@boxrt
+\newcount\rail@boxhht
+\newcount\rail@ovalsp
+\newcount\rail@framesp
+\newcount\rail@textlf
+\newcount\rail@textrt
+\newcount\rail@textup
+\newcount\rail@joinsz
+\newcount\rail@joinhsz
+\newcount\rail@joinadj
+
+\newcommand\railinit{
+\rail@extra=\railextra
+\divide\rail@extra by \railunit
+\rail@boxht=\railboxheight
+\divide\rail@boxht by \railunit
+\rail@boxsp=\railboxskip
+\divide\rail@boxsp by \railunit
+\rail@boxlf=\railboxleft
+\divide\rail@boxlf by \railunit
+\rail@boxrt=\railboxright
+\divide\rail@boxrt by \railunit
+\rail@boxhht=\railboxheight
+\divide\rail@boxhht by \railunit
+\divide\rail@boxhht by 2
+\rail@ovalsp=\railovalspace
+\divide\rail@ovalsp by \railunit
+\rail@framesp=\railframespace
+\divide\rail@framesp by \railunit
+\rail@textlf=\railtextleft
+\divide\rail@textlf by \railunit
+\rail@textrt=\railtextright
+\divide\rail@textrt by \railunit
+\rail@textup=\railtextup
+\divide\rail@textup by \railunit
+\rail@joinsz=\railjoinsize
+\divide\rail@joinsz by \railunit
+\rail@joinhsz=\railjoinsize
+\divide\rail@joinhsz by \railunit
+\divide\rail@joinhsz by 2
+\rail@joinadj=\railjoinadjust
+\divide\rail@joinadj by \railunit
+}
+
+\AtBeginDocument{\railinit}
+
+% \rail@param : declarations for list environment
+%
+% \railparam{TEXT} : sets \rail@param to TEXT
+%
+% \rail@reserved : characters reserved for grammar
+
+\newcommand\railparam[1]{
+\def\rail@param{
+ \setlength\leftmargin{0pt}\setlength\rightmargin{0pt}
+ \setlength\labelwidth{0pt}\setlength\labelsep{0pt}
+ \setlength\itemindent{0pt}\setlength\listparindent{0pt}
+ #1
+}
+}
+\railparam{}
+
+\newtoks\rail@reserved
+\rail@reserved={:;|*+?[]()'"}
+
+% \rail@termfont : format setup for terminals
+%
+% \rail@nontfont : format setup for nonterminals
+%
+% \rail@annofont : format setup for annotations
+%
+% \rail@rulefont : format setup for rule names
+%
+% \rail@indexfont : format setup for index entry
+%
+% \railtermfont{TEXT} : set terminal format setup to TEXT
+%
+% \railnontermfont{TEXT} : set nonterminal format setup to TEXT
+%
+% \railannotatefont{TEXT} : set annotation format setup to TEXT
+%
+% \railnamefont{TEXT} : set rule name format setup to TEXT
+%
+% \railindexfont{TEXT} : set index entry format setup to TEXT
+
+\def\rail@termfont{\ttfamily\upshape}
+\def\rail@nontfont{\rmfamily\upshape}
+\def\rail@annofont{\rmfamily\itshape}
+\def\rail@namefont{\rmfamily\itshape}
+\def\rail@indexfont{\rmfamily\itshape}
+
+\newcommand\railtermfont[1]{
+\def\rail@termfont{#1}
+}
+
+\newcommand\railnontermfont[1]{
+\def\rail@nontfont{#1}
+}
+
+\newcommand\railannotatefont[1]{
+\def\rail@annofont{#1}
+}
+
+\newcommand\railnamefont[1]{
+\def\rail@namefont{#1}
+}
+
+\newcommand\railindexfont[1]{
+\def\rail@indexfont{#1}
+}
+
+% railroad read/write macros
+%
+% \begin{rail} TEXT \end{rail} : TEXT is written out to the .rai file,
+% as \rail@i{NR}{TEXT}. Then the matching
+% \rail@o{NR}{FMT} from the .rao file is
+% executed (if defined).
+%
+% \railoptions{OPTIONS} : OPTIONS are written out to the .rai file,
+% as \rail@p{OPTIONS}.
+%
+% \railterm{IDENT,IDENT,...} : format IDENT as terminals. writes out
+% \rail@t{IDENT} to the .rai file
+%
+% \railalias{IDENT}{TEXT} : format IDENT as TEXT. defines \rail@t@IDENT as
+% TEXT.
+%
+% \railtoken{IDENT}{TEXT} : abbreviates \railalias{IDENT}{TEXT}\railterm{IDENT}
+% (for backward compatibility)
+%
+% \rail@setcodes : guards special characters
+%
+% \rail@makeother{CHARACTER} : sets \catcode of CHARACTER to "other"
+% used inside a loop for \rail@setcodes
+%
+% \rail@nr : railroad diagram counter
+%
+% \ifrail@match : current \rail@i{NR}{TEXT} matches
+%
+% \rail@first : actions to be done first. read in .rao file,
+% open .rai file if \@filesw true, undefine \rail@first.
+% executed from \begin{rail}, \railoptions and \railterm.
+%
+% \rail@i{NR}{TEXT} : defines \rail@i@NR as TEXT. written to the .rai
+% file by \rail, read from the .rao file by
+% \rail@first
+%
+% \rail@t{IDENT} : tells Rail that IDENT is to be custom formatted,
+% written to the .rai file by \railterm.
+%
+% \rail@o{NR}{TEXT} : defines \rail@o@NR as TEXT, read from the .rao
+% file by \rail@first.
+%
+% \rail@p{OPTIONS} : pass options to rail, written to the .rai file by
+% \railoptions
+%
+% \rail@write{TEXT} : write TEXT to the .rai file
+%
+% \rail@warn : warn user for mismatching diagrams
+%
+% \rail@endwarn : either \relax or \rail@warn
+%
+% \ifrail@all : checked at the end of the document
+
+\def\rail@makeother#1{
+ \expandafter\catcode\expandafter`\csname\string #1\endcsname=12
+}
+
+\def\rail@setcodes{
+\let\par=\relax
+\let\\=\relax
+\expandafter\@tfor\expandafter\rail@symbol\expandafter:\expandafter=%
+ \the\rail@reserved
+\do{\expandafter\rail@makeother\rail@symbol}
+}
+
+\newcount\rail@nr
+
+\newif\ifrail@all
+\rail@alltrue
+
+\newif\ifrail@match
+
+\def\rail@first{
+\begingroup
+\makeatletter
+\rail@setcodes
+\InputIfFileExists{\jobname.rao}{}{\PackageInfo{rail}{No file \jobname.rao}}
+\makeatother
+\endgroup
+\if@filesw
+\newwrite\tf@rai
+\immediate\openout\tf@rai=\jobname.rai
+\fi
+\global\let\rail@first=\relax
+}
+
+\long\def\rail@body#1\end{
+{
+\newlinechar=`^^J
+\def\par{\string\par^^J}
+\rail@write{\string\rail@i{\number\rail@nr}{#1}}
+}
+\xdef\rail@i@{#1}
+\end
+}
+
+\newenvironment{rail}{
+\global\advance\rail@nr by 1
+\rail@first
+\begingroup
+\rail@setcodes
+\rail@body
+}{
+\endgroup
+\rail@matchtrue
+\@ifundefined{rail@o@\number\rail@nr}{\rail@matchfalse}{}
+\expandafter\ifx\csname rail@i@\number\rail@nr\endcsname\rail@i@
+\else
+\rail@matchfalse
+\fi
+\ifrail@match
+\csname rail@o@\number\rail@nr\endcsname
+\else
+\PackageWarning{rail}{Railroad diagram {\number\rail@nr} doesn't match}
+\global\let\rail@endwarn=\rail@warn
+\begin{list}{}{\rail@param}
+\rail@begin{1}{}
+\rail@setbox{\bfseries ???}
+\rail@oval
+\rail@end
+\end{list}
+\fi
+}
+
+\newcommand\railoptions[1]{
+\rail@first
+\rail@write{\string\rail@p{#1}}
+}
+
+\newcommand\railterm[1]{
+\rail@first
+\@for\rail@@:=#1\do{
+\rail@write{\string\rail@t{\rail@@}}
+}
+}
+
+\newcommand\railalias[2]{
+\expandafter\def\csname rail@t@#1\endcsname{#2}
+}
+
+\newcommand\railtoken[2]{\railalias{#1}{#2}\railterm{#1}}
+
+\long\def\rail@i#1#2{
+\expandafter\gdef\csname rail@i@#1\endcsname{#2}
+}
+
+\def\rail@o#1#2{
+\expandafter\gdef\csname rail@o@#1\endcsname{
+\begin{list}{}{\rail@param}
+#2
+\end{list}
+}
+}
+
+\def\rail@t#1{}
+
+\def\rail@p#1{}
+
+\long\def\rail@write#1{\@ifundefined{tf@rai}{}{\immediate\write\tf@rai{#1}}}
+
+\def\rail@warn{
+\PackageWarningNoLine{rail}{Railroad diagram(s) may have changed.
+ Use 'rail' and rerun}
+}
+
+\let\rail@endwarn=\relax
+
+\AtEndDocument{\rail@endwarn}
+
+% index entry macro
+%
+% \rail@index{IDENT} : add index entry for IDENT
+
+\def\rail@index#1{
+\index{\rail@indexfont#1}
+}
+
+% railroad formatting primitives
+%
+% \rail@x : current x
+% \rail@y : current y
+% \rail@ex : current end x
+% \rail@sx : starting x for \rail@cr
+% \rail@rx : rightmost previous x for \rail@cr
+%
+% \rail@tmpa : temporary count
+% \rail@tmpb : temporary count
+% \rail@tmpc : temporary count
+%
+% \rail@put : put at (\rail@x,\rail@y)
+% \rail@vput : put vector at (\rail@x,\rail@y)
+%
+% \rail@eline : end line by drawing from \rail@ex to \rail@x
+%
+% \rail@vreline : end line by drawing a vector from \rail@x to \rail@ex
+%
+% \rail@vleline : end line by drawing a vector from \rail@ex to \rail@x
+%
+% \rail@sety{LEVEL} : set \rail@y to level LEVEL
+
+\newcount\rail@x
+\newcount\rail@y
+\newcount\rail@ex
+\newcount\rail@sx
+\newcount\rail@rx
+
+\newcount\rail@tmpa
+\newcount\rail@tmpb
+\newcount\rail@tmpc
+
+\def\rail@put{\put(\number\rail@x,\number\rail@y)}
+
+\def\rail@vput{\put(\number\rail@ex,\number\rail@y)}
+
+\def\rail@eline{
+\rail@tmpb=\rail@x
+\advance\rail@tmpb by -\rail@ex
+\rail@put{\line(-1,0){\number\rail@tmpb}}
+}
+
+\def\rail@vreline{
+\rail@tmpb=\rail@x
+\advance\rail@tmpb by -\rail@ex
+\rail@vput{\vector(1,0){\number\rail@tmpb}}
+}
+
+\def\rail@vleline{
+\rail@tmpb=\rail@x
+\advance\rail@tmpb by -\rail@ex
+\rail@put{\vector(-1,0){\number\rail@tmpb}}
+}
+
+\def\rail@sety#1{
+\rail@y=#1
+\multiply\rail@y by -\rail@boxsp
+\advance\rail@y by -\rail@boxht
+}
+
+% \rail@begin{HEIGHT}{NAME} : begin a railroad diagram of height HEIGHT
+%
+% \rail@end : end a railroad diagram
+%
+% \rail@expand{IDENT} : expand IDENT
+
+\def\rail@begin#1#2{
+\item
+\begin{minipage}[t]{\linewidth}
+\ifx\@empty#2\else
+{\rail@namefont \rail@expand{#2}}\\*[\railnamesep]
+\fi
+\unitlength=\railunit
+\rail@tmpa=#1
+\multiply\rail@tmpa by \rail@boxsp
+\begin{picture}(0,\number\rail@tmpa)(0,-\number\rail@tmpa)
+\rail@ex=0
+\rail@rx=0
+\rail@x=\rail@extra
+\rail@sx=\rail@x
+\rail@sety{0}
+}
+
+\def\rail@end{
+\advance\rail@x by \rail@extra
+\rail@eline
+\end{picture}
+\end{minipage}
+}
+
+\def\rail@vend{
+\advance\rail@x by \rail@extra
+\rail@vreline
+\end{picture}
+\end{minipage}
+}
+
+\def\rail@expand#1{\@ifundefined{rail@t@#1}{#1}{\csname rail@t@#1\endcsname}}
+
+% \rail@token{TEXT}[ANNOT] : format token TEXT with annotation
+% \rail@ltoken{TEXT}[ANNOT] : format token TEXT with annotation, arrow left
+% \rail@rtoken{TEXT}[ANNOT] : format token TEXT with annotation, arrow right
+%
+% \rail@ctoken{TEXT}[ANNOT] : format token TEXT centered with annotation
+% \rail@lctoken{TEXT}[ANNOT] : format token TEXT centered with annotation, arrow left
+% \rail@rctoken{TEXT}[ANNOT] : format token TEXT centered with annotation, arrow right
+%
+% \rail@nont{TEXT}[ANNOT] : format nonterminal TEXT with annotation
+% \rail@lnont{TEXT}[ANNOT] : format nonterminal TEXT with annotation, arrow left
+% \rail@rnont{TEXT}[ANNOT] : format nonterminal TEXT with annotation. arrow right
+%
+% \rail@cnont{TEXT}[ANNOT] : format nonterminal TEXT centered with annotation
+% \rail@lcnont{TEXT}[ANNOT] : format nonterminal TEXT centered with annotation,
+% arrow left
+% \rail@rcnont{TEXT}[ANNOT] : format nonterminal TEXT centered with annotation,
+% arrow right
+%
+% \rail@term{TEXT}[ANNOT] : format terminal TEXT with annotation
+% \rail@lterm{TEXT}[ANNOT] : format terminal TEXT with annotation, arrow left
+% \rail@rterm{TEXT}[ANNOT] : format terminal TEXT with annotation, arrow right
+%
+% \rail@cterm{TEXT}[ANNOT] : format terminal TEXT centered with annotation
+% \rail@lcterm{TEXT}[ANNOT] : format terminal TEXT centered with annotation, arrow left
+% \rail@rcterm{TEXT}[ANNOT] : format terminal TEXT centered with annotation,
+% arrow right
+%
+% \rail@annote[TEXT] : format TEXT as annotation
+
+\def\rail@token#1[#2]{
+\rail@setbox{%
+{\rail@termfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@oval
+}
+
+\def\rail@ltoken#1[#2]{
+\rail@setbox{%
+{\rail@termfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vloval
+}
+
+\def\rail@rtoken#1[#2]{
+\rail@setbox{%
+{\rail@termfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vroval
+}
+
+\def\rail@ctoken#1[#2]{
+\rail@setbox{%
+{\rail@termfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@coval
+}
+
+\def\rail@lctoken#1[#2]{
+\rail@setbox{%
+{\rail@termfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vlcoval
+}
+
+\def\rail@rctoken#1[#2]{
+\rail@setbox{%
+{\rail@termfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vrcoval
+}
+
+\def\rail@nont#1[#2]{
+\rail@setbox{%
+{\rail@nontfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@frame
+}
+
+\def\rail@lnont#1[#2]{
+\rail@setbox{%
+{\rail@nontfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vlframe
+}
+
+\def\rail@rnont#1[#2]{
+\rail@setbox{%
+{\rail@nontfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vrframe
+}
+
+\def\rail@cnont#1[#2]{
+\rail@setbox{%
+{\rail@nontfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@cframe
+}
+
+\def\rail@lcnont#1[#2]{
+\rail@setbox{%
+{\rail@nontfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vlcframe
+}
+
+\def\rail@rcnont#1[#2]{
+\rail@setbox{%
+{\rail@nontfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vrcframe
+}
+
+\def\rail@term#1[#2]{
+\rail@setbox{%
+{\rail@termfont #1}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@oval
+}
+
+\def\rail@lterm#1[#2]{
+\rail@setbox{%
+{\rail@termfont #1}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vloval
+}
+
+\def\rail@rterm#1[#2]{
+\rail@setbox{%
+{\rail@termfont #1}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vroval
+}
+
+\def\rail@cterm#1[#2]{
+\rail@setbox{%
+{\rail@termfont #1}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@coval
+}
+
+\def\rail@lcterm#1[#2]{
+\rail@setbox{%
+{\rail@termfont #1}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vlcoval
+}
+
+\def\rail@rcterm#1[#2]{
+\rail@setbox{%
+{\rail@termfont #1}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vrcoval
+}
+
+\def\rail@annote[#1]{
+\rail@setbox{\rail@annofont #1}
+\rail@text
+}
+
+% \rail@box : temporary box for \rail@oval and \rail@frame
+%
+% \rail@setbox{TEXT} : set \rail@box to TEXT, set \rail@tmpa to width
+%
+% \rail@oval : format \rail@box of width \rail@tmpa inside an oval
+% \rail@vloval : format \rail@box of width \rail@tmpa inside an oval, vector left
+% \rail@vroval : format \rail@box of width \rail@tmpa inside an oval, vector right
+%
+% \rail@coval : same as \rail@oval, but centered between \rail@x and
+% \rail@mx
+% \rail@vlcoval : same as \rail@oval, but centered between \rail@x and
+% \rail@mx, vector left
+% \rail@vrcoval : same as \rail@oval, but centered between \rail@x and
+% \rail@mx, vector right
+%
+% \rail@frame : format \rail@box of width \rail@tmpa inside a frame
+% \rail@vlframe : format \rail@box of width \rail@tmpa inside a frame, vector left
+% \rail@vrframe : format \rail@box of width \rail@tmpa inside a frame, vector right
+%
+% \rail@cframe : same as \rail@frame, but centered between \rail@x and
+% \rail@mx
+% \rail@vlcframe : same as \rail@frame, but centered between \rail@x and
+% \rail@mx, vector left
+% \rail@vrcframe : same as \rail@frame, but centered between \rail@x and
+% \rail@mx, vector right
+%
+% \rail@text : format \rail@box of width \rail@tmpa above the line
+
+\newbox\rail@box
+
+\def\rail@setbox#1{
+\setbox\rail@box\hbox{\strut#1}
+\rail@tmpa=\wd\rail@box
+\divide\rail@tmpa by \railunit
+}
+
+\def\rail@oval{
+\advance\rail@x by \rail@boxlf
+\rail@eline
+\advance\rail@tmpa by \rail@ovalsp
+\ifnum\rail@tmpa<\rail@boxht\rail@tmpa=\rail@boxht\fi
+\rail@tmpb=\rail@tmpa
+\divide\rail@tmpb by 2
+\advance\rail@y by -\rail@boxhht
+\rail@put{\makebox(\number\rail@tmpa,\number\rail@boxht){\box\rail@box}}
+\advance\rail@y by \rail@boxhht
+\advance\rail@x by \rail@tmpb
+\rail@put{\oval(\number\rail@tmpa,\number\rail@boxht)}
+\advance\rail@x by \rail@tmpb
+\rail@ex=\rail@x
+\advance\rail@x by \rail@boxrt
+}
+
+\def\rail@vloval{
+\advance\rail@x by \rail@boxlf
+\rail@eline
+\advance\rail@tmpa by \rail@ovalsp
+\ifnum\rail@tmpa<\rail@boxht\rail@tmpa=\rail@boxht\fi
+\rail@tmpb=\rail@tmpa
+\divide\rail@tmpb by 2
+\advance\rail@y by -\rail@boxhht
+\rail@put{\makebox(\number\rail@tmpa,\number\rail@boxht){\box\rail@box}}
+\advance\rail@y by \rail@boxhht
+\advance\rail@x by \rail@tmpb
+\rail@put{\oval(\number\rail@tmpa,\number\rail@boxht)}
+\advance\rail@x by \rail@tmpb
+\rail@ex=\rail@x
+\advance\rail@x by \rail@boxrt
+\rail@vleline
+}
+
+\def\rail@vroval{
+\advance\rail@x by \rail@boxlf
+\rail@vreline
+\advance\rail@tmpa by \rail@ovalsp
+\ifnum\rail@tmpa<\rail@boxht\rail@tmpa=\rail@boxht\fi
+\rail@tmpb=\rail@tmpa
+\divide\rail@tmpb by 2
+\advance\rail@y by -\rail@boxhht
+\rail@put{\makebox(\number\rail@tmpa,\number\rail@boxht){\box\rail@box}}
+\advance\rail@y by \rail@boxhht
+\advance\rail@x by \rail@tmpb
+\rail@put{\oval(\number\rail@tmpa,\number\rail@boxht)}
+\advance\rail@x by \rail@tmpb
+\rail@ex=\rail@x
+\advance\rail@x by \rail@boxrt
+}
+
+\def\rail@coval{
+\rail@tmpb=\rail@tmpa
+\advance\rail@tmpb by \rail@ovalsp
+\ifnum\rail@tmpb<\rail@boxht\rail@tmpb=\rail@boxht\fi
+\advance\rail@tmpb by \rail@boxlf
+\advance\rail@tmpb by \rail@boxrt
+\rail@tmpc=\rail@mx
+\advance\rail@tmpc by -\rail@x
+\advance\rail@tmpc by -\rail@tmpb
+\divide\rail@tmpc by 2
+\ifnum\rail@tmpc>0
+\advance\rail@x by \rail@tmpc
+\fi
+\rail@oval
+}
+
+\def\rail@vlcoval{
+\rail@tmpb=\rail@tmpa
+\advance\rail@tmpb by \rail@ovalsp
+\ifnum\rail@tmpb<\rail@boxht\rail@tmpb=\rail@boxht\fi
+\advance\rail@tmpb by \rail@boxlf
+\advance\rail@tmpb by \rail@boxrt
+\rail@tmpc=\rail@mx
+\advance\rail@tmpc by -\rail@x
+\advance\rail@tmpc by -\rail@tmpb
+\divide\rail@tmpc by 2
+\ifnum\rail@tmpc>0
+\advance\rail@x by \rail@tmpc
+\fi
+\rail@vloval
+}
+
+\def\rail@vrcoval{
+\rail@tmpb=\rail@tmpa
+\advance\rail@tmpb by \rail@ovalsp
+\ifnum\rail@tmpb<\rail@boxht\rail@tmpb=\rail@boxht\fi
+\advance\rail@tmpb by \rail@boxlf
+\advance\rail@tmpb by \rail@boxrt
+\rail@tmpc=\rail@mx
+\advance\rail@tmpc by -\rail@x
+\advance\rail@tmpc by -\rail@tmpb
+\divide\rail@tmpc by 2
+\ifnum\rail@tmpc>0
+\advance\rail@x by \rail@tmpc
+\fi
+\rail@vroval
+}
+
+\def\rail@frame{
+\advance\rail@x by \rail@boxlf
+\rail@eline
+\advance\rail@tmpa by \rail@framesp
+\ifnum\rail@tmpa<\rail@boxht\rail@tmpa=\rail@boxht\fi
+\advance\rail@y by -\rail@boxhht
+\rail@put{\framebox(\number\rail@tmpa,\number\rail@boxht){\box\rail@box}}
+\advance\rail@y by \rail@boxhht
+\advance\rail@x by \rail@tmpa
+\rail@ex=\rail@x
+\advance\rail@x by \rail@boxrt
+}
+
+\def\rail@vlframe{
+\advance\rail@x by \rail@boxlf
+\rail@eline
+\advance\rail@tmpa by \rail@framesp
+\ifnum\rail@tmpa<\rail@boxht\rail@tmpa=\rail@boxht\fi
+\advance\rail@y by -\rail@boxhht
+\rail@put{\framebox(\number\rail@tmpa,\number\rail@boxht){\box\rail@box}}
+\advance\rail@y by \rail@boxhht
+\advance\rail@x by \rail@tmpa
+\rail@ex=\rail@x
+\advance\rail@x by \rail@boxrt
+\rail@vleline
+}
+
+\def\rail@vrframe{
+\advance\rail@x by \rail@boxlf
+\rail@vreline
+\advance\rail@tmpa by \rail@framesp
+\ifnum\rail@tmpa<\rail@boxht\rail@tmpa=\rail@boxht\fi
+\advance\rail@y by -\rail@boxhht
+\rail@put{\framebox(\number\rail@tmpa,\number\rail@boxht){\box\rail@box}}
+\advance\rail@y by \rail@boxhht
+\advance\rail@x by \rail@tmpa
+\rail@ex=\rail@x
+\advance\rail@x by \rail@boxrt
+}
+
+\def\rail@cframe{
+\rail@tmpb=\rail@tmpa
+\advance\rail@tmpb by \rail@framesp
+\ifnum\rail@tmpb<\rail@boxht\rail@tmpb=\rail@boxht\fi
+\advance\rail@tmpb by \rail@boxlf
+\advance\rail@tmpb by \rail@boxrt
+\rail@tmpc=\rail@mx
+\advance\rail@tmpc by -\rail@x
+\advance\rail@tmpc by -\rail@tmpb
+\divide\rail@tmpc by 2
+\ifnum\rail@tmpc>0
+\advance\rail@x by \rail@tmpc
+\fi
+\rail@frame
+}
+
+\def\rail@vlcframe{
+\rail@tmpb=\rail@tmpa
+\advance\rail@tmpb by \rail@framesp
+\ifnum\rail@tmpb<\rail@boxht\rail@tmpb=\rail@boxht\fi
+\advance\rail@tmpb by \rail@boxlf
+\advance\rail@tmpb by \rail@boxrt
+\rail@tmpc=\rail@mx
+\advance\rail@tmpc by -\rail@x
+\advance\rail@tmpc by -\rail@tmpb
+\divide\rail@tmpc by 2
+\ifnum\rail@tmpc>0
+\advance\rail@x by \rail@tmpc
+\fi
+\rail@vlframe
+}
+
+\def\rail@vrcframe{
+\rail@tmpb=\rail@tmpa
+\advance\rail@tmpb by \rail@framesp
+\ifnum\rail@tmpb<\rail@boxht\rail@tmpb=\rail@boxht\fi
+\advance\rail@tmpb by \rail@boxlf
+\advance\rail@tmpb by \rail@boxrt
+\rail@tmpc=\rail@mx
+\advance\rail@tmpc by -\rail@x
+\advance\rail@tmpc by -\rail@tmpb
+\divide\rail@tmpc by 2
+\ifnum\rail@tmpc>0
+\advance\rail@x by \rail@tmpc
+\fi
+\rail@vrframe
+}
+
+\def\rail@text{
+\advance\rail@x by \rail@textlf
+\advance\rail@y by \rail@textup
+\rail@put{\box\rail@box}
+\advance\rail@y by -\rail@textup
+\advance\rail@x by \rail@tmpa
+\advance\rail@x by \rail@textrt
+}
+
+% alternatives
+%
+% \rail@jx \rail@jy : current join point
+%
+% \rail@gx \rail@gy \rail@gex \rail@grx : global versions of \rail@x etc,
+% to pass values over group closings
+%
+% \rail@mx : maximum x so far
+%
+% \rail@sy : starting \rail@y for alternatives
+%
+% \rail@jput : put at (\rail@jx,\rail@jy)
+%
+% \rail@joval[PART] : put \oval[PART] with adjust
+
+\newcount\rail@jx
+\newcount\rail@jy
+
+\newcount\rail@gx
+\newcount\rail@gy
+\newcount\rail@gex
+\newcount\rail@grx
+
+\newcount\rail@sy
+\newcount\rail@mx
+
+\def\rail@jput{
+\put(\number\rail@jx,\number\rail@jy)
+}
+
+\def\rail@joval[#1]{
+\advance\rail@jx by \rail@joinadj
+\rail@jput{\oval(\number\rail@joinsz,\number\rail@joinsz)[#1]}
+\advance\rail@jx by -\rail@joinadj
+}
+
+% \rail@barsplit : incoming split for '|'
+%
+% \rail@plussplit : incoming split for '+'
+%
+
+\def\rail@barsplit{
+\advance\rail@jy by -\rail@joinhsz
+\rail@joval[tr]
+\advance\rail@jx by \rail@joinhsz
+}
+
+\def\rail@plussplit{
+\advance\rail@jy by -\rail@joinhsz
+\advance\rail@jx by \rail@joinsz
+\rail@joval[tl]
+\advance\rail@jx by -\rail@joinhsz
+}
+
+% \rail@alt{SPLIT} : start alternatives with incoming split SPLIT
+%
+
+\def\rail@alt#1{
+\rail@sy=\rail@y
+\rail@jx=\rail@x
+\rail@jy=\rail@y
+\advance\rail@x by \rail@joinsz
+\rail@mx=0
+\let\rail@list=\@empty
+\let\rail@comma=\@empty
+\let\rail@split=#1
+\begingroup
+\rail@sx=\rail@x
+\rail@rx=0
+}
+
+% \rail@nextalt{FIX}{Y} : start next alternative at vertical position Y
+% and fix-up FIX
+%
+
+\def\rail@nextalt#1#2{
+\global\rail@gx=\rail@x
+\global\rail@gy=\rail@y
+\global\rail@gex=\rail@ex
+\global\rail@grx=\rail@rx
+\endgroup
+#1
+\ifnum\rail@gx>\rail@mx\rail@mx=\rail@gx\fi
+\ifnum\rail@grx>\rail@mx\rail@mx=\rail@grx\fi
+\edef\rail@list{\rail@list\rail@comma\number\rail@gex:\number\rail@gy}
+\def\rail@comma{,}
+\rail@split
+\let\rail@split=\@empty
+\rail@sety{#2}
+\rail@tmpa=\rail@jy
+\advance\rail@tmpa by -\rail@y
+\advance\rail@tmpa by -\rail@joinhsz
+\rail@jput{\line(0,-1){\number\rail@tmpa}}
+\rail@jy=\rail@y
+\advance\rail@jy by \rail@joinhsz
+\advance\rail@jx by \rail@joinhsz
+\rail@joval[bl]
+\advance\rail@jx by -\rail@joinhsz
+\rail@ex=\rail@x
+\begingroup
+\rail@sx=\rail@x
+\rail@rx=0
+}
+
+% \rail@barjoin : outgoing join for first '|' alternative
+%
+% \rail@plusjoin : outgoing join for first '+' alternative
+%
+% \rail@altjoin : join for subsequent alternative
+%
+
+\def\rail@barjoin{
+\ifnum\rail@y<\rail@sy
+\global\rail@gex=\rail@jx
+\else
+\global\rail@gex=\rail@ex
+\fi
+\advance\rail@jy by -\rail@joinhsz
+\rail@joval[tl]
+\advance\rail@jx by -\rail@joinhsz
+\ifnum\rail@y<\rail@sy
+\rail@altjoin
+\fi
+}
+
+\def\rail@plusjoin{
+\global\rail@gex=\rail@ex
+\advance\rail@jy by -\rail@joinhsz
+\advance\rail@jx by -\rail@joinsz
+\rail@joval[tr]
+\advance\rail@jx by \rail@joinhsz
+}
+
+\def\rail@altjoin{
+\rail@eline
+\rail@tmpa=\rail@jy
+\advance\rail@tmpa by -\rail@y
+\advance\rail@tmpa by -\rail@joinhsz
+\rail@jput{\line(0,-1){\number\rail@tmpa}}
+\rail@jy=\rail@y
+\advance\rail@jy by \rail@joinhsz
+\advance\rail@jx by -\rail@joinhsz
+\rail@joval[br]
+\advance\rail@jx by \rail@joinhsz
+}
+
+% \rail@eltsplit EX:Y; : split EX:Y into \rail@ex \rail@y
+%
+% \rail@endalt{JOIN} : end alternatives with outgoing join JOIN
+
+\def\rail@eltsplit#1:#2;{\rail@ex=#1\rail@y=#2}
+
+\def\rail@endalt#1{
+\global\rail@gx=\rail@x
+\global\rail@gy=\rail@y
+\global\rail@gex=\rail@ex
+\global\rail@grx=\rail@rx
+\endgroup
+\ifnum\rail@gx>\rail@mx\rail@mx=\rail@gx\fi
+\ifnum\rail@grx>\rail@mx\rail@mx=\rail@grx\fi
+\edef\rail@list{\rail@list\rail@comma\number\rail@gex:\number\rail@gy}
+\rail@x=\rail@mx
+\rail@jx=\rail@x
+\rail@jy=\rail@sy
+\advance\rail@jx by \rail@joinsz
+\let\rail@join=#1
+\@for\rail@elt:=\rail@list\do{
+\expandafter\rail@eltsplit\rail@elt;
+\rail@join
+\let\rail@join=\rail@altjoin
+}
+\rail@x=\rail@mx
+\rail@y=\rail@sy
+\rail@ex=\rail@gex
+\advance\rail@x by \rail@joinsz
+}
+
+% \rail@bar : start '|' alternatives
+%
+% \rail@nextbar : next '|' alternative
+%
+% \rail@endbar : end '|' alternatives
+%
+
+\def\rail@bar{
+\rail@alt\rail@barsplit
+}
+
+\def\rail@nextbar{
+\rail@nextalt\relax
+}
+
+\def\rail@endbar{
+\rail@endalt\rail@barjoin
+}
+
+% \rail@plus : start '+' alternatives
+%
+% \rail@nextplus: next '+' alternative
+%
+% \rail@endplus : end '+' alternatives
+%
+
+\def\rail@plus{
+\rail@alt\rail@plussplit
+}
+
+\def\rail@nextplus{
+\rail@nextalt\rail@fixplus
+}
+
+\def\rail@fixplus{
+\ifnum\rail@gy<\rail@sy
+\begingroup
+\rail@x=\rail@gx
+\rail@y=\rail@gy
+\rail@ex=\rail@gex
+\rail@rx=\rail@grx
+\ifnum\rail@x<\rail@rx
+\rail@x=\rail@rx
+\fi
+\rail@eline
+\rail@jx=\rail@x
+\rail@jy=\rail@y
+\advance\rail@jy by \rail@joinhsz
+\rail@joval[br]
+\advance\rail@jx by \rail@joinhsz
+\rail@tmpa=\rail@sy
+\advance\rail@tmpa by -\rail@joinhsz
+\advance\rail@tmpa by -\rail@jy
+\rail@jput{\line(0,1){\number\rail@tmpa}}
+\rail@jy=\rail@sy
+\advance\rail@jy by -\rail@joinhsz
+\advance\rail@jx by \rail@joinhsz
+\rail@joval[tl]
+\advance\rail@jy by \rail@joinhsz
+\global\rail@gx=\rail@jx
+\global\rail@gy=\rail@jy
+\global\rail@gex=\rail@gx
+\global\rail@grx=\rail@rx
+\endgroup
+\fi
+}
+
+\def\rail@endplus{
+\rail@endalt\rail@plusjoin
+}
+
+% \rail@cr{Y} : carriage return to vertical position Y
+
+\def\rail@cr#1{
+\rail@tmpa=\rail@sx
+\advance\rail@tmpa by \rail@joinsz
+\ifnum\rail@x<\rail@tmpa\rail@x=\rail@tmpa\fi
+\rail@eline
+\rail@jx=\rail@x
+\rail@jy=\rail@y
+\advance\rail@x by \rail@joinsz
+\ifnum\rail@x>\rail@rx\rail@rx=\rail@x\fi
+\advance\rail@jy by -\rail@joinhsz
+\rail@joval[tr]
+\advance\rail@jx by \rail@joinhsz
+\rail@sety{#1}
+\rail@tmpa=\rail@jy
+\advance\rail@tmpa by -\rail@y
+\advance\rail@tmpa by -\rail@boxsp
+\advance\rail@tmpa by -\rail@joinhsz
+\rail@jput{\line(0,-1){\number\rail@tmpa}}
+\rail@jy=\rail@y
+\advance\rail@jy by \rail@boxsp
+\advance\rail@jy by \rail@joinhsz
+\advance\rail@jx by -\rail@joinhsz
+\rail@joval[br]
+\advance\rail@jy by -\rail@joinhsz
+\rail@tmpa=\rail@jx
+\advance\rail@tmpa by -\rail@sx
+\advance\rail@tmpa by -\rail@joinhsz
+\rail@jput{\line(-1,0){\number\rail@tmpa}}
+\rail@jx=\rail@sx
+\advance\rail@jx by \rail@joinhsz
+\advance\rail@jy by -\rail@joinhsz
+\rail@joval[tl]
+\advance\rail@jx by -\rail@joinhsz
+\rail@tmpa=\rail@boxsp
+\advance\rail@tmpa by -\rail@joinsz
+\rail@jput{\line(0,-1){\number\rail@tmpa}}
+\advance\rail@jy by -\rail@tmpa
+\advance\rail@jx by \rail@joinhsz
+\rail@joval[bl]
+\rail@x=\rail@jx
+\rail@ex=\rail@x
+}
+
+% default setup for Isabelle
+\newenvironment{railoutput}%
+{\begin{list}{}{\rail@param}\def\rail@expand{\relax}\makeatletter}{\makeatother\end{list}}
+
+\def\rail@termfont{\isabellestyle{tt}}
+\def\rail@nontfont{\isabellestyle{it}}
+\def\rail@namefont{\isabellestyle{it}}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/tex/document/root.aux Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,335 @@
+\relax
+\providecommand\hyper@newdestlabel[2]{}
+\providecommand\HyperFirstAtBeginDocument{\AtBeginDocument}
+\HyperFirstAtBeginDocument{\ifx\hyper@anchor\@undefined
+\global\let\oldcontentsline\contentsline
+\gdef\contentsline#1#2#3#4{\oldcontentsline{#1}{#2}{#3}}
+\global\let\oldnewlabel\newlabel
+\gdef\newlabel#1#2{\newlabelxx{#1}#2}
+\gdef\newlabelxx#1#2#3#4#5#6{\oldnewlabel{#1}{{#2}{#3}}}
+\AtEndDocument{\ifx\hyper@anchor\@undefined
+\let\contentsline\oldcontentsline
+\let\newlabel\oldnewlabel
+\fi}
+\fi}
+\global\let\hyper@last\relax
+\gdef\HyperFirstAtBeginDocument#1{#1}
+\providecommand\HyField@AuxAddToFields[1]{}
+\providecommand\HyField@AuxAddToCoFields[2]{}
+\@writefile{toc}{\beamer@endinputifotherversion {3.32pt}}
+\@writefile{nav}{\beamer@endinputifotherversion {3.32pt}}
+\@writefile{nav}{\headcommand {\slideentry {0}{0}{1}{1/1}{}{0}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {1}{1}}}
+\@writefile{nav}{\headcommand {\slideentry {0}{0}{2}{2/2}{}{0}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {2}{2}}}
+\@writefile{nav}{\headcommand {\partentry {Sorting}{6}}}
+\@writefile{nav}{\headcommand {\beamer@partpages {1}{2}}}
+\@writefile{nav}{\headcommand {\beamer@sectionpages {1}{2}}}
+\@writefile{nav}{\headcommand {\beamer@subsectionpages {1}{2}}}
+\@writefile{nav}{\headcommand {\slideentry {0}{0}{3}{3/3}{}{6}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {3}{3}}}
+\@writefile{nav}{\headcommand {\slideentry {0}{0}{4}{4/4}{}{6}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {4}{4}}}
+\@writefile{toc}{\beamer@sectionintoc {1}{Correctness}{5}{6}{1}}
+\@writefile{nav}{\headcommand {\sectionentry {1}{Correctness}{5}{Correctness}{6}}}
+\@writefile{nav}{\headcommand {\beamer@sectionpages {3}{4}}}
+\@writefile{nav}{\headcommand {\beamer@subsectionpages {3}{4}}}
+\@writefile{nav}{\headcommand {\slideentry {1}{0}{5}{5/5}{}{6}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {5}{5}}}
+\@writefile{nav}{\headcommand {\slideentry {1}{0}{6}{6/6}{}{6}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {6}{6}}}
+\@writefile{nav}{\headcommand {\slideentry {1}{0}{7}{7/7}{}{6}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {7}{7}}}
+\@writefile{nav}{\headcommand {\slideentry {1}{0}{8}{8/8}{}{6}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {8}{8}}}
+\@writefile{toc}{\beamer@sectionintoc {2}{Insertion Sort}{9}{6}{2}}
+\@writefile{nav}{\headcommand {\sectionentry {2}{Insertion Sort}{9}{Insertion Sort}{6}}}
+\@writefile{nav}{\headcommand {\beamer@sectionpages {5}{8}}}
+\@writefile{nav}{\headcommand {\beamer@subsectionpages {5}{8}}}
+\@writefile{nav}{\headcommand {\slideentry {2}{0}{9}{9/9}{}{6}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {9}{9}}}
+\@writefile{nav}{\headcommand {\slideentry {2}{0}{10}{10/10}{}{6}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {10}{10}}}
+\@writefile{toc}{\beamer@sectionintoc {3}{Time}{11}{6}{3}}
+\@writefile{nav}{\headcommand {\sectionentry {3}{Time}{11}{Time}{6}}}
+\@writefile{nav}{\headcommand {\beamer@sectionpages {9}{10}}}
+\@writefile{nav}{\headcommand {\beamer@subsectionpages {9}{10}}}
+\@writefile{nav}{\headcommand {\slideentry {3}{0}{11}{11/11}{}{6}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {11}{11}}}
+\@writefile{nav}{\headcommand {\slideentry {3}{0}{12}{12/12}{}{6}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {12}{12}}}
+\@writefile{nav}{\headcommand {\slideentry {3}{0}{13}{13/13}{}{6}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {13}{13}}}
+\@writefile{nav}{\headcommand {\slideentry {3}{0}{14}{14/14}{}{6}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {14}{14}}}
+\@writefile{nav}{\headcommand {\slideentry {3}{0}{15}{15/15}{}{6}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {15}{15}}}
+\@writefile{nav}{\headcommand {\slideentry {3}{0}{16}{16/16}{}{6}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {16}{16}}}
+\@writefile{nav}{\headcommand {\slideentry {3}{0}{17}{17/17}{}{6}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {17}{17}}}
+\@writefile{nav}{\headcommand {\slideentry {3}{0}{18}{18/18}{}{6}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {18}{18}}}
+\@writefile{toc}{\beamer@sectionintoc {4}{Merge Sort}{19}{6}{4}}
+\@writefile{nav}{\headcommand {\sectionentry {4}{Merge Sort}{19}{Merge Sort}{6}}}
+\@writefile{nav}{\headcommand {\beamer@sectionpages {11}{18}}}
+\@writefile{nav}{\headcommand {\beamer@subsectionpages {11}{18}}}
+\@writefile{nav}{\headcommand {\slideentry {4}{0}{19}{19/19}{}{6}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {19}{19}}}
+\@writefile{nav}{\headcommand {\slideentry {4}{0}{20}{20/20}{}{6}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {20}{20}}}
+\@writefile{nav}{\headcommand {\slideentry {4}{0}{21}{21/21}{}{6}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {21}{21}}}
+\@writefile{nav}{\headcommand {\partentry {Binary Trees}{7}}}
+\@writefile{nav}{\headcommand {\beamer@partpages {3}{21}}}
+\@writefile{nav}{\headcommand {\beamer@sectionpages {19}{21}}}
+\@writefile{nav}{\headcommand {\beamer@subsectionpages {19}{21}}}
+\@writefile{nav}{\headcommand {\slideentry {4}{0}{22}{22/22}{}{7}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {22}{22}}}
+\@writefile{nav}{\headcommand {\slideentry {4}{0}{23}{23/23}{}{7}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {23}{23}}}
+\@writefile{toc}{\beamer@sectionintoc {5}{Binary Trees}{24}{7}{5}}
+\@writefile{nav}{\headcommand {\sectionentry {5}{Binary Trees}{24}{Binary Trees}{7}}}
+\@writefile{nav}{\headcommand {\beamer@sectionpages {22}{23}}}
+\@writefile{nav}{\headcommand {\beamer@subsectionpages {22}{23}}}
+\@writefile{nav}{\headcommand {\slideentry {5}{0}{24}{24/24}{}{7}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {24}{24}}}
+\@writefile{nav}{\headcommand {\slideentry {5}{0}{25}{25/25}{}{7}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {25}{25}}}
+\@writefile{nav}{\headcommand {\slideentry {5}{0}{26}{26/26}{}{7}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {26}{26}}}
+\@writefile{toc}{\beamer@sectionintoc {6}{Basic Functions}{27}{7}{6}}
+\@writefile{nav}{\headcommand {\sectionentry {6}{Basic Functions}{27}{Basic Functions}{7}}}
+\@writefile{nav}{\headcommand {\beamer@sectionpages {24}{26}}}
+\@writefile{nav}{\headcommand {\beamer@subsectionpages {24}{26}}}
+\@writefile{nav}{\headcommand {\slideentry {6}{0}{27}{27/27}{}{7}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {27}{27}}}
+\@writefile{nav}{\headcommand {\slideentry {6}{0}{28}{28/28}{}{7}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {28}{28}}}
+\@writefile{nav}{\headcommand {\slideentry {6}{0}{29}{29/29}{}{7}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {29}{29}}}
+\@writefile{nav}{\headcommand {\slideentry {6}{0}{30}{30/30}{}{7}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {30}{30}}}
+\@writefile{nav}{\headcommand {\slideentry {6}{0}{31}{31/31}{}{7}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {31}{31}}}
+\@writefile{nav}{\headcommand {\slideentry {6}{0}{32}{32/32}{}{7}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {32}{32}}}
+\@writefile{toc}{\beamer@sectionintoc {7}{Complete and Balanced Trees}{33}{7}{7}}
+\@writefile{nav}{\headcommand {\sectionentry {7}{Complete and Balanced Trees}{33}{Complete and Balanced Trees}{7}}}
+\@writefile{nav}{\headcommand {\beamer@sectionpages {27}{32}}}
+\@writefile{nav}{\headcommand {\beamer@subsectionpages {27}{32}}}
+\@writefile{nav}{\headcommand {\slideentry {7}{0}{33}{33/33}{}{7}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {33}{33}}}
+\@writefile{nav}{\headcommand {\slideentry {7}{0}{34}{34/34}{}{7}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {34}{34}}}
+\@writefile{nav}{\headcommand {\slideentry {7}{0}{35}{35/35}{}{7}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {35}{35}}}
+\@writefile{nav}{\headcommand {\slideentry {7}{0}{36}{36/36}{}{7}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {36}{36}}}
+\@writefile{nav}{\headcommand {\slideentry {7}{0}{37}{37/37}{}{7}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {37}{37}}}
+\@writefile{nav}{\headcommand {\slideentry {7}{0}{38}{38/38}{}{7}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {38}{38}}}
+\@writefile{nav}{\headcommand {\partentry {Search Trees}{8}}}
+\@writefile{nav}{\headcommand {\beamer@partpages {22}{38}}}
+\@writefile{nav}{\headcommand {\beamer@sectionpages {33}{38}}}
+\@writefile{nav}{\headcommand {\beamer@subsectionpages {33}{38}}}
+\@writefile{nav}{\headcommand {\slideentry {7}{0}{39}{39/39}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {39}{39}}}
+\@writefile{nav}{\headcommand {\slideentry {7}{0}{40}{40/40}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {40}{40}}}
+\@writefile{nav}{\headcommand {\slideentry {7}{0}{41}{41/41}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {41}{41}}}
+\@writefile{nav}{\headcommand {\slideentry {7}{0}{42}{42/42}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {42}{42}}}
+\@writefile{nav}{\headcommand {\slideentry {7}{0}{43}{43/43}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {43}{43}}}
+\@writefile{nav}{\headcommand {\slideentry {7}{0}{44}{44/44}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {44}{44}}}
+\@writefile{nav}{\headcommand {\slideentry {7}{0}{45}{45/45}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {45}{45}}}
+\@writefile{nav}{\headcommand {\slideentry {7}{0}{46}{46/46}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {46}{46}}}
+\@writefile{toc}{\beamer@sectionintoc {8}{Unbalanced BST}{47}{8}{8}}
+\@writefile{nav}{\headcommand {\sectionentry {8}{Unbalanced BST}{47}{Unbalanced BST}{8}}}
+\@writefile{nav}{\headcommand {\beamer@sectionpages {39}{46}}}
+\@writefile{nav}{\headcommand {\beamer@subsectionpages {39}{46}}}
+\@writefile{nav}{\headcommand {\slideentry {8}{0}{47}{47/47}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {47}{47}}}
+\@writefile{nav}{\headcommand {\slideentry {8}{0}{48}{48/48}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {48}{48}}}
+\@writefile{nav}{\headcommand {\slideentry {8}{0}{49}{49/49}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {49}{49}}}
+\@writefile{nav}{\headcommand {\slideentry {8}{0}{50}{50/50}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {50}{50}}}
+\@writefile{toc}{\beamer@subsectionintoc {8}{1}{Correctness}{51}{8}{8}}
+\@writefile{nav}{\headcommand {\beamer@subsectionentry {8}{8}{1}{51}{Correctness}}\headcommand {\beamer@subsectionpages {47}{50}}}
+\@writefile{nav}{\headcommand {\slideentry {8}{1}{1}{51/51}{Correctness}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {51}{51}}}
+\@writefile{nav}{\headcommand {\slideentry {8}{1}{2}{52/52}{Correctness}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {52}{52}}}
+\@writefile{nav}{\headcommand {\slideentry {8}{1}{3}{53/53}{Correctness}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {53}{53}}}
+\@writefile{toc}{\beamer@subsectionintoc {8}{2}{Correctness Proof Method Based on Sorted Lists}{54}{8}{8}}
+\@writefile{nav}{\headcommand {\beamer@subsectionentry {8}{8}{2}{54}{Correctness Proof Method Based on Sorted Lists}}\headcommand {\beamer@subsectionpages {51}{53}}}
+\@writefile{nav}{\headcommand {\slideentry {8}{2}{1}{54/54}{Correctness Proof Method Based on Sorted Lists}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {54}{54}}}
+\@writefile{nav}{\headcommand {\slideentry {8}{2}{2}{55/55}{Correctness Proof Method Based on Sorted Lists}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {55}{55}}}
+\@writefile{nav}{\headcommand {\slideentry {8}{2}{3}{56/56}{Correctness Proof Method Based on Sorted Lists}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {56}{56}}}
+\@writefile{nav}{\headcommand {\slideentry {8}{2}{4}{57/57}{Correctness Proof Method Based on Sorted Lists}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {57}{57}}}
+\@writefile{nav}{\headcommand {\slideentry {8}{2}{5}{58/58}{Correctness Proof Method Based on Sorted Lists}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {58}{58}}}
+\@writefile{nav}{\headcommand {\slideentry {8}{2}{6}{59/59}{Correctness Proof Method Based on Sorted Lists}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {59}{59}}}
+\@writefile{nav}{\headcommand {\slideentry {8}{2}{7}{60/60}{Correctness Proof Method Based on Sorted Lists}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {60}{60}}}
+\@writefile{toc}{\beamer@sectionintoc {9}{2-3 Trees}{61}{8}{9}}
+\@writefile{nav}{\headcommand {\sectionentry {9}{2-3 Trees}{61}{2-3 Trees}{8}}}
+\@writefile{nav}{\headcommand {\beamer@sectionpages {47}{60}}}
+\@writefile{nav}{\headcommand {\beamer@subsectionpages {54}{60}}}
+\@writefile{nav}{\headcommand {\slideentry {9}{0}{8}{61/61}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {61}{61}}}
+\@writefile{nav}{\headcommand {\slideentry {9}{0}{9}{62/62}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {62}{62}}}
+\@writefile{nav}{\headcommand {\slideentry {9}{0}{10}{63/63}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {63}{63}}}
+\@writefile{nav}{\headcommand {\slideentry {9}{0}{11}{64/64}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {64}{64}}}
+\@writefile{nav}{\headcommand {\slideentry {9}{0}{12}{65/65}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {65}{65}}}
+\@writefile{nav}{\headcommand {\slideentry {9}{0}{13}{66/66}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {66}{66}}}
+\@writefile{nav}{\headcommand {\slideentry {9}{0}{14}{67/67}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {67}{67}}}
+\@writefile{nav}{\headcommand {\slideentry {9}{0}{15}{68/68}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {68}{68}}}
+\@writefile{nav}{\headcommand {\slideentry {9}{0}{16}{69/69}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {69}{69}}}
+\@writefile{nav}{\headcommand {\slideentry {9}{0}{17}{70/70}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {70}{70}}}
+\@writefile{nav}{\headcommand {\slideentry {9}{0}{18}{71/71}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {71}{71}}}
+\@writefile{nav}{\headcommand {\slideentry {9}{0}{19}{72/72}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {72}{72}}}
+\@writefile{nav}{\headcommand {\slideentry {9}{0}{20}{73/73}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {73}{73}}}
+\@writefile{nav}{\headcommand {\slideentry {9}{0}{21}{74/74}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {74}{74}}}
+\@writefile{nav}{\headcommand {\slideentry {9}{0}{22}{75/75}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {75}{75}}}
+\@writefile{nav}{\headcommand {\slideentry {9}{0}{23}{76/76}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {76}{76}}}
+\@writefile{nav}{\headcommand {\slideentry {9}{0}{24}{77/77}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {77}{77}}}
+\@writefile{nav}{\headcommand {\slideentry {9}{0}{25}{78/78}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {78}{78}}}
+\@writefile{toc}{\beamer@sectionintoc {10}{Red-Black Trees}{79}{8}{10}}
+\@writefile{nav}{\headcommand {\sectionentry {10}{Red-Black Trees}{79}{Red-Black Trees}{8}}}
+\@writefile{nav}{\headcommand {\beamer@sectionpages {61}{78}}}
+\@writefile{nav}{\headcommand {\beamer@subsectionpages {61}{78}}}
+\@writefile{nav}{\headcommand {\slideentry {10}{0}{26}{79/79}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {79}{79}}}
+\@writefile{nav}{\headcommand {\slideentry {10}{0}{27}{80/80}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {80}{80}}}
+\@writefile{nav}{\headcommand {\slideentry {10}{0}{28}{81/81}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {81}{81}}}
+\@writefile{nav}{\headcommand {\slideentry {10}{0}{29}{82/82}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {82}{82}}}
+\@writefile{nav}{\headcommand {\slideentry {10}{0}{30}{83/83}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {83}{83}}}
+\@writefile{nav}{\headcommand {\slideentry {10}{0}{31}{84/84}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {84}{84}}}
+\@writefile{nav}{\headcommand {\slideentry {10}{0}{32}{85/85}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {85}{85}}}
+\@writefile{nav}{\headcommand {\slideentry {10}{0}{33}{86/86}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {86}{86}}}
+\@writefile{nav}{\headcommand {\slideentry {10}{0}{34}{87/87}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {87}{87}}}
+\@writefile{nav}{\headcommand {\slideentry {10}{0}{35}{88/88}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {88}{88}}}
+\@writefile{nav}{\headcommand {\slideentry {10}{0}{36}{89/89}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {89}{89}}}
+\@writefile{nav}{\headcommand {\slideentry {10}{0}{37}{90/90}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {90}{90}}}
+\@writefile{nav}{\headcommand {\slideentry {10}{0}{38}{91/91}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {91}{91}}}
+\@writefile{nav}{\headcommand {\slideentry {10}{0}{39}{92/92}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {92}{92}}}
+\@writefile{toc}{\beamer@sectionintoc {11}{Tries and Patricia Tries}{93}{8}{11}}
+\@writefile{nav}{\headcommand {\sectionentry {11}{Tries and Patricia Tries}{93}{Tries and Patricia Tries}{8}}}
+\@writefile{nav}{\headcommand {\beamer@sectionpages {79}{92}}}
+\@writefile{nav}{\headcommand {\beamer@subsectionpages {79}{92}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{0}{40}{93/93}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {93}{93}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{0}{41}{94/94}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {94}{94}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{0}{42}{95/95}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {95}{95}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{0}{43}{96/96}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {96}{96}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{0}{44}{97/97}{}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {97}{97}}}
+\@writefile{toc}{\beamer@subsectionintoc {11}{1}{Data Refinement Basics}{98}{8}{11}}
+\@writefile{nav}{\headcommand {\beamer@subsectionentry {8}{11}{1}{98}{Data Refinement Basics}}\headcommand {\beamer@subsectionpages {93}{97}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{1}{1}{98/98}{Data Refinement Basics}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {98}{98}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{1}{2}{99/99}{Data Refinement Basics}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {99}{99}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{1}{3}{100/100}{Data Refinement Basics}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {100}{100}}}
+\@writefile{toc}{\beamer@subsectionintoc {11}{2}{Trie}{101}{8}{11}}
+\@writefile{nav}{\headcommand {\beamer@subsectionentry {8}{11}{2}{101}{Trie}}\headcommand {\beamer@subsectionpages {98}{100}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{2}{1}{101/101}{Trie}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {101}{101}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{2}{2}{102/102}{Trie}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {102}{102}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{2}{3}{103/103}{Trie}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {103}{103}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{2}{4}{104/104}{Trie}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {104}{104}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{2}{5}{105/105}{Trie}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {105}{105}}}
+\@writefile{toc}{\beamer@subsectionintoc {11}{3}{Patricia Trie}{106}{8}{11}}
+\@writefile{nav}{\headcommand {\beamer@subsectionentry {8}{11}{3}{106}{Patricia Trie}}\headcommand {\beamer@subsectionpages {101}{105}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{3}{1}{106/106}{Patricia Trie}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {106}{106}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{3}{2}{107/107}{Patricia Trie}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {107}{107}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{3}{3}{108/108}{Patricia Trie}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {108}{108}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{3}{4}{109/109}{Patricia Trie}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {109}{109}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{3}{5}{110/110}{Patricia Trie}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {110}{110}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{3}{6}{111/111}{Patricia Trie}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {111}{111}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{3}{7}{112/112}{Patricia Trie}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {112}{112}}}
+\@writefile{toc}{\beamer@subsectionintoc {11}{4}{Patricia Trie for Words}{113}{8}{11}}
+\@writefile{nav}{\headcommand {\beamer@subsectionentry {8}{11}{4}{113}{Patricia Trie for Words}}\headcommand {\beamer@subsectionpages {106}{112}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{4}{1}{113/113}{Patricia Trie for Words}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {113}{113}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{4}{2}{114/114}{Patricia Trie for Words}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {114}{114}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{4}{3}{115/115}{Patricia Trie for Words}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {115}{115}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{4}{4}{116/116}{Patricia Trie for Words}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {116}{116}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{4}{5}{117/117}{Patricia Trie for Words}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {117}{117}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{4}{6}{118/118}{Patricia Trie for Words}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {118}{118}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{4}{7}{119/119}{Patricia Trie for Words}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {119}{119}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{4}{8}{120/120}{Patricia Trie for Words}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {120}{120}}}
+\@writefile{nav}{\headcommand {\slideentry {11}{4}{9}{121/121}{Patricia Trie for Words}{8}}}
+\@writefile{nav}{\headcommand {\beamer@framepages {121}{121}}}
+\@writefile{nav}{\headcommand {\beamer@partpages {39}{121}}}
+\@writefile{nav}{\headcommand {\beamer@subsectionpages {113}{121}}}
+\@writefile{nav}{\headcommand {\beamer@sectionpages {93}{121}}}
+\@writefile{nav}{\headcommand {\beamer@documentpages {121}}}
+\@writefile{nav}{\headcommand {\def \inserttotalframenumber {121}}}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/tex/document/root.log Fri Jun 23 18:08:31 2017 +0200
@@ -0,0 +1,1320 @@
+This is pdfTeX, Version 3.1415926-2.5-1.40.14 (TeX Live 2013) (format=pdflatex 2013.12.17) 23 JUN 2017 18:07
+entering extended mode
+ restricted \write18 enabled.
+ %&-line parsing enabled.
+**\nonstopmode\input{root.tex}
+(./root.tex (/usr/local/texlive/2013/texmf-dist/tex/latex/beamer/beamer.cls
+(/usr/local/texlive/2013/texmf-dist/tex/latex/beamer/beamerbasercs.sty
+Package: beamerbasercs 2013/11/24 (rcs-revision 8ee54e934cd6)
+)
+Document Class: beamer 2013/04/04 3.32 A class for typesetting presentations (r
+cs-revision dca7db0ccda1)
+(/usr/local/texlive/2013/texmf-dist/tex/latex/beamer/beamerbasemodes.sty
+Package: beamerbasemodes 2013/09/03 (rcs-revision 768f2d98ca64)
+\beamer@tempbox=\box26
+\beamer@tempcount=\count79
+\c@beamerpauses=\count80
+
+(/usr/local/texlive/2013/texmf-dist/tex/latex/beamer/beamerbasedecode.sty
+Package: beamerbasedecode 2010/05/01 (rcs-revision efa082c6111d)
+\beamer@slideinframe=\count81
+\beamer@minimum=\count82
+)
+\beamer@commentbox=\box27
+\beamer@modecount=\count83
+)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/oberdiek/ifpdf.sty
+Package: ifpdf 2011/01/30 v2.3 Provides the ifpdf switch (HO)
+Package ifpdf Info: pdfTeX in PDF mode is detected.
+)
+\headheight=\dimen102
+\headdp=\dimen103
+\footheight=\dimen104
+\sidebarheight=\dimen105
+\beamer@tempdim=\dimen106
+\beamer@finalheight=\dimen107
+\beamer@animht=\dimen108
+\beamer@animdp=\dimen109
+\beamer@animwd=\dimen110
+\beamer@leftmargin=\dimen111
+\beamer@rightmargin=\dimen112
+\beamer@leftsidebar=\dimen113
+\beamer@rightsidebar=\dimen114
+\beamer@boxsize=\dimen115
+\beamer@vboxoffset=\dimen116
+\beamer@descdefault=\dimen117
+\beamer@descriptionwidth=\dimen118
+\beamer@lastskip=\skip41
+\beamer@areabox=\box28
+\beamer@animcurrent=\box29
+\beamer@animshowbox=\box30
+\beamer@sectionbox=\box31
+\beamer@logobox=\box32
+\beamer@linebox=\box33
+\beamer@sectioncount=\count84
+\beamer@subsubsectionmax=\count85
+\beamer@subsectionmax=\count86
+\beamer@sectionmax=\count87
+\beamer@totalheads=\count88
+\beamer@headcounter=\count89
+\beamer@partstartpage=\count90
+\beamer@sectionstartpage=\count91
+\beamer@subsectionstartpage=\count92
+\beamer@animationtempa=\count93
+\beamer@animationtempb=\count94
+\beamer@xpos=\count95
+\beamer@ypos=\count96
+\beamer@showpartnumber=\count97
+\beamer@currentsubsection=\count98
+\beamer@coveringdepth=\count99
+\beamer@sectionadjust=\count100
+\beamer@tocsectionnumber=\count101
+
+(/usr/local/texlive/2013/texmf-dist/tex/latex/beamer/beamerbaseoptions.sty
+Package: beamerbaseoptions 2013/03/10 (rcs-revision 47431932db0d)
+
+(/usr/local/texlive/2013/texmf-dist/tex/latex/graphics/keyval.sty
+Package: keyval 1999/03/16 v1.13 key=value parser (DPC)
+\KV@toks@=\toks14
+))
+\beamer@paperwidth=\skip42
+\beamer@paperheight=\skip43
+
+(/usr/local/texlive/2013/texmf-dist/tex/latex/geometry/geometry.sty
+Package: geometry 2010/09/12 v5.6 Page Geometry
+
+(/usr/local/texlive/2013/texmf-dist/tex/generic/oberdiek/ifvtex.sty
+Package: ifvtex 2010/03/01 v1.5 Detect VTeX and its facilities (HO)
+Package ifvtex Info: VTeX not detected.
+)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/ifxetex/ifxetex.sty
+Package: ifxetex 2010/09/12 v0.6 Provides ifxetex conditional
+)
+\Gm@cnth=\count102
+\Gm@cntv=\count103
+\c@Gm@tempcnt=\count104
+\Gm@bindingoffset=\dimen119
+\Gm@wd@mp=\dimen120
+\Gm@odd@mp=\dimen121
+\Gm@even@mp=\dimen122
+\Gm@layoutwidth=\dimen123
+\Gm@layoutheight=\dimen124
+\Gm@layouthoffset=\dimen125
+\Gm@layoutvoffset=\dimen126
+\Gm@dimlist=\toks15
+)
+(/usr/local/texlive/2013/texmf-dist/tex/latex/extsizes/size14.clo
+File: size14.clo 1999/11/11 v1.4a NON-Standard LaTeX file (size option)
+)
+(/usr/local/texlive/2013/texmf-dist/tex/latex/pgf/basiclayer/pgfcore.sty
+(/usr/local/texlive/2013/texmf-dist/tex/latex/graphics/graphicx.sty
+Package: graphicx 1999/02/16 v1.0f Enhanced LaTeX Graphics (DPC,SPQR)
+
+(/usr/local/texlive/2013/texmf-dist/tex/latex/graphics/graphics.sty
+Package: graphics 2009/02/05 v1.0o Standard LaTeX Graphics (DPC,SPQR)
+
+(/usr/local/texlive/2013/texmf-dist/tex/latex/graphics/trig.sty
+Package: trig 1999/03/16 v1.09 sin cos tan (DPC)
+)
+(/usr/local/texlive/2013/texmf-dist/tex/latex/latexconfig/graphics.cfg
+File: graphics.cfg 2010/04/23 v1.9 graphics configuration of TeX Live
+)
+Package graphics Info: Driver file: pdftex.def on input line 91.
+
+(/usr/local/texlive/2013/texmf-dist/tex/latex/pdftex-def/pdftex.def
+File: pdftex.def 2011/05/27 v0.06d Graphics/color for pdfTeX
+
+(/usr/local/texlive/2013/texmf-dist/tex/generic/oberdiek/infwarerr.sty
+Package: infwarerr 2010/04/08 v1.3 Providing info/warning/error messages (HO)
+)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/oberdiek/ltxcmds.sty
+Package: ltxcmds 2011/11/09 v1.22 LaTeX kernel commands for general use (HO)
+)
+\Gread@gobject=\count105
+))
+\Gin@req@height=\dimen127
+\Gin@req@width=\dimen128
+)
+(/usr/local/texlive/2013/texmf-dist/tex/latex/pgf/systemlayer/pgfsys.sty
+(/usr/local/texlive/2013/texmf-dist/tex/latex/pgf/utilities/pgfrcs.sty
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/utilities/pgfutil-common.te
+x
+\pgfutil@everybye=\toks16
+)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/utilities/pgfutil-latex.def
+\pgfutil@abb=\box34
+(/usr/local/texlive/2013/texmf-dist/tex/latex/ms/everyshi.sty
+Package: everyshi 2001/05/15 v3.00 EveryShipout Package (MS)
+))
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/utilities/pgfrcs.code.tex
+Package: pgfrcs 2010/10/25 v2.10 (rcs-revision 1.24)
+))
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/systemlayer/pgfsys.code.tex
+Package: pgfsys 2010/06/30 v2.10 (rcs-revision 1.37)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/utilities/pgfkeys.code.tex
+\pgfkeys@pathtoks=\toks17
+\pgfkeys@temptoks=\toks18
+
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/utilities/pgfkeysfiltered.c
+ode.tex
+\pgfkeys@tmptoks=\toks19
+))
+\pgf@x=\dimen129
+\pgf@y=\dimen130
+\pgf@xa=\dimen131
+\pgf@ya=\dimen132
+\pgf@xb=\dimen133
+\pgf@yb=\dimen134
+\pgf@xc=\dimen135
+\pgf@yc=\dimen136
+\w@pgf@writea=\write3
+\r@pgf@reada=\read1
+\c@pgf@counta=\count106
+\c@pgf@countb=\count107
+\c@pgf@countc=\count108
+\c@pgf@countd=\count109
+
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/systemlayer/pgf.cfg
+File: pgf.cfg 2008/05/14 (rcs-revision 1.7)
+)
+Package pgfsys Info: Driver file for pgf: pgfsys-pdftex.def on input line 900.
+
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/systemlayer/pgfsys-pdftex.d
+ef
+File: pgfsys-pdftex.def 2009/05/22 (rcs-revision 1.26)
+
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/systemlayer/pgfsys-common-p
+df.def
+File: pgfsys-common-pdf.def 2008/05/19 (rcs-revision 1.10)
+)))
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/systemlayer/pgfsyssoftpath.
+code.tex
+File: pgfsyssoftpath.code.tex 2008/07/18 (rcs-revision 1.7)
+\pgfsyssoftpath@smallbuffer@items=\count110
+\pgfsyssoftpath@bigbuffer@items=\count111
+)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/systemlayer/pgfsysprotocol.
+code.tex
+File: pgfsysprotocol.code.tex 2006/10/16 (rcs-revision 1.4)
+)) (/usr/local/texlive/2013/texmf-dist/tex/latex/xcolor/xcolor.sty
+Package: xcolor 2007/01/21 v2.11 LaTeX color extensions (UK)
+
+(/usr/local/texlive/2013/texmf-dist/tex/latex/latexconfig/color.cfg
+File: color.cfg 2007/01/18 v1.5 color configuration of teTeX/TeXLive
+)
+Package xcolor Info: Driver file: pdftex.def on input line 225.
+Package xcolor Info: Model `cmy' substituted by `cmy0' on input line 1337.
+Package xcolor Info: Model `hsb' substituted by `rgb' on input line 1341.
+Package xcolor Info: Model `RGB' extended on input line 1353.
+Package xcolor Info: Model `HTML' substituted by `rgb' on input line 1355.
+Package xcolor Info: Model `Hsb' substituted by `hsb' on input line 1356.
+Package xcolor Info: Model `tHsb' substituted by `hsb' on input line 1357.
+Package xcolor Info: Model `HSB' substituted by `hsb' on input line 1358.
+Package xcolor Info: Model `Gray' substituted by `gray' on input line 1359.
+Package xcolor Info: Model `wave' substituted by `hsb' on input line 1360.
+)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/basiclayer/pgfcore.code.tex
+Package: pgfcore 2010/04/11 v2.10 (rcs-revision 1.7)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/math/pgfmath.code.tex
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/math/pgfmathcalc.code.tex
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/math/pgfmathutil.code.tex)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/math/pgfmathparser.code.tex
+\pgfmath@dimen=\dimen137
+\pgfmath@count=\count112
+\pgfmath@box=\box35
+\pgfmath@toks=\toks20
+\pgfmath@stack@operand=\toks21
+\pgfmath@stack@operation=\toks22
+)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/math/pgfmathfunctions.code.
+tex
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/math/pgfmathfunctions.basic
+.code.tex)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/math/pgfmathfunctions.trigo
+nometric.code.tex)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/math/pgfmathfunctions.rando
+m.code.tex)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/math/pgfmathfunctions.compa
+rison.code.tex)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/math/pgfmathfunctions.base.
+code.tex)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/math/pgfmathfunctions.round
+.code.tex)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/math/pgfmathfunctions.misc.
+code.tex)))
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/math/pgfmathfloat.code.tex
+\c@pgfmathroundto@lastzeros=\count113
+))
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/basiclayer/pgfcorepoints.co
+de.tex
+File: pgfcorepoints.code.tex 2010/04/09 (rcs-revision 1.20)
+\pgf@picminx=\dimen138
+\pgf@picmaxx=\dimen139
+\pgf@picminy=\dimen140
+\pgf@picmaxy=\dimen141
+\pgf@pathminx=\dimen142
+\pgf@pathmaxx=\dimen143
+\pgf@pathminy=\dimen144
+\pgf@pathmaxy=\dimen145
+\pgf@xx=\dimen146
+\pgf@xy=\dimen147
+\pgf@yx=\dimen148
+\pgf@yy=\dimen149
+\pgf@zx=\dimen150
+\pgf@zy=\dimen151
+)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/basiclayer/pgfcorepathconst
+ruct.code.tex
+File: pgfcorepathconstruct.code.tex 2010/08/03 (rcs-revision 1.24)
+\pgf@path@lastx=\dimen152
+\pgf@path@lasty=\dimen153
+)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/basiclayer/pgfcorepathusage
+.code.tex
+File: pgfcorepathusage.code.tex 2008/04/22 (rcs-revision 1.12)
+\pgf@shorten@end@additional=\dimen154
+\pgf@shorten@start@additional=\dimen155
+)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/basiclayer/pgfcorescopes.co
+de.tex
+File: pgfcorescopes.code.tex 2010/09/08 (rcs-revision 1.34)
+\pgfpic=\box36
+\pgf@hbox=\box37
+\pgf@layerbox@main=\box38
+\pgf@picture@serial@count=\count114
+)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/basiclayer/pgfcoregraphicst
+ate.code.tex
+File: pgfcoregraphicstate.code.tex 2008/04/22 (rcs-revision 1.9)
+\pgflinewidth=\dimen156
+)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/basiclayer/pgfcoretransform
+ations.code.tex
+File: pgfcoretransformations.code.tex 2009/06/10 (rcs-revision 1.11)
+\pgf@pt@x=\dimen157
+\pgf@pt@y=\dimen158
+\pgf@pt@temp=\dimen159
+)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/basiclayer/pgfcorequick.cod
+e.tex
+File: pgfcorequick.code.tex 2008/10/09 (rcs-revision 1.3)
+)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/basiclayer/pgfcoreobjects.c
+ode.tex
+File: pgfcoreobjects.code.tex 2006/10/11 (rcs-revision 1.2)
+)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/basiclayer/pgfcorepathproce
+ssing.code.tex
+File: pgfcorepathprocessing.code.tex 2008/10/09 (rcs-revision 1.8)
+)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/basiclayer/pgfcorearrows.co
+de.tex
+File: pgfcorearrows.code.tex 2008/04/23 (rcs-revision 1.11)
+)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/basiclayer/pgfcoreshade.cod
+e.tex
+File: pgfcoreshade.code.tex 2008/11/23 (rcs-revision 1.13)
+\pgf@max=\dimen160
+\pgf@sys@shading@range@num=\count115
+)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/basiclayer/pgfcoreimage.cod
+e.tex
+File: pgfcoreimage.code.tex 2010/03/25 (rcs-revision 1.16)
+
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/basiclayer/pgfcoreexternal.
+code.tex
+File: pgfcoreexternal.code.tex 2010/09/01 (rcs-revision 1.17)
+\pgfexternal@startupbox=\box39
+))
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/basiclayer/pgfcorelayers.co
+de.tex
+File: pgfcorelayers.code.tex 2010/08/27 (rcs-revision 1.2)
+)
+(/usr/local/texlive/2013/texmf-dist/tex/generic/pgf/basiclayer/pgfcoretranspare
+ncy.code.tex