added CPP slides
authorChristian Urban <christian dot urban at kcl dot ac dot uk>
Fri, 13 Dec 2013 10:37:25 +1100
changeset 18 cfd4b8219c87
parent 17 a87e2181d6b6
child 19 c3517b281164
added CPP slides
ROOT
Slides/Slides1.thy
Slides/document/beamerthemeplaincu.sty
Slides/document/build
Slides/document/chunhan.jpg
Slides/document/mathpartir.sty
Slides/document/root.tex
Slides/document/xingyuan.jpg
slides1.pdf
--- a/ROOT	Fri Sep 20 12:22:04 2013 +0100
+++ b/ROOT	Fri Dec 13 10:37:25 2013 +1100
@@ -10,4 +10,9 @@
   theories
     "Paper"
 
-
+session "Slides1" in Slides = RC +
+  options [document = pdf, document_output = "..", document_variants = "slides1"]
+  theories
+    "Slides1"
+  files
+    "document/root.tex"
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/Slides1.thy	Fri Dec 13 10:37:25 2013 +1100
@@ -0,0 +1,507 @@
+(*<*)
+theory Slides1
+imports "../rc_theory" "../final_theorems" "../rc_theory" "../os_rc" 
+begin
+
+notation (Rule output)
+  "==>"  ("\<^raw:\mbox{}\inferrule{\mbox{>_\<^raw:}}>\<^raw:{\mbox{>_\<^raw:}}>")
+
+syntax (Rule output)
+  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
+  ("\<^raw:\mbox{}\inferrule{>_\<^raw:}>\<^raw:{\mbox{>_\<^raw:}}>")
+
+  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" 
+  ("\<^raw:\mbox{>_\<^raw:}\\>/ _")
+
+  "_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>")
+
+notation (Axiom output)
+  "Trueprop"  ("\<^raw:\mbox{}\inferrule{\mbox{}}{\mbox{>_\<^raw:}}>")
+
+notation (IfThen output)
+  "==>"  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
+syntax (IfThen output)
+  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
+  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
+  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}> /\<^raw:{\normalsize \,>and\<^raw:\,}>/ _")
+  "_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>")
+
+notation (IfThenNoBox output)
+  "==>"  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
+syntax (IfThenNoBox output)
+  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
+  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
+  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("_ /\<^raw:{\normalsize \,>and\<^raw:\,}>/ _")
+  "_asm" :: "prop \<Rightarrow> asms" ("_")
+
+consts DUMMY::'a
+
+abbreviation
+  "is_parent f pf \<equiv> (parent f = Some pf)"
+
+context tainting_s_sound begin 
+
+notation (latex output)
+  source_dir ("anchor") and
+  SProc ("P_\<^bsup>_\<^esup>") and
+  SFile ("F_\<^bsup>_\<^esup>") and
+  SIPC ("I'(_')\<^bsup>_\<^esup>") and
+  READ ("Read") and 
+  WRITE ("Write") and
+  EXECUTE ("Execute") and
+  CHANGE_OWNER ("ChangeOwner") and
+  CREATE ("Create") and
+  SEND ("Send") and
+  RECEIVE ("Receive") and
+  DELETE ("Delete") and
+  compatible ("permissions") and
+  comproles ("compatible") and
+  DUMMY  ("\<^raw:\mbox{$\_$}>") and
+  Cons ("_::_" [78,77] 79) and 
+  Proc ("") and
+  File ("") and
+  File_type ("") and
+  Proc_type ("") and
+  IPC ("") and
+  init_processes ("init'_procs") and
+  os_grant ("admissible") and
+  rc_grant ("granted") and
+  exists ("alive") and
+  default_fd_create_type ("default'_type") and
+  InheritParent_file_type ("InheritPatentType") and
+  NormalFile_type ("NormalFileType") and
+  deleted ("deleted _ _" [50, 100] 100) and
+  taintable_s ("taintable\<^sup>s") and
+  tainted_s ("tainted\<^sup>s") and
+  all_sobjs ("reachable\<^sup>s") and
+  init_obj2sobj ("\<lbrakk>_\<rbrakk>") and
+  erole_functor ("erole'_aux") 
+
+declare [[show_question_marks = false]]
+
+abbreviation
+  "is_process_type s p t \<equiv> (type_of_process s p = Some t)"
+
+abbreviation
+  "is_current_role s p r \<equiv> (currentrole s p = Some r)"
+
+abbreviation
+  "is_file_type s f t \<equiv> (etype_of_file s f = Some t)"
+
+
+lemma osgrant10: (* modified by chunhan *)
+  "\<lbrakk>p \<in> current_procs \<tau>; p' \<notin> current_procs \<tau>\<rbrakk> \<Longrightarrow> os_grant \<tau> (Clone p p')"
+by simp 
+
+lemma rcgrant4:
+  "\<lbrakk>is_current_role s p r; is_file_type s f t; (r, File_type t, EXECUTE) \<in> compatible\<rbrakk>
+   \<Longrightarrow> rc_grant s (Execute p f)"
+by simp
+
+(*>*)
+
+text_raw {*
+  \tikzset{alt/.code args={<#1>#2#3#4}{%
+  \alt<#1>{\pgfkeysalso{#2}}{\pgfkeysalso{#3}} % \pgfkeysalso doesn't change the path
+  }}
+  \renewcommand{\slidecaption}{CPP, 13 December 2013}
+  \newcommand{\bl}[1]{\textcolor{blue}{#1}}                        
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+  \mode<presentation>{
+  \begin{frame}
+  \frametitle{%
+  \begin{tabular}{@ {\hspace{-2mm}}c@ {}}
+  \\[-3mm]
+  \Large A Formalisation of an\\[-1mm] 
+  \Large Access Control Framework
+  \end{tabular}}
+  
+  \begin{center}
+  \begin{tabular}{c@ {\hspace{15mm}}c}
+  \includegraphics[scale=0.034]{chunhan.jpg} &
+  \includegraphics[scale=0.034]{xingyuan.jpg}\\[-3mm]
+  \end{tabular}
+  \end{center}
+ 
+  \begin{center}
+  \small joint work with Chunhan Wu and Xingyuan Zhang from the PLA
+  University of Science and Technology in Nanjing
+  \end{center}
+
+  \begin{center}
+  \small Christian Urban\\
+  \small King's College London
+  \end{center}
+
+  \end{frame}}
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+*}
+
+text_raw {*
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+  \begin{frame}[c]
+  \frametitle{Access Control}
+
+\only<1>{
+  Perhaps most known are
+
+  \begin{itemize}
+  \item Unix-style access control systems\\ 
+  (root super-user, setuid mechanism)\bigskip
+
+  \begin{center}\footnotesize
+  \begin{semiverbatim}
+> ls -ld . * */*
+
+drwxr-xr-x 1 alice staff\ \ \ \  32768\ \ Apr\ \  2 2010 .
+
+-rw----r-- 1 alice students 31359\ \ Jul 24 2011 manual.txt
+
+-rwsr--r-x 1 bob\ \ \  students 141359 Jun\ \  1 2013 microedit
+
+dr--r-xr-x 1 bob\ \ \  staff\ \ \ \  32768\ \ Jul 23 2011 src
+
+-rw-r--r-- 1 bob\ \ \  staff\ \ \ \  81359\ \ Feb 28 2012 src/code.c
+\end{semiverbatim}
+  \end{center}
+\end{itemize}}
+
+\only<2->{
+More fine-grained access control is provided by\medskip
+
+\begin{itemize}
+  \item SELinux\\ (security enhanced Linux devloped by the NSA;\\ mandatory access control system)\bigskip 
+  \item Role-Compatibility Model\\ (developed by Amon Ott;\\ main application in the
+  Apache server) 
+  \end{itemize}}
+
+  \end{frame}
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
+
+*}
+
+text_raw {*
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+  \mode<presentation>{
+  \begin{frame}[c]
+  \frametitle{Operations in the OS}
+  \bigskip
+
+  using Paulson's inductive method a \alert{\bf state of the system} is 
+  a \alert{\bf trace}, a list of events (system calls):
+
+  \begin{center}\large
+  \bl{$[e_1,\ldots, e_2]$}
+  \end{center}\bigskip
+  
+
+   \begin{isabelle}\small
+  \mbox{
+  \begin{tabular}{@ {}r@ {\hspace{1.5mm}}c@ {\hspace{1.5mm}}l@ {\hspace{2mm}}l@ 
+                     {\hspace{1.5mm}}l@ {\hspace{2mm}}l@ {\hspace{1.5mm}}l@ 
+                     {\hspace{2mm}}l@ {}}
+  @{text e} 
+  & @{text "::="} & @{term "CreateFile p f"} & @{text "|"} & @{term "ReadFile p f"} & @{text "|"} & @{term "Send p i"}     \\
+  & @{text "|"} & @{term "WriteFile p f"}    & @{text "|"} & @{term "Execute p f"} & @{text "|"} & @{term "Recv p i"}\\
+  & @{text "|"} & @{term "DeleteFile p f"}   & @{text "|"} & @{term "Clone p p'"} & @{text "|"} & @{term "CreateIPC p i"}  \\
+  & @{text "|"} & @{term "ChangeOwner p u"}  & @{text "|"} & @{term "ChangeRole p r"} & @{text "|"} & @{term "DeleteIPC p i"}\\
+  & @{text "|"} & @{term "Kill p p'"}
+  \end{tabular}}
+  \end{isabelle}
+  \end{frame}}
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
+
+*}
+
+
+
+text_raw {*
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+  \mode<presentation>{
+  \begin{frame}[c]
+  \frametitle{Valid Traces}
+
+  we need to restrict the traces to \alert{\bf valid traces}:\bigskip\bigskip\bigskip 
+
+  
+  \begin{center}
+  \begin{tabular}{@ {}c@ {}}
+  \bl{@{thm [mode=Axiom] valid.intros(1)}}\hspace{5mm}
+  \bl{@{thm [mode=Rule] valid.intros(2)}} 
+  \end{tabular}
+  \end{center}
+
+  \begin{textblock}{3}(9,6.1)
+  \onslide<2-3>{
+  \begin{tikzpicture}
+  \node at (0,0) [single arrow, fill=red,text=white, rotate=50, shape border rotate=180]{OS};
+  \end{tikzpicture}}
+  \end{textblock}
+
+  \begin{textblock}{3}(6,11.2)
+  \only<3>{
+  \begin{tikzpicture}
+  \draw (0,0) node[inner sep=2mm,fill=cream, ultra thick, draw=red, rounded corners=2mm] 
+  {\normalsize\color{darkgray}
+  \begin{minipage}{5.6cm}\raggedright
+  \bl{@{thm[mode=Rule] osgrant10[of _ "s"]}}
+  \end{minipage}};
+  \end{tikzpicture}}
+  \end{textblock}
+
+  \begin{textblock}{6}(12.6,6.1)
+  \onslide<4->{
+  \begin{tikzpicture}
+  \node at (0,0) [single arrow, fill=red,text=white, rotate=50, shape border rotate=180]{RC};
+  \end{tikzpicture}}
+  \end{textblock}
+
+  \begin{textblock}{3}(4,11.3)
+  \only<4->{
+  \begin{tikzpicture}
+  \draw (0,0) node[inner sep=2mm,fill=cream, ultra thick, draw=red, rounded corners=2mm] 
+  {\normalsize\color{darkgray}
+  \begin{minipage}{8.5cm}\raggedright
+  \bl{@{thm[mode=Rule] rcgrant4}}
+  \end{minipage}};
+  \end{tikzpicture}}
+  \end{textblock}
+
+
+  \end{frame}}
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
+
+*}
+
+
+text_raw {*
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+  \mode<presentation>{
+  \begin{frame}[c]
+  \frametitle{Design of AC-Policies}
+
+  \Large
+  \begin{quote}
+  ''what you specify is what you get but not necessarily what you want\ldots''
+  \end{quote}
+
+  \end{frame}}
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
+
+*}
+
+text_raw {*
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+  \mode<presentation>{
+  \begin{frame}[c]
+  \frametitle{Testing Policies}
+
+  \begin{center}
+  \begin{tikzpicture}[scale=1]
+  %\draw[black!10,step=2mm] (-5,-3) grid (5,4);
+  %\draw[black!10,thick,step=10mm] (-5,-3) grid (5,4);
+  \draw[white,thick,step=10mm] (-5,-3) grid (5,4);
+
+  \draw [black!20, line width=1mm] (0,0) circle (1cm);
+  \draw [line width=1mm] (0,0) circle (3cm);
+  \node [black!20] at (0,0) {\begin{tabular}{c}core\\[-1mm] system\end{tabular}};
+  
+  \draw [red, line width=2mm, <-] (-2.1,2.1) -- (-3.5,3.2);
+  \node at (-3.5,3.6) {your system};
+  \node at (-4.8,0) {\Large policy $+$};
+
+  
+  \only<2>{
+  \draw [black, fill=red, line width=0.5mm] (2,1) circle (3mm);
+  \draw [red, line width=2mm, <-] (2.3,1.2) -- (3.5,2);
+  \node at (3.8,2.3) {a seed};}
+
+  \only<3>{
+  \draw [black, fill=red, line width=0.5mm] (2,1) circle (7mm);
+  \node[white] at (2,1) {\small tainted};}
+
+  \only<4>{
+  \begin{scope}
+  \draw [clip] (0,0) circle (2.955cm);
+  \draw [black, fill=red, line width=0.5mm] (2,1) circle (9mm);
+  \node[white] at (2,1) {\small tainted};
+  \end{scope}}
+
+  \only<5->{
+  \begin{scope}
+  \draw [clip] (0,0) circle (2.955cm);
+  \draw [black, fill=red, line width=0.5mm] (2,1) circle (13mm);
+  \node[white] at (2,1) {\small tainted};
+  \end{scope}}
+
+  \only<6->{
+  \draw[fill=white, line width=1mm] (1.265,2.665) arc (-35:183:5mm);
+  \draw[fill=white, line width=1mm] (1.25,3.25) arc (-35:183:3mm);
+  \node[black, rotate=10] at (1.9,3.6) {\LARGE\ldots};
+  }
+
+  \end{tikzpicture}
+  \end{center}
+
+
+  \begin{textblock}{3}(1,11.9)
+  \only<5->{
+  \begin{tikzpicture}
+  \draw (0,0) node[inner sep=2mm,fill=cream, ultra thick, draw=red, rounded corners=2mm] 
+  {\normalsize\color{darkgray}
+  \begin{minipage}{6.9cm}\raggedright\small
+  \bl{@{thm [mode=Rule] tainted.intros(6)}}
+  \end{minipage}};
+  \end{tikzpicture}}
+  \end{textblock}
+
+  \end{frame}}
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
+*}
+
+text_raw {*
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+  \mode<presentation>{
+  \begin{frame}[c]
+  \frametitle{\begin{tabular}{@ {}c@ {}}A Sound and Complete Test\end{tabular}}
+
+  \begin{itemize}
+  \item working purely in the \emph{dynamic world} does not work -\!-\!- infinite state space\bigskip
+
+  \item working purely on \emph{static} policies also does not\\ work -\!-\!- because of over approximation
+  \smallskip
+  \begin{itemize}
+  \item suppose a tainted file has type \emph{bin} and
+  \item there is a role \bl{$r$} which can both read and write \emph{bin}-files\pause
+  \item then we would conclude that the tainted file can spread\medskip\pause
+  \item but if there is no process with role \bl{$r$} and it will never been
+  created, then the file actually does not spread
+  \end{itemize}\bigskip\pause
+
+  \item \alert{our solution:} take a middle ground and record precisely the information
+  of the initial state, but be less precise about every newly created object.
+  \end{itemize}
+
+  \end{frame}}
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
+*}
+
+text_raw {*
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+  \mode<presentation>{
+  \begin{frame}[c]
+  \frametitle{Results about our Test}
+
+  \begin{itemize}
+  \item we can show that the objects (files, processes, \ldots) we need to consider 
+  are only finite (at some point
+  it does not matter if we create another \emph{bin}-file)
+  \end{itemize}
+
+  \colorbox{cream}{
+  \begin{minipage}{10.5cm}
+  {\bf Thm (Soundness)}\\ If our test says an object is taintable, then 
+  it is taintable, and we produce a sequence of events that show how it can
+  be tainted.
+  \end{minipage}}\bigskip\pause
+
+  \colorbox{cream}{
+  \begin{minipage}{10.5cm}
+  {\bf Thm (Completeness)}\\ 
+  If an object is taintable and \emph{undeletable$^\star$}, then 
+  our test will find out that it is taintable.
+  \end{minipage}}\medskip
+
+  \small
+  $^\star$ an object is \emph{undeleteable} if it exists in the initial state,
+  but there exists no valid state in which it could have been deleted
+
+  \end{frame}}
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
+*}
+
+text_raw {*
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+  \mode<presentation>{
+  \begin{frame}[c]
+  \frametitle{Why the Restriction?}
+
+  \begin{itemize}
+  \item assume a process with \bl{\emph{ID}} is tainted but gets killed by another process
+  \item after that a proces with \bl{\emph{ID}} gets \emph{re-created} by cloning an untainted process\medskip
+  \item clearly the new process should be considered untainted\pause
+  \end{itemize}
+
+  unfortunately our test will not be able to detect the difference (we are 
+  less precise about newly created processes)\bigskip\medskip\pause
+
+  \alert{Is this a serious restriction? We think not \ldots}
+  
+  \end{frame}}
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
+*}
+
+text_raw {*
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+  \mode<presentation>{
+  \begin{frame}[c]
+  \frametitle{Core System}
+
+  Admins usually ask whether their policy is strong enough to protect their
+  core system?
+  
+  \begin{center}
+  \begin{tikzpicture}[scale=0.8]
+
+  \draw [black!100, line width=1mm, fill=red] (0,0) circle (1.2cm);
+  \draw [line width=1mm] (0,0) circle (3cm);
+  \node [white] at (0,0) {\begin{tabular}{c}core\\[-1mm] system\end{tabular}};
+  
+  \begin{scope}
+  \draw [clip] (0,0) circle (2.955cm);
+  \draw [black, fill=red, line width=0.5mm] (2,1) circle (10mm);
+  \node[white] at (2,1) {\small tainted};
+  \end{scope}
+
+  \end{tikzpicture}
+  \end{center}
+
+  core system files are typically undeletable
+
+  \end{frame}}
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
+*}
+
+text_raw {*
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+  \mode<presentation>{
+  \begin{frame}[c]
+  \frametitle{Conclusion}
+
+  \begin{itemize}
+  \item we considered Role-Compatibility Model for the Apache Server\medskip \\
+  13 events, 13 rules for OS admisibility, 14 rules for RC-granting, 10 rules for 
+  tainted
+  \end{itemize}\bigskip
+
+  \begin{itemize}
+  \item we can scale this to SELinux\medskip\\
+  more fine-grainded OS events (inodes, hard-links, shared memory, \ldots)
+  \end{itemize}\pause\bigskip
+
+  \begin{itemize}
+  \item hard sell to Ott (who designed the RC-model)
+  \item hard sell to the community working on access control (beyond \emph{good science})
+  \end{itemize}
+  \end{frame}}
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
+*}
+
+
+
+(*<*)
+end
+end
+(*>*)
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/document/beamerthemeplaincu.sty	Fri Dec 13 10:37:25 2013 +1100
@@ -0,0 +1,108 @@
+%%\Providespackage{beamerthemeplainculight}[2003/11/07 ver 0.93]
+\NeedsTeXFormat{LaTeX2e}[1995/12/01]
+
+% Copyright 2003 by Till Tantau <tantau@cs.tu-berlin.de>.
+%
+% This program can be redistributed and/or modified under the terms
+% of the LaTeX Project Public License Distributed from CTAN
+% archives in directory macros/latex/base/lppl.txt.
+ 
+\newcommand{\slidecaption}{}
+
+\mode<presentation>
+
+\usepackage{fontspec}
+\usefonttheme{serif}
+\setmainfont{Hoefler Text}
+\renewcommand\ttdefault{lmtt}
+
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+% comic fonts fonts
+%\DeclareFontFamily{T1}{comic}{}%
+%\DeclareFontShape{T1}{comic}{m}{n}{<->s*[.9]comic8t}{}%
+%\DeclareFontShape{T1}{comic}{m}{it}{<->s*[.9]comic8t}{}%
+%\DeclareFontShape{T1}{comic}{m}{sc}{<->s*[.9]comic8t}{}%
+%\DeclareFontShape{T1}{comic}{b}{n}{<->s*[.9]comicbd8t}{}%
+%\DeclareFontShape{T1}{comic}{b}{it}{<->s*[.9]comicbd8t}{}%
+%\DeclareFontShape{T1}{comic}{m}{sl}{<->ssub * comic/m/it}{}%
+%\DeclareFontShape{T1}{comic}{b}{sc}{<->sub * comic/m/sc}{}%
+%\DeclareFontShape{T1}{comic}{b}{sl}{<->ssub * comic/b/it}{}%
+%\DeclareFontShape{T1}{comic}{bx}{n}{<->ssub * comic/b/n}{}%
+%\DeclareFontShape{T1}{comic}{bx}{it}{<->ssub * comic/b/it}{}%
+%\DeclareFontShape{T1}{comic}{bx}{sc}{<->sub * comic/m/sc}{}%
+%\DeclareFontShape{T1}{comic}{bx}{sl}{<->ssub * comic/b/it}{}%
+%
+%\renewcommand{\rmdefault}{comic}%
+%\renewcommand{\sfdefault}{comic}%
+\renewcommand{\mathfamilydefault}{cmr}% mathfont should be still the old one
+%
+\DeclareMathVersion{bold}% mathfont needs to be bold
+\DeclareSymbolFont{operators}{OT1}{cmr}{b}{n}%
+\SetSymbolFont{operators}{bold}{OT1}{cmr}{b}{n}%
+\DeclareSymbolFont{letters}{OML}{cmm}{b}{it}%
+\SetSymbolFont{letters}{bold}{OML}{cmm}{b}{it}%
+\DeclareSymbolFont{symbols}{OMS}{cmsy}{b}{n}%
+\SetSymbolFont{symbols}{bold}{OMS}{cmsy}{b}{n}%
+\DeclareSymbolFont{largesymbols}{OMX}{cmex}{b}{n}%
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+
+
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+% Frametitles
+
+\setbeamerfont{frametitle}{size={\LARGE}}
+\setbeamerfont{frametitle}{family={\fontspec{Hoefler Text Black}}}
+%\setbeamerfont{frametitle}{family={\usefont{T1}{ptm}{b}{n}}
+\setbeamercolor{frametitle}{fg=ProcessBlue,bg=white}
+
+\setbeamertemplate{frametitle}{%
+\vskip 2mm  % distance from the top margin
+\hskip -3mm % distance from left margin
+\vbox{%
+\begin{minipage}{1.05\textwidth}%
+\centering%
+\insertframetitle%
+\end{minipage}}%
+}
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+
+
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+% Foot
+%
+\setbeamertemplate{navigation symbols}{} 
+\usefoottemplate{%
+\vbox{%
+  \tinyline{%
+    \tiny\hfill\textcolor{gray!50}{\slidecaption{} --
+      p.~\insertframenumber/\inserttotalframenumber}}}%
+}
+
+
+
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+\beamertemplateballitem
+\setlength\leftmargini{2mm}
+\setlength\leftmarginii{0.6cm}
+\setlength\leftmarginiii{1.5cm}
+
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+% blocks
+%\definecolor{cream}{rgb}{1,1,.65}
+\definecolor{cream}{rgb}{1,1,.8}
+\setbeamerfont{block title}{size=\normalsize}
+\setbeamercolor{block title}{fg=black,bg=cream}
+\setbeamercolor{block body}{fg=black,bg=cream}
+
+\setbeamertemplate{blocks}[rounded][shadow=true]
+
+\setbeamercolor{boxcolor}{fg=black,bg=cream}
+
+\mode
+<all>
+
+
+
+
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/document/build	Fri Dec 13 10:37:25 2013 +1100
@@ -0,0 +1,11 @@
+#!/bin/bash
+
+set -e
+
+FORMAT="$1"
+VARIANT="$2"
+
+ISABELLE_PDFLATEX="xelatex"
+
+"$ISABELLE_TOOL" latex -o "$FORMAT"
+"$ISABELLE_TOOL" latex -o "$FORMAT"
Binary file Slides/document/chunhan.jpg has changed
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/document/mathpartir.sty	Fri Dec 13 10:37:25 2013 +1100
@@ -0,0 +1,446 @@
+%  Mathpartir --- Math Paragraph for Typesetting Inference Rules
+%
+%  Copyright (C) 2001, 2002, 2003, 2004, 2005 Didier Rémy
+%
+%  Author         : Didier Remy 
+%  Version        : 1.2.0
+%  Bug Reports    : to author
+%  Web Site       : http://pauillac.inria.fr/~remy/latex/
+% 
+%  Mathpartir 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, or (at your option)
+%  any later version.
+%  
+%  Mathpartir 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 
+%  (http://pauillac.inria.fr/~remy/license/GPL).
+%
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+%  File mathpartir.sty (LaTeX macros)
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+
+\NeedsTeXFormat{LaTeX2e}
+\ProvidesPackage{mathpartir}
+    [2005/12/20 version 1.2.0 Math Paragraph for Typesetting Inference Rules]
+
+%%
+
+%% Identification
+%% Preliminary declarations
+
+\RequirePackage {keyval}
+
+%% Options
+%% More declarations
+
+%% PART I: Typesetting maths in paragraphe mode
+
+%% \newdimen \mpr@tmpdim
+%% Dimens are a precious ressource. Uses seems to be local.
+\let \mpr@tmpdim \@tempdima
+
+% To ensure hevea \hva compatibility, \hva should expands to nothing 
+% in mathpar or in inferrule
+\let \mpr@hva \empty
+
+%% normal paragraph parametters, should rather be taken dynamically
+\def \mpr@savepar {%
+  \edef \MathparNormalpar
+     {\noexpand \lineskiplimit \the\lineskiplimit
+      \noexpand \lineskip \the\lineskip}%
+  }
+
+\def \mpr@rulelineskip {\lineskiplimit=0.3em\lineskip=0.2em plus 0.1em}
+\def \mpr@lesslineskip {\lineskiplimit=0.6em\lineskip=0.5em plus 0.2em}
+\def \mpr@lineskip  {\lineskiplimit=1.2em\lineskip=1.2em plus 0.2em}
+\let \MathparLineskip \mpr@lineskip
+\def \mpr@paroptions {\MathparLineskip}
+\let \mpr@prebindings \relax
+
+\newskip \mpr@andskip \mpr@andskip 2em plus 0.5fil minus 0.5em
+
+\def \mpr@goodbreakand
+   {\hskip -\mpr@andskip  \penalty -1000\hskip \mpr@andskip}
+\def \mpr@and {\hskip \mpr@andskip}
+\def \mpr@andcr {\penalty 50\mpr@and}
+\def \mpr@cr {\penalty -10000\mpr@and}
+\def \mpr@eqno #1{\mpr@andcr #1\hskip 0em plus -1fil \penalty 10}
+
+\def \mpr@bindings {%
+  \let \and \mpr@andcr
+  \let \par \mpr@andcr
+  \let \\\mpr@cr
+  \let \eqno \mpr@eqno
+  \let \hva \mpr@hva
+  } 
+\let \MathparBindings \mpr@bindings
+
+% \@ifundefined {ignorespacesafterend}
+%    {\def \ignorespacesafterend {\aftergroup \ignorespaces}
+
+\newenvironment{mathpar}[1][]
+  {$$\mpr@savepar \parskip 0em \hsize \linewidth \centering
+     \vbox \bgroup \mpr@prebindings \mpr@paroptions #1\ifmmode $\else
+     \noindent $\displaystyle\fi
+     \MathparBindings}
+  {\unskip \ifmmode $\fi\egroup $$\ignorespacesafterend}
+
+\newenvironment{mathparpagebreakable}[1][]
+  {\begingroup 
+   \par
+   \mpr@savepar \parskip 0em \hsize \linewidth \centering
+      \mpr@prebindings \mpr@paroptions #1%
+      \vskip \abovedisplayskip \vskip -\lineskip%
+     \ifmmode  \else  $\displaystyle\fi
+     \MathparBindings
+  }
+  {\unskip
+   \ifmmode $\fi \par\endgroup
+   \vskip \belowdisplayskip
+   \noindent
+  \ignorespacesafterend}
+
+% \def \math@mathpar #1{\setbox0 \hbox {$\displaystyle #1$}\ifnum
+%     \wd0 < \hsize  $$\box0$$\else \bmathpar #1\emathpar \fi}
+
+%%% HOV BOXES
+
+\def \mathvbox@ #1{\hbox \bgroup \mpr@normallineskip 
+  \vbox \bgroup \tabskip 0em \let \\ \cr
+  \halign \bgroup \hfil $##$\hfil\cr #1\crcr \egroup \egroup
+  \egroup}
+
+\def \mathhvbox@ #1{\setbox0 \hbox {\let \\\qquad $#1$}\ifnum \wd0 < \hsize
+      \box0\else \mathvbox {#1}\fi}
+
+
+%% Part II -- operations on lists
+
+\newtoks \mpr@lista
+\newtoks \mpr@listb
+
+\long \def\mpr@cons #1\mpr@to#2{\mpr@lista {\\{#1}}\mpr@listb \expandafter
+{#2}\edef #2{\the \mpr@lista \the \mpr@listb}}
+
+\long \def\mpr@snoc #1\mpr@to#2{\mpr@lista {\\{#1}}\mpr@listb \expandafter
+{#2}\edef #2{\the \mpr@listb\the\mpr@lista}}
+
+\long \def \mpr@concat#1=#2\mpr@to#3{\mpr@lista \expandafter {#2}\mpr@listb
+\expandafter {#3}\edef #1{\the \mpr@listb\the\mpr@lista}}
+
+\def \mpr@head #1\mpr@to #2{\expandafter \mpr@head@ #1\mpr@head@ #1#2}
+\long \def \mpr@head@ #1#2\mpr@head@ #3#4{\def #4{#1}\def#3{#2}}
+
+\def \mpr@flatten #1\mpr@to #2{\expandafter \mpr@flatten@ #1\mpr@flatten@ #1#2}
+\long \def \mpr@flatten@ \\#1\\#2\mpr@flatten@ #3#4{\def #4{#1}\def #3{\\#2}}
+
+\def \mpr@makelist #1\mpr@to #2{\def \mpr@all {#1}%
+   \mpr@lista {\\}\mpr@listb \expandafter {\mpr@all}\edef \mpr@all {\the
+   \mpr@lista \the \mpr@listb \the \mpr@lista}\let #2\empty 
+   \def \mpr@stripof ##1##2\mpr@stripend{\def \mpr@stripped{##2}}\loop
+     \mpr@flatten \mpr@all \mpr@to \mpr@one
+     \expandafter \mpr@snoc \mpr@one \mpr@to #2\expandafter \mpr@stripof
+     \mpr@all \mpr@stripend  
+     \ifx \mpr@stripped \empty \let \mpr@isempty 0\else \let \mpr@isempty 1\fi
+     \ifx 1\mpr@isempty
+   \repeat
+}
+
+\def \mpr@rev #1\mpr@to #2{\let \mpr@tmp \empty
+   \def \\##1{\mpr@cons ##1\mpr@to \mpr@tmp}#1\let #2\mpr@tmp}
+
+%% Part III -- Type inference rules
+
+\newif \if@premisse
+\newbox \mpr@hlist
+\newbox \mpr@vlist
+\newif \ifmpr@center \mpr@centertrue
+\def \mpr@htovlist {%
+   \setbox \mpr@hlist
+      \hbox {\strut
+             \ifmpr@center \hskip -0.5\wd\mpr@hlist\fi
+             \unhbox \mpr@hlist}%
+   \setbox \mpr@vlist
+      \vbox {\if@premisse  \box \mpr@hlist \unvbox \mpr@vlist
+             \else \unvbox \mpr@vlist \box \mpr@hlist
+             \fi}%
+}
+% OLD version
+% \def \mpr@htovlist {%
+%    \setbox \mpr@hlist
+%       \hbox {\strut \hskip -0.5\wd\mpr@hlist \unhbox \mpr@hlist}%
+%    \setbox \mpr@vlist
+%       \vbox {\if@premisse  \box \mpr@hlist \unvbox \mpr@vlist
+%              \else \unvbox \mpr@vlist \box \mpr@hlist
+%              \fi}%
+% }
+
+\def \mpr@item #1{$\displaystyle #1$}
+\def \mpr@sep{2em}
+\def \mpr@blank { }
+\def \mpr@hovbox #1#2{\hbox
+  \bgroup
+  \ifx #1T\@premissetrue
+  \else \ifx #1B\@premissefalse
+  \else
+     \PackageError{mathpartir}
+       {Premisse orientation should either be T or B}
+       {Fatal error in Package}%
+  \fi \fi
+  \def \@test {#2}\ifx \@test \mpr@blank\else
+  \setbox \mpr@hlist \hbox {}%
+  \setbox \mpr@vlist \vbox {}%
+  \if@premisse \let \snoc \mpr@cons \else \let \snoc \mpr@snoc \fi
+  \let \@hvlist \empty \let \@rev \empty
+  \mpr@tmpdim 0em
+  \expandafter \mpr@makelist #2\mpr@to \mpr@flat
+  \if@premisse \mpr@rev \mpr@flat \mpr@to \@rev \else \let \@rev \mpr@flat \fi
+  \def \\##1{%
+     \def \@test {##1}\ifx \@test \empty
+        \mpr@htovlist
+        \mpr@tmpdim 0em %%% last bug fix not extensively checked
+     \else
+      \setbox0 \hbox{\mpr@item {##1}}\relax
+      \advance \mpr@tmpdim by \wd0
+      %\mpr@tmpdim 1.02\mpr@tmpdim
+      \ifnum \mpr@tmpdim < \hsize
+         \ifnum \wd\mpr@hlist > 0
+           \if@premisse
+             \setbox \mpr@hlist 
+                \hbox {\unhbox0 \hskip \mpr@sep \unhbox \mpr@hlist}%
+           \else
+             \setbox \mpr@hlist
+                \hbox {\unhbox \mpr@hlist  \hskip \mpr@sep \unhbox0}%
+           \fi
+         \else 
+         \setbox \mpr@hlist \hbox {\unhbox0}%
+         \fi
+      \else
+         \ifnum \wd \mpr@hlist > 0
+            \mpr@htovlist 
+            \mpr@tmpdim \wd0
+         \fi
+         \setbox \mpr@hlist \hbox {\unhbox0}%
+      \fi
+      \advance \mpr@tmpdim by \mpr@sep
+   \fi
+   }%
+   \@rev
+   \mpr@htovlist
+   \ifmpr@center \hskip \wd\mpr@vlist\fi \box \mpr@vlist
+   \fi
+   \egroup
+}
+
+%%% INFERENCE RULES
+
+\@ifundefined{@@over}{%
+    \let\@@over\over % fallback if amsmath is not loaded
+    \let\@@overwithdelims\overwithdelims
+    \let\@@atop\atop \let\@@atopwithdelims\atopwithdelims
+    \let\@@above\above \let\@@abovewithdelims\abovewithdelims
+  }{}
+
+%% The default
+
+\def \mpr@@fraction #1#2{\hbox {\advance \hsize by -0.5em
+    $\displaystyle {#1\mpr@over #2}$}}
+\def \mpr@@nofraction #1#2{\hbox {\advance \hsize by -0.5em
+    $\displaystyle {#1\@@atop #2}$}}
+
+\let \mpr@fraction \mpr@@fraction
+
+%% A generic solution to arrow
+
+\def \mpr@make@fraction #1#2#3#4#5{\hbox {%
+     \def \mpr@tail{#1}%
+     \def \mpr@body{#2}%
+     \def \mpr@head{#3}%
+     \setbox1=\hbox{$#4$}\setbox2=\hbox{$#5$}%
+     \setbox3=\hbox{$\mkern -3mu\mpr@body\mkern -3mu$}%
+     \setbox3=\hbox{$\mkern -3mu \mpr@body\mkern -3mu$}%
+     \dimen0=\dp1\advance\dimen0 by \ht3\relax\dp1\dimen0\relax
+     \dimen0=\ht2\advance\dimen0 by \dp3\relax\ht2\dimen0\relax
+     \setbox0=\hbox {$\box1 \@@atop \box2$}%
+     \dimen0=\wd0\box0
+     \box0 \hskip -\dimen0\relax
+     \hbox to \dimen0 {$%
+       \mathrel{\mpr@tail}\joinrel
+       \xleaders\hbox{\copy3}\hfil\joinrel\mathrel{\mpr@head}%
+     $}}}
+
+%% Old stuff should be removed in next version
+\def \mpr@@nothing #1#2
+    {$\lower 0.01pt \mpr@@nofraction {#1}{#2}$}
+\def \mpr@@reduce #1#2{\hbox
+    {$\lower 0.01pt \mpr@@fraction {#1}{#2}\mkern -15mu\rightarrow$}}
+\def \mpr@@rewrite #1#2#3{\hbox
+    {$\lower 0.01pt \mpr@@fraction {#2}{#3}\mkern -8mu#1$}}
+\def \mpr@infercenter #1{\vcenter {\mpr@hovbox{T}{#1}}}
+
+\def \mpr@empty {}
+\def \mpr@inferrule
+  {\bgroup
+     \ifnum \linewidth<\hsize \hsize \linewidth\fi
+     \mpr@rulelineskip
+     \let \and \qquad
+     \let \hva \mpr@hva
+     \let \@rulename \mpr@empty
+     \let \@rule@options \mpr@empty
+     \let \mpr@over \@@over
+     \mpr@inferrule@}
+\newcommand {\mpr@inferrule@}[3][]
+  {\everymath={\displaystyle}%       
+   \def \@test {#2}\ifx \empty \@test
+      \setbox0 \hbox {$\vcenter {\mpr@hovbox{B}{#3}}$}%
+   \else 
+   \def \@test {#3}\ifx \empty \@test
+      \setbox0 \hbox {$\vcenter {\mpr@hovbox{T}{#2}}$}%
+   \else
+   \setbox0 \mpr@fraction {\mpr@hovbox{T}{#2}}{\mpr@hovbox{B}{#3}}%
+   \fi \fi
+   \def \@test {#1}\ifx \@test\empty \box0
+   \else \vbox 
+%%% Suggestion de Francois pour les etiquettes longues
+%%%   {\hbox to \wd0 {\RefTirName {#1}\hfil}\box0}\fi
+      {\hbox {\RefTirName {#1}}\box0}\fi
+   \egroup}
+
+\def \mpr@vdotfil #1{\vbox to #1{\leaders \hbox{$\cdot$} \vfil}}
+
+% They are two forms
+% \inferrule [label]{[premisses}{conclusions}
+% or
+% \inferrule* [options]{[premisses}{conclusions}
+%
+% Premisses and conclusions are lists of elements separated by \\
+% Each \\ produces a break, attempting horizontal breaks if possible, 
+% and  vertical breaks if needed. 
+% 
+% An empty element obtained by \\\\ produces a vertical break in all cases. 
+%
+% The former rule is aligned on the fraction bar. 
+% The optional label appears on top of the rule
+% The second form to be used in a derivation tree is aligned on the last
+% line of its conclusion
+% 
+% The second form can be parameterized, using the key=val interface. The
+% folloiwng keys are recognized:
+%       
+%  width                set the width of the rule to val
+%  narrower             set the width of the rule to val\hsize
+%  before               execute val at the beginning/left
+%  lab                  put a label [Val] on top of the rule
+%  lskip                add negative skip on the right
+%  left                 put a left label [Val]
+%  Left                 put a left label [Val],  ignoring its width 
+%  right                put a right label [Val]
+%  Right                put a right label [Val], ignoring its width
+%  leftskip             skip negative space on the left-hand side
+%  rightskip            skip negative space on the right-hand side
+%  vdots                lift the rule by val and fill vertical space with dots
+%  after                execute val at the end/right
+%  
+%  Note that most options must come in this order to avoid strange
+%  typesetting (in particular  leftskip must preceed left and Left and
+%  rightskip must follow Right or right; vdots must come last 
+%  or be only followed by rightskip. 
+%  
+
+%% Keys that make sence in all kinds of rules
+\def \mprset #1{\setkeys{mprset}{#1}}
+\define@key {mprset}{andskip}[]{\mpr@andskip=#1}
+\define@key {mprset}{lineskip}[]{\lineskip=#1}
+\define@key {mprset}{flushleft}[]{\mpr@centerfalse}
+\define@key {mprset}{center}[]{\mpr@centertrue}
+\define@key {mprset}{rewrite}[]{\let \mpr@fraction \mpr@@rewrite}
+\define@key {mprset}{atop}[]{\let \mpr@fraction \mpr@@nofraction}
+\define@key {mprset}{myfraction}[]{\let \mpr@fraction #1}
+\define@key {mprset}{fraction}[]{\def \mpr@fraction {\mpr@make@fraction #1}}
+\define@key {mprset}{sep}{\def\mpr@sep{#1}}
+
+\newbox \mpr@right
+\define@key {mpr}{flushleft}[]{\mpr@centerfalse}
+\define@key {mpr}{center}[]{\mpr@centertrue}
+\define@key {mpr}{rewrite}[]{\let \mpr@fraction \mpr@@rewrite}
+\define@key {mpr}{myfraction}[]{\let \mpr@fraction #1}
+\define@key {mpr}{fraction}[]{\def \mpr@fraction {\mpr@make@fraction #1}}
+\define@key {mpr}{left}{\setbox0 \hbox {$\TirName {#1}\;$}\relax
+     \advance \hsize by -\wd0\box0}
+\define@key {mpr}{width}{\hsize #1}
+\define@key {mpr}{sep}{\def\mpr@sep{#1}}
+\define@key {mpr}{before}{#1}
+\define@key {mpr}{lab}{\let \RefTirName \TirName \def \mpr@rulename {#1}}
+\define@key {mpr}{Lab}{\let \RefTirName \TirName \def \mpr@rulename {#1}}
+\define@key {mpr}{narrower}{\hsize #1\hsize}
+\define@key {mpr}{leftskip}{\hskip -#1}
+\define@key {mpr}{reduce}[]{\let \mpr@fraction \mpr@@reduce}
+\define@key {mpr}{rightskip}
+  {\setbox \mpr@right \hbox {\unhbox \mpr@right \hskip -#1}}
+\define@key {mpr}{LEFT}{\setbox0 \hbox {$#1$}\relax
+     \advance \hsize by -\wd0\box0}
+\define@key {mpr}{left}{\setbox0 \hbox {$\TirName {#1}\;$}\relax
+     \advance \hsize by -\wd0\box0}
+\define@key {mpr}{Left}{\llap{$\TirName {#1}\;$}}
+\define@key {mpr}{right}
+  {\setbox0 \hbox {$\;\TirName {#1}$}\relax \advance \hsize by -\wd0
+   \setbox \mpr@right \hbox {\unhbox \mpr@right \unhbox0}}
+\define@key {mpr}{RIGHT}
+  {\setbox0 \hbox {$#1$}\relax \advance \hsize by -\wd0
+   \setbox \mpr@right \hbox {\unhbox \mpr@right \unhbox0}}
+\define@key {mpr}{Right}
+  {\setbox \mpr@right \hbox {\unhbox \mpr@right \rlap {$\;\TirName {#1}$}}}
+\define@key {mpr}{vdots}{\def \mpr@vdots {\@@atop \mpr@vdotfil{#1}}}
+\define@key {mpr}{after}{\edef \mpr@after {\mpr@after #1}}
+
+\newcommand \mpr@inferstar@ [3][]{\setbox0
+  \hbox {\let \mpr@rulename \mpr@empty \let \mpr@vdots \relax
+         \setbox \mpr@right \hbox{}%
+         $\setkeys{mpr}{#1}%
+          \ifx \mpr@rulename \mpr@empty \mpr@inferrule {#2}{#3}\else
+          \mpr@inferrule [{\mpr@rulename}]{#2}{#3}\fi
+          \box \mpr@right \mpr@vdots$}
+  \setbox1 \hbox {\strut}
+  \@tempdima \dp0 \advance \@tempdima by -\dp1
+  \raise \@tempdima \box0}
+
+\def \mpr@infer {\@ifnextchar *{\mpr@inferstar}{\mpr@inferrule}}
+\newcommand \mpr@err@skipargs[3][]{}
+\def \mpr@inferstar*{\ifmmode 
+    \let \@do \mpr@inferstar@
+  \else 
+    \let \@do \mpr@err@skipargs
+    \PackageError {mathpartir}
+      {\string\inferrule* can only be used in math mode}{}%
+  \fi \@do}
+
+
+%%% Exports
+
+% Envirnonment mathpar
+
+\let \inferrule \mpr@infer
+
+% make a short name \infer is not already defined
+\@ifundefined {infer}{\let \infer \mpr@infer}{}
+
+\def \TirNameStyle #1{\small \textsc{#1}}
+\def \tir@name #1{\hbox {\small \TirNameStyle{#1}}}
+\let \TirName \tir@name
+\let \DefTirName \TirName
+\let \RefTirName \TirName
+
+%%% Other Exports
+
+% \let \listcons \mpr@cons
+% \let \listsnoc \mpr@snoc
+% \let \listhead \mpr@head
+% \let \listmake \mpr@makelist
+
+
+
+
+\endinput
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Slides/document/root.tex	Fri Dec 13 10:37:25 2013 +1100
@@ -0,0 +1,155 @@
+\documentclass[dvipsnames, 14pt, t, xelatex]{beamer}
+\usepackage{beamerthemeplaincu}
+%%\usepackage{ulem}
+%\usepackage[T1]{fontenc}
+\usepackage{proof}
+%\usepackage[latin1]{inputenc}
+\usepackage{isabelle}
+\usepackage{isabellesym}
+\usepackage{mathpartir}
+\usepackage[absolute, overlay]{textpos}
+\usepackage{proof}
+\usepackage{ifthen}
+\usepackage{animate}
+\usepackage{tikz}
+\usepackage{pgf}
+\usetikzlibrary{arrows}
+\usetikzlibrary{automata}
+\usetikzlibrary{shapes}
+\usetikzlibrary{shadows}
+\usetikzlibrary{calc}
+\usetikzlibrary{shapes.callouts,shapes.arrows,decorations.text}
+\usetikzlibrary{decorations.pathreplacing}
+
+% Isabelle configuration
+%%\urlstyle{rm}
+\isabellestyle{rm}
+\renewcommand{\isastyle}{\rm}%
+\renewcommand{\isastyleminor}{\fontspec{STIXGeneral}}%
+\renewcommand{\isastylescript}{\fontsize{8}{10}\selectfont\fontspec{STIXGeneral}\slshape}%
+
+
+\renewcommand{\isatagproof}{}
+\renewcommand{\endisatagproof}{}
+\renewcommand{\isamarkupcmt}[1]{#1}
+
+% Isabelle characters
+\renewcommand{\isacharunderscore}{\_}
+\renewcommand{\isacharbar}{\isamath{\mid}}
+\renewcommand{\isasymiota}{}
+\renewcommand{\isacharbraceleft}{\{}
+\renewcommand{\isacharbraceright}{\}}
+%\renewcommand{\isacharless}{$\langle$}
+%\renewcommand{\isachargreater}{$\rangle$}
+\renewcommand{\isasymsharp}{\isamath{\#}}
+\renewcommand{\isasymdots}{\isamath{...}}
+\renewcommand{\isasymbullet}{\act}
+\defaultfontfeatures{Ligatures=TeX}
+
+% mathpatir
+\mprset{sep=1em}
+
+% general math stuff
+\newcommand{\dn}{\stackrel{\mbox{\scriptsize def}}{=}}% for definitions
+\newcommand{\dnn}{\stackrel{\mbox{\Large def}}{=}}
+\renewcommand{\isasymequiv}{$\dn$}
+\renewcommand{\emptyset}{\varnothing}% nice round empty set
+%%\renewcommand{\Gamma}{\varGamma} 
+\DeclareRobustCommand{\flqq}{\mbox{\guillemotleft}}
+\DeclareRobustCommand{\frqq}{\mbox{\guillemotright}}
+\newcommand{\smath}[1]{\textcolor{blue}{\ensuremath{#1}}}
+\newcommand{\fresh}{\mathrel{\#}}
+\newcommand{\act}{{\raisebox{-0.5mm}{\Large$\boldsymbol{\cdot}$}}}% swapping action
+\newcommand{\swap}[2]{(#1\,#2)}% swapping operation
+
+% beamer stuff 
+\renewcommand{\slidecaption}{Salvador, 26.~August 2008}
+
+
+% colours for Isar Code (in article mode everything is black and white)
+\mode<presentation>{
+\definecolor{isacol:brown}{rgb}{.823,.411,.117}
+\definecolor{isacol:lightblue}{rgb}{.274,.509,.705}
+\definecolor{isacol:green}{rgb}{0,.51,0.14}
+\definecolor{isacol:red}{rgb}{.803,0,0}
+\definecolor{isacol:blue}{rgb}{0,0,.803}
+\definecolor{isacol:darkred}{rgb}{.545,0,0}
+\definecolor{isacol:black}{rgb}{0,0,0}}
+\mode<article>{
+\definecolor{isacol:brown}{rgb}{0,0,0}
+\definecolor{isacol:lightblue}{rgb}{0,0,0}
+\definecolor{isacol:green}{rgb}{0,0,0}
+\definecolor{isacol:red}{rgb}{0,0,0}
+\definecolor{isacol:blue}{rgb}{0,0,0}
+\definecolor{isacol:darkred}{rgb}{0,0,0}
+\definecolor{isacol:black}{rgb}{0,0,0}
+}
+
+
+\newcommand{\strong}[1]{{\bfseries {#1}}}
+\newcommand{\bluecmd}[1]{{\color{isacol:lightblue}{\strong{#1}}}}
+\newcommand{\browncmd}[1]{{\color{isacol:brown}{\strong{#1}}}}
+\newcommand{\redcmd}[1]{{\color{isacol:red}{\strong{#1}}}}
+
+\renewcommand{\isakeyword}[1]{%
+\ifthenelse{\equal{#1}{show}}{\browncmd{#1}}{%
+\ifthenelse{\equal{#1}{case}}{\browncmd{#1}}{%
+\ifthenelse{\equal{#1}{assume}}{\browncmd{#1}}{%
+\ifthenelse{\equal{#1}{obtain}}{\browncmd{#1}}{%
+\ifthenelse{\equal{#1}{fix}}{\browncmd{#1}}{%
+\ifthenelse{\equal{#1}{oops}}{\redcmd{#1}}{%
+\ifthenelse{\equal{#1}{thm}}{\redcmd{#1}}{%
+{\bluecmd{#1}}}}}}}}}}%
+
+% inner syntax colour
+\chardef\isachardoublequoteopen=`\"%
+\chardef\isachardoublequoteclose=`\"%
+\chardef\isacharbackquoteopen=`\`%
+\chardef\isacharbackquoteclose=`\`%
+\newenvironment{innersingle}%
+{\isacharbackquoteopen\color{isacol:green}}%
+{\color{isacol:black}\isacharbackquoteclose}
+\newenvironment{innerdouble}%
+{\isachardoublequoteopen\color{isacol:green}}%
+{\color{isacol:black}\isachardoublequoteclose}
+
+%% misc
+\newcommand{\gb}[1]{\textcolor{isacol:green}{#1}}
+\newcommand{\rb}[1]{\textcolor{red}{#1}}
+
+%% animations
+\newcounter{growcnt}
+\newcommand{\grow}[2]
+{\begin{tikzpicture}[baseline=(n.base)]%
+    \node[scale=(0.1 *#1 + 0.001),inner sep=0pt] (n) {#2};
+  \end{tikzpicture}%
+}
+
+%% isatabbing
+%\renewcommand{\isamarkupcmt}[1]%
+%{\ifthenelse{\equal{TABSET}{#1}}{\=}%
+% {\ifthenelse{\equal{TAB}{#1}}{\>}%
+%  {\ifthenelse{\equal{NEWLINE}{#1}}{\\}%
+%   {\ifthenelse{\equal{DOTS}{#1}}{\ldots}{\isastylecmt--- {#1}}}%
+%  }%
+% }%
+%}%
+
+
+
+\newenvironment{isatabbing}%
+{\renewcommand{\isanewline}{\\}\begin{tabbing}}%
+{\end{tabbing}}
+
+\begin{document}
+\nocite{*}
+\input{session}
+\end{document}
+
+%%% Local Variables:  
+%%% mode: latex
+%%% TeX-master: t
+%%% TeX-command-default: "Slides"
+%%% TeX-view-style: (("." "kghostview --landscape --scale 0.45 --geometry 605x505 %f"))
+%%% End: 
+
Binary file Slides/document/xingyuan.jpg has changed
Binary file slides1.pdf has changed