merged
authorblanchet
Wed, 21 Oct 2009 16:57:57 +0200
changeset 33055 5a733f325939
parent 33054 dd1192a96968 (current diff)
parent 33052 6f071d92960b (diff)
child 33056 791a4655cae3
merged
src/HOL/Isar_examples/Basic_Logic.thy
src/HOL/Isar_examples/Cantor.thy
src/HOL/Isar_examples/Drinker.thy
src/HOL/Isar_examples/Expr_Compiler.thy
src/HOL/Isar_examples/Fibonacci.thy
src/HOL/Isar_examples/Group.thy
src/HOL/Isar_examples/Hoare.thy
src/HOL/Isar_examples/Hoare_Ex.thy
src/HOL/Isar_examples/Knaster_Tarski.thy
src/HOL/Isar_examples/Mutilated_Checkerboard.thy
src/HOL/Isar_examples/Nested_Datatype.thy
src/HOL/Isar_examples/Peirce.thy
src/HOL/Isar_examples/Puzzle.thy
src/HOL/Isar_examples/README.html
src/HOL/Isar_examples/ROOT.ML
src/HOL/Isar_examples/Summation.thy
src/HOL/Isar_examples/document/proof.sty
src/HOL/Isar_examples/document/root.bib
src/HOL/Isar_examples/document/root.tex
src/HOL/Isar_examples/document/style.tex
src/HOL/MetisExamples/Abstraction.thy
src/HOL/MetisExamples/BT.thy
src/HOL/MetisExamples/BigO.thy
src/HOL/MetisExamples/Message.thy
src/HOL/MetisExamples/ROOT.ML
src/HOL/MetisExamples/Tarski.thy
src/HOL/MetisExamples/TransClosure.thy
src/HOL/MetisExamples/set.thy
src/HOL/SET-Protocol/Cardholder_Registration.thy
src/HOL/SET-Protocol/EventSET.thy
src/HOL/SET-Protocol/Merchant_Registration.thy
src/HOL/SET-Protocol/MessageSET.thy
src/HOL/SET-Protocol/PublicSET.thy
src/HOL/SET-Protocol/Purchase.thy
src/HOL/SET-Protocol/ROOT.ML
src/HOL/SET-Protocol/document/root.tex
src/HOL/SMT/SMT_Definitions.thy
src/HOL/Tools/record.ML
src/HOL/Tools/refute.ML
--- a/Admin/Benchmarks/HOL-datatype/IsaMakefile	Wed Oct 21 16:54:04 2009 +0200
+++ b/Admin/Benchmarks/HOL-datatype/IsaMakefile	Wed Oct 21 16:57:57 2009 +0200
@@ -20,13 +20,13 @@
 ## HOL-datatype
 
 HOL:
-	@cd $(SRC)/HOL; $(ISATOOL) make HOL
+	@cd $(SRC)/HOL; $(ISABELLE_TOOL) make HOL
 
 HOL-datatype: HOL $(LOG)/HOL-datatype.gz
 
 $(LOG)/HOL-datatype.gz: $(OUT)/HOLBrackin.thy Instructions.thy SML.thy \
   Verilog.thy
-	@cd ..; $(ISATOOL) usedir -s HOL-datatype $(OUT)/HOL HOL-datatype
+	@cd ..; $(ISABELLE_TOOL) usedir -s HOL-datatype $(OUT)/HOL HOL-datatype
 
 
 ## clean
--- a/Admin/isatest/isatest-makedist	Wed Oct 21 16:54:04 2009 +0200
+++ b/Admin/isatest/isatest-makedist	Wed Oct 21 16:57:57 2009 +0200
@@ -91,7 +91,7 @@
 
 ## spawn test runs
 
-$SSH sunbroy2 "$MAKEALL $HOME/settings/sun-poly"
+#$SSH sunbroy2 "$MAKEALL $HOME/settings/sun-poly"
 # give test some time to copy settings and start
 sleep 15
 $SSH macbroy22 "$MAKEALL $HOME/settings/at-poly"
--- a/Admin/isatest/isatest-stats	Wed Oct 21 16:54:04 2009 +0200
+++ b/Admin/isatest/isatest-stats	Wed Oct 21 16:57:57 2009 +0200
@@ -21,13 +21,13 @@
   HOL-HoareParallel \
   HOL-Lambda \
   HOL-Library \
-  HOL-MetisExamples \
+  HOL-Metis_Examples \
   HOL-MicroJava \
   HOL-NSA \
   HOL-Nominal-Examples \
   HOL-Number_Theory \
   HOL-Old_Number_Theory \
-  HOL-SET-Protocol \
+  HOL-SET_Protocol \
   HOL-UNITY \
   HOL-Word \
   HOL-ex \
--- a/CONTRIBUTORS	Wed Oct 21 16:54:04 2009 +0200
+++ b/CONTRIBUTORS	Wed Oct 21 16:57:57 2009 +0200
@@ -7,6 +7,15 @@
 Contributions to this Isabelle version
 --------------------------------------
 
+* October 2009: Sascha Boehme, TUM
+  Extension of SMT method: proof-reconstruction for the SMT solver Z3
+
+* October 2009: Florian Haftmann, TUM
+  Refinement of parts of the HOL datatype package
+
+* October 2009: Florian Haftmann, TUM
+  Generic term styles for term antiquotations
+
 * September 2009: Thomas Sewell, NICTA
   More efficient HOL/record implementation
 
@@ -14,7 +23,7 @@
   SMT method using external SMT solvers
 
 * September 2009: Florian Haftmann, TUM
-  Refinement of Sets and Lattices
+  Refinement of sets and lattices
 
 * July 2009: Jeremy Avigad and Amine Chaieb
   New number theory
--- a/NEWS	Wed Oct 21 16:54:04 2009 +0200
+++ b/NEWS	Wed Oct 21 16:57:57 2009 +0200
@@ -23,6 +23,9 @@
 to print all interpretations of locale l in the theory.  Interpretations
 in proofs are not shown.
 
+* Thoroughly revised locales tutorial.  New section on conditional
+interpretation.
+
 
 *** document preparation ***
 
@@ -43,7 +46,9 @@
 arithmetic, and fixed-size bitvectors; there is also basic
 support for higher-order features (esp. lambda abstractions).
 It is an incomplete decision procedure based on external SMT
-solvers using the oracle mechanism.
+solvers using the oracle mechanism; for the SMT solver Z3,
+this method is proof-producing. Certificates are provided to
+avoid calling the external solvers solely for re-checking proofs.
 
 * Reorganization of number theory:
   * former session NumberTheory now named Old_Number_Theory
@@ -79,8 +84,10 @@
 works well in practice on quantifier-free real arithmetic with +, -,
 *, ^, =, <= and <, i.e. boolean combinations of equalities and
 inequalities between polynomials. It makes use of external
-semidefinite programming solvers.  For more information and examples
-see src/HOL/Library/Sum_Of_Squares.
+semidefinite programming solvers. Method "sos" generates a certificate
+that can be pasted into the proof thus avoiding the need to call an external
+tool every time the proof is checked.
+For more information and examples see src/HOL/Library/Sum_Of_Squares.
 
 * Code generator attributes follow the usual underscore convention:
     code_unfold     replaces    code unfold
@@ -146,6 +153,10 @@
 this.  Fix using O_assoc[symmetric].  The same applies to the curried
 version "R OO S".
 
+* Function "Inv" is renamed to "inv_onto" and function "inv" is now an
+abbreviation for "inv_onto UNIV". Lemmas are renamed accordingly.
+INCOMPATIBILITY.
+
 * ML antiquotation @{code_datatype} inserts definition of a datatype
 generated by the code generator; see Predicate.thy for an example.
 
@@ -202,6 +213,9 @@
 
 *** ML ***
 
+* Removed some old-style infix operations using polymorphic equality.
+INCOMPATIBILITY.
+
 * Structure Synchronized (cf. src/Pure/Concurrent/synchronized.ML)
 provides a high-level programming interface to synchronized state
 variables with atomic update.  This works via pure function
@@ -250,6 +264,9 @@
 Syntax.pretty_typ/term directly, preferably with proper context
 instead of global theory.
 
+* Operations of structure Skip_Proof (formerly SkipProof) no longer
+require quick_and_dirty mode, which avoids critical setmp.
+
 
 *** System ***
 
--- a/doc-src/Intro/bool.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/Intro/bool.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -1,5 +1,5 @@
 Bool = FOL +
-types 	bool 0
-arities bool 	:: term
-consts tt,ff	:: "bool"
+types   bool 0
+arities bool    :: term
+consts tt,ff    :: "bool"
 end
--- a/doc-src/Intro/list.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/Intro/list.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -1,6 +1,6 @@
 List = FOL +
-types 	list 1
-arities list	:: (term)term
-consts	Nil	:: "'a list"
-   	Cons	:: "['a, 'a list] => 'a list" 
+types   list 1
+arities list    :: (term)term
+consts  Nil     :: "'a list"
+        Cons    :: "['a, 'a list] => 'a list" 
 end
--- a/doc-src/IsarImplementation/Thy/ML.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/IsarImplementation/Thy/ML.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -229,7 +229,7 @@
   view being presented to the user.
 
   Occasionally, such global process flags are treated like implicit
-  arguments to certain operations, by using the @{ML setmp} combinator
+  arguments to certain operations, by using the @{ML setmp_CRITICAL} combinator
   for safe temporary assignment.  Its traditional purpose was to
   ensure proper recovery of the original value when exceptions are
   raised in the body, now the functionality is extended to enter the
@@ -237,7 +237,7 @@
   parallelism).
 
   Note that recovery of plain value passing semantics via @{ML
-  setmp}~@{text "ref value"} assumes that this @{text "ref"} is
+  setmp_CRITICAL}~@{text "ref value"} assumes that this @{text "ref"} is
   exclusively manipulated within the critical section.  In particular,
   any persistent global assignment of @{text "ref := value"} needs to
   be marked critical as well, to prevent intruding another threads
@@ -258,7 +258,7 @@
   \begin{mldecls}
   @{index_ML NAMED_CRITICAL: "string -> (unit -> 'a) -> 'a"} \\
   @{index_ML CRITICAL: "(unit -> 'a) -> 'a"} \\
-  @{index_ML setmp: "'a Unsynchronized.ref -> 'a -> ('b -> 'c) -> 'b -> 'c"} \\
+  @{index_ML setmp_CRITICAL: "'a Unsynchronized.ref -> 'a -> ('b -> 'c) -> 'b -> 'c"} \\
   \end{mldecls}
 
   \begin{description}
@@ -272,7 +272,7 @@
   \item @{ML CRITICAL} is the same as @{ML NAMED_CRITICAL} with empty
   name argument.
 
-  \item @{ML setmp}~@{text "ref value f x"} evaluates @{text "f x"}
+  \item @{ML setmp_CRITICAL}~@{text "ref value f x"} evaluates @{text "f x"}
   while staying within the critical section and having @{text "ref :=
   value"} assigned temporarily.  This recovers a value-passing
   semantics involving global references, regardless of exceptions or
--- a/doc-src/Locales/Locales/Examples.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/Locales/Locales/Examples.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -1,5 +1,5 @@
 theory Examples
-imports Main GCD
+imports Main
 begin
 
 hide %invisible const Lattices.lattice
@@ -12,9 +12,9 @@
   primitives are universal quantification (@{text "\<And>"}), entailment
   (@{text "\<Longrightarrow>"}) and equality (@{text "\<equiv>"}).  Variables (not bound
   variables) are sometimes preceded by a question mark.  The logic is
-  typed.  Type variables are denoted by @{text "'a"}, @{text "'b"}
-  etc., and @{text "\<Rightarrow>"} is the function type.  Double brackets @{text
-  "\<lbrakk>"} and @{text "\<rbrakk>"} are used to abbreviate nested entailment.
+  typed.  Type variables are denoted by~@{text "'a"},~@{text "'b"}
+  etc., and~@{text "\<Rightarrow>"} is the function type.  Double brackets~@{text
+  "\<lbrakk>"} and~@{text "\<rbrakk>"} are used to abbreviate nested entailment.
 *}
 *)
 
@@ -26,25 +26,27 @@
 \[
   @{text "\<And>x\<^sub>1\<dots>x\<^sub>n. \<lbrakk> A\<^sub>1; \<dots> ;A\<^sub>m \<rbrakk> \<Longrightarrow> \<dots>"}
 \]
-  where variables @{text "x\<^sub>1"}, \ldots, @{text "x\<^sub>n"} are called
-  \emph{parameters} and the premises $@{text "A\<^sub>1"}, \ldots,
-  @{text "A\<^sub>m"}$ \emph{assumptions}.  A formula @{text "C"}
+  where variables~@{text "x\<^sub>1"}, \ldots,~@{text "x\<^sub>n"} are called
+  \emph{parameters} and the premises $@{text "A\<^sub>1"}, \ldots,~@{text
+  "A\<^sub>m"}$ \emph{assumptions}.  A formula~@{text "C"}
   is a \emph{theorem} in the context if it is a conclusion
 \[
-%\label{eq-fact-in-context}
   @{text "\<And>x\<^sub>1\<dots>x\<^sub>n. \<lbrakk> A\<^sub>1; \<dots> ;A\<^sub>m \<rbrakk> \<Longrightarrow> C"}.
 \]
   Isabelle/Isar's notion of context goes beyond this logical view.
   Its contexts record, in a consecutive order, proved
-  conclusions along with attributes, which
-  may control proof procedures.  Contexts also contain syntax information
-  for parameters and for terms depending on them.
+  conclusions along with \emph{attributes}, which can provide context
+  specific configuration information for proof procedures and concrete
+  syntax.  From a logical perspective, locales are just contexts that
+  have been made persistent.  To the user, though, they provide
+  powerful means for declaring and combining contexts, and for the
+  reuse of theorems proved in these contexts.
   *}
 
 section {* Simple Locales *}
 
 text {*
-  Locales can be seen as persistent contexts.  In its simplest form, a
+  In its simplest form, a
   \emph{locale declaration} consists of a sequence of context elements
   declaring parameters (keyword \isakeyword{fixes}) and assumptions
   (keyword \isakeyword{assumes}).  The following is the specification of
@@ -57,22 +59,67 @@
       and anti_sym [intro]: "\<lbrakk> x \<sqsubseteq> y; y \<sqsubseteq> x \<rbrakk> \<Longrightarrow> x = y"
       and trans [trans]: "\<lbrakk> x \<sqsubseteq> y; y \<sqsubseteq> z \<rbrakk> \<Longrightarrow> x \<sqsubseteq> z"
 
-text {* The parameter of this locale is @{term le}, with infix syntax
-  @{text \<sqsubseteq>}.  There is an implicit type parameter @{typ "'a"}.  It
-  is not necessary to declare parameter types: most general types will
-  be inferred from the context elements for all parameters.
+text (in partial_order) {* The parameter of this locale is~@{text le},
+  which is a binary predicate with infix syntax~@{text \<sqsubseteq>}.  The
+  parameter syntax is available in the subsequent
+  assumptions, which are the familiar partial order axioms.
+
+  Isabelle recognises unbound names as free variables.  In locale
+  assumptions, these are implicitly universally quantified.  That is,
+  @{term "\<lbrakk> x \<sqsubseteq> y; y \<sqsubseteq> z \<rbrakk> \<Longrightarrow> x \<sqsubseteq> z"} in fact means
+  @{term "\<And>x y z. \<lbrakk> x \<sqsubseteq> y; y \<sqsubseteq> z \<rbrakk> \<Longrightarrow> x \<sqsubseteq> z"}.
 
-  The above declaration not only introduces the locale, it also
-  defines the \emph{locale predicate} @{term partial_order} with
-  definition @{thm [source] partial_order_def}:
+  Two commands are provided to inspect locales:
+  \isakeyword{print\_locales} lists the names of all locales of the
+  current theory; \isakeyword{print\_locale}~$n$ prints the parameters
+  and assumptions of locale $n$; the variation \isakeyword{print\_locale!}~$n$
+  additionally outputs the conclusions that are stored in the locale.
+  We may inspect the new locale
+  by issuing \isakeyword{print\_locale!} @{term partial_order}.  The output
+  is the following list of context elements.
+\begin{small}
+\begin{alltt}
+  \isakeyword{fixes} le :: "'a \(\Rightarrow\) 'a \(\Rightarrow\)  bool" (\isakeyword{infixl} "\(\sqsubseteq\)" 50)
+  \isakeyword{assumes} "partial_order op \(\sqsubseteq\)"
+  \isakeyword{notes} assumption
+    refl [intro, simp] = `?x \(\sqsubseteq\) ?x`
+    \isakeyword{and}
+    anti_sym [intro] = `\(\isasymlbrakk\)?x \(\sqsubseteq\) ?y; ?y \(\sqsubseteq\) ?x\(\isasymrbrakk\) \(\Longrightarrow\) ?x = ?y`
+    \isakeyword{and}
+    trans [trans] = `\(\isasymlbrakk\)?x \(\sqsubseteq\) ?y; ?y \(\sqsubseteq\) ?z\(\isasymrbrakk\) \(\Longrightarrow\) ?x \(\sqsubseteq\) ?z`
+\end{alltt}
+\end{small}
+  The keyword \isakeyword{notes} denotes a conclusion element.  There
+  is one conclusion, which was added automatically.  Instead, there is
+  only one assumption, namely @{term "partial_order le"}.  The locale
+  declaration has introduced the predicate @{term
+  partial_order} to the theory.  This predicate is the
+  \emph{locale predicate}.  Its definition may be inspected by
+  issuing \isakeyword{thm} @{thm [source] partial_order_def}.
   @{thm [display, indent=2] partial_order_def}
+  In our example, this is a unary predicate over the parameter of the
+  locale.  It is equivalent to the original assumptions, which have
+  been turned into conclusions and are
+  available as theorems in the context of the locale.  The names and
+  attributes from the locale declaration are associated to these
+  theorems and are effective in the context of the locale.
 
+  Each conclusion has a \emph{foundational theorem} as counterpart
+  in the theory.  Technically, this is simply the theorem composed
+  of context and conclusion.  For the transitivity theorem, this is
+  @{thm [source] partial_order.trans}:
+  @{thm [display, indent=2] partial_order_def}
+*}
+
+subsection {* Targets: Extending Locales *}
+
+text {*
   The specification of a locale is fixed, but its list of conclusions
   may be extended through Isar commands that take a \emph{target} argument.
   In the following, \isakeyword{definition} and 
   \isakeyword{theorem} are illustrated.
   Table~\ref{tab:commands-with-target} lists Isar commands that accept
-  a target.  There are various ways of specifying the target.  A
+  a target.  Isar provides various ways of specifying the target.  A
   target for a single command may be indicated with keyword
   \isakeyword{in} in the following way:
 
@@ -101,33 +148,39 @@
     less :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "\<sqsubset>" 50)
     where "(x \<sqsubset> y) = (x \<sqsubseteq> y \<and> x \<noteq> y)"
 
-text {* A definition in a locale depends on the locale parameters.
-  Here, a global constant @{term partial_order.less} is declared, which is lifted over the
-  locale parameter @{term le}.  Its definition is the global theorem
-  @{thm [source] partial_order.less_def}:
+text (in partial_order) {* The strict order @{text less} with infix
+  syntax~@{text \<sqsubset>} is
+  defined in terms of the locale parameter~@{text le} and the general
+  equality of the object logic we work in.  The definition generates a
+  \emph{foundational constant}
+  @{term partial_order.less} with definition @{thm [source]
+  partial_order.less_def}:
   @{thm [display, indent=2] partial_order.less_def}
   At the same time, the locale is extended by syntax transformations
-  hiding this construction in the context of the locale.  That is,
-  @{term "partial_order.less le"} is printed and parsed as infix
-  @{text \<sqsubset>}. *}
+  hiding this construction in the context of the locale.  Here, the
+  abbreviation @{text less} is available for
+  @{text "partial_order.less le"}, and it is printed
+  and parsed as infix~@{text \<sqsubset>}.  Finally, the conclusion @{thm [source]
+  less_def} is added to the locale:
+  @{thm [display, indent=2] less_def}
+*}
 
-text (in partial_order) {* Finally, the conclusion of the definition
-  is added to the locale, @{thm [source] less_def}:
-  @{thm [display, indent=2] less_def}
-  *}
-
-text {* As an example of a theorem statement in the locale, here is the
-  derivation of a transitivity law. *}
+text {* The treatment of theorem statements is more straightforward.
+  As an example, here is the derivation of a transitivity law for the
+  strict order relation. *}
 
   lemma (in partial_order) less_le_trans [trans]:
     "\<lbrakk> x \<sqsubset> y; y \<sqsubseteq> z \<rbrakk> \<Longrightarrow> x \<sqsubset> z"
     unfolding %visible less_def by %visible (blast intro: trans)
 
-text {* In the context of the proof, assumptions and theorems of the
-  locale may be used.  Attributes are effective: @{text anti_sym} was
+text {* In the context of the proof, conclusions of the
+  locale may be used like theorems.  Attributes are effective: @{text
+  anti_sym} was
   declared as introduction rule, hence it is in the context's set of
   rules used by the classical reasoner by default.  *}
 
+subsection {* Context Blocks *}
+
 text {* When working with locales, sequences of commands with the same
   target are frequent.  A block of commands, delimited by
   \isakeyword{begin} and \isakeyword{end}, makes a theory-like style
@@ -139,7 +192,7 @@
 
   This style of working is illustrated in the block below, where
   notions of infimum and supremum for partial orders are introduced,
-  together with theorems.  *}
+  together with theorems about their uniqueness.  *}
 
   context partial_order begin
 
@@ -164,20 +217,20 @@
     by (unfold is_inf_def) blast
 
   theorem is_inf_uniq: "\<lbrakk>is_inf x y i; is_inf x y i'\<rbrakk> \<Longrightarrow> i = i'"
-  proof -
+    proof -
     assume inf: "is_inf x y i"
     assume inf': "is_inf x y i'"
     show ?thesis
     proof (rule anti_sym)
       from inf' show "i \<sqsubseteq> i'"
       proof (rule is_inf_greatest)
-	from inf show "i \<sqsubseteq> x" ..
-	from inf show "i \<sqsubseteq> y" ..
+        from inf show "i \<sqsubseteq> x" ..
+        from inf show "i \<sqsubseteq> y" ..
       qed
       from inf show "i' \<sqsubseteq> i"
       proof (rule is_inf_greatest)
-	from inf' show "i' \<sqsubseteq> x" ..
-	from inf' show "i' \<sqsubseteq> y" ..
+        from inf' show "i' \<sqsubseteq> x" ..
+        from inf' show "i' \<sqsubseteq> y" ..
       qed
     qed
   qed
@@ -206,20 +259,20 @@
     by (unfold is_sup_def) blast
 
   theorem is_sup_uniq: "\<lbrakk>is_sup x y s; is_sup x y s'\<rbrakk> \<Longrightarrow> s = s'"
-  proof -
+    proof -
     assume sup: "is_sup x y s"
     assume sup': "is_sup x y s'"
     show ?thesis
     proof (rule anti_sym)
       from sup show "s \<sqsubseteq> s'"
       proof (rule is_sup_least)
-	from sup' show "x \<sqsubseteq> s'" ..
-	from sup' show "y \<sqsubseteq> s'" ..
+        from sup' show "x \<sqsubseteq> s'" ..
+        from sup' show "y \<sqsubseteq> s'" ..
       qed
       from sup' show "s' \<sqsubseteq> s"
       proof (rule is_sup_least)
-	from sup show "x \<sqsubseteq> s" ..
-	from sup show "y \<sqsubseteq> s" ..
+        from sup show "x \<sqsubseteq> s" ..
+        from sup show "y \<sqsubseteq> s" ..
       qed
     qed
   qed
@@ -238,17 +291,12 @@
 
   end
 
-text {* Two commands are provided to inspect locales:
-  \isakeyword{print\_locales} lists the names of all locales of the
-  current theory; \isakeyword{print\_locale}~$n$ prints the parameters
-  and assumptions of locale $n$; \isakeyword{print\_locale!}~$n$
-  additionally outputs the conclusions.
-
-  The syntax of the locale commands discussed in this tutorial is
-  shown in Table~\ref{tab:commands}.  The grammer is complete with the
-  exception of additional context elements not discussed here.  See the
-  Isabelle/Isar Reference Manual~\cite{IsarRef}
-  for full documentation.  *}
+text {* The syntax of the locale commands discussed in this tutorial is
+  shown in Table~\ref{tab:commands}.  The grammar is complete with the
+  exception of the context elements \isakeyword{constrains} and
+  \isakeyword{defines}, which are provided for backward
+  compatibility.  See the Isabelle/Isar Reference
+  Manual~\cite{IsarRef} for full documentation.  *}
 
 
 section {* Import \label{sec:import} *}
@@ -257,11 +305,13 @@
   Algebraic structures are commonly defined by adding operations and
   properties to existing structures.  For example, partial orders
   are extended to lattices and total orders.  Lattices are extended to
-  distributive lattices.
+  distributive lattices. *}
 
-  With locales, this inheritance is achieved through \emph{import} of a
-  locale.  Import is a separate entity in the locale declaration.  If
-  present, it precedes the context elements.
+text {*
+  With locales, this kind of inheritance is achieved through
+  \emph{import} of locales.  The import part of a locale declaration,
+  if present, precedes the context elements.  Here is an example,
+  where partial orders are extended to lattices.
   *}
 
   locale lattice = partial_order +
@@ -270,12 +320,11 @@
   begin
 
 text {* These assumptions refer to the predicates for infimum
-  and supremum defined in @{text partial_order}.  We may now introduce
-  the notions of meet and join.  *}
+  and supremum defined for @{text partial_order} in the previous
+  section.  We now introduce the notions of meet and join.  *}
 
   definition
     meet (infixl "\<sqinter>" 70) where "x \<sqinter> y = (THE inf. is_inf x y inf)"
-
   definition
     join (infixl "\<squnion>" 65) where "x \<squnion> y = (THE sup. is_sup x y sup)"
 
@@ -346,9 +395,9 @@
       show "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> x" ..
       show "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> y"
       proof -
-	have "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> y \<sqinter> z" ..
-	also have "\<dots> \<sqsubseteq> y" ..
-	finally show ?thesis .
+        have "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> y \<sqinter> z" ..
+        also have "\<dots> \<sqsubseteq> y" ..
+        finally show ?thesis .
       qed
     qed
     show "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> z"
@@ -362,19 +411,19 @@
     proof
       show "w \<sqsubseteq> x"
       proof -
-	have "w \<sqsubseteq> x \<sqinter> y" by fact
-	also have "\<dots> \<sqsubseteq> x" ..
-	finally show ?thesis .
+        have "w \<sqsubseteq> x \<sqinter> y" by fact
+        also have "\<dots> \<sqsubseteq> x" ..
+        finally show ?thesis .
       qed
       show "w \<sqsubseteq> y \<sqinter> z"
       proof
-	show "w \<sqsubseteq> y"
-	proof -
-	  have "w \<sqsubseteq> x \<sqinter> y" by fact
-	  also have "\<dots> \<sqsubseteq> y" ..
-	  finally show ?thesis .
-	qed
-	show "w \<sqsubseteq> z" by fact
+        show "w \<sqsubseteq> y"
+        proof -
+          have "w \<sqsubseteq> x \<sqinter> y" by fact
+          also have "\<dots> \<sqsubseteq> y" ..
+          finally show ?thesis .
+        qed
+        show "w \<sqsubseteq> z" by fact
       qed
     qed
   qed
@@ -402,9 +451,9 @@
       show "x \<sqsubseteq> x \<squnion> (y \<squnion> z)" ..
       show "y \<sqsubseteq> x \<squnion> (y \<squnion> z)"
       proof -
-	have "y \<sqsubseteq> y \<squnion> z" ..
-	also have "... \<sqsubseteq> x \<squnion> (y \<squnion> z)" ..
-	finally show ?thesis .
+        have "y \<sqsubseteq> y \<squnion> z" ..
+        also have "... \<sqsubseteq> x \<squnion> (y \<squnion> z)" ..
+        finally show ?thesis .
       qed
     qed
     show "z \<sqsubseteq> x \<squnion> (y \<squnion> z)"
@@ -418,19 +467,19 @@
     proof
       show "x \<sqsubseteq> w"
       proof -
-	have "x \<sqsubseteq> x \<squnion> y" ..
-	also have "\<dots> \<sqsubseteq> w" by fact
-	finally show ?thesis .
+        have "x \<sqsubseteq> x \<squnion> y" ..
+        also have "\<dots> \<sqsubseteq> w" by fact
+        finally show ?thesis .
       qed
       show "y \<squnion> z \<sqsubseteq> w"
       proof
-	show "y \<sqsubseteq> w"
-	proof -
-	  have "y \<sqsubseteq> x \<squnion> y" ..
-	  also have "... \<sqsubseteq> w" by fact
-	  finally show ?thesis .
-	qed
-	show "z \<sqsubseteq> w" by fact
+        show "y \<sqsubseteq> w"
+        proof -
+          have "y \<sqsubseteq> x \<squnion> y" ..
+          also have "... \<sqsubseteq> w" by fact
+          finally show ?thesis .
+        qed
+        show "z \<sqsubseteq> w" by fact
       qed
     qed
   qed
@@ -518,8 +567,10 @@
 
   end
 
-text {* Locales for total orders and distributive lattices follow.
-  Each comes with an example theorem. *}
+text {* Locales for total orders and distributive lattices follow to
+  establish a sufficiently rich landscape of locales for
+  further examples in this tutorial.  Each comes with an example
+  theorem. *}
 
   locale total_order = partial_order +
     assumes total: "x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
@@ -543,12 +594,13 @@
   qed
 
 text {*
-  The locale hierachy obtained through these declarations is shown in Figure~\ref{fig:lattices}(a).
+  The locale hierarchy obtained through these declarations is shown in
+  Figure~\ref{fig:lattices}(a).
 
 \begin{figure}
 \hrule \vspace{2ex}
 \begin{center}
-\subfigure[Declared hierachy]{
+\subfigure[Declared hierarchy]{
 \begin{tikzpicture}
   \node (po) at (0,0) {@{text partial_order}};
   \node (lat) at (-1.5,-1) {@{text lattice}};
@@ -594,21 +646,50 @@
   \label{sec:changing-the-hierarchy} *}
 
 text {*
-  Total orders are lattices.  Hence, by deriving the lattice
-  axioms for total orders, the hierarchy may be changed
-  and @{text lattice} be placed between @{text partial_order}
-  and @{text total_order}, as shown in Figure~\ref{fig:lattices}(b).
-  Changes to the locale hierarchy may be declared
-  with the \isakeyword{sublocale} command. *}
+  Locales enable to prove theorems abstractly, relative to
+  sets of assumptions.  These theorems can then be used in other
+  contexts where the assumptions themselves, or
+  instances of the assumptions, are theorems.  This form of theorem
+  reuse is called \emph{interpretation}.  Locales generalise
+  interpretation from theorems to conclusions, enabling the reuse of
+  definitions and other constructs that are not part of the
+  specifications of the locales.
+
+  The first from of interpretation we will consider in this tutorial
+  is provided by the \isakeyword{sublocale} command.  It enables to
+  modify the import hierarchy to reflect the \emph{logical} relation
+  between locales.
+
+  Consider the locale hierarchy from Figure~\ref{fig:lattices}(a).
+  Total orders are lattices, although this is not reflected here, and
+  definitions, theorems and other conclusions
+  from @{term lattice} are not available in @{term total_order}.  To
+  obtain the situation in Figure~\ref{fig:lattices}(b), it is
+  sufficient to add the conclusions of the latter locale to the former.
+  The \isakeyword{sublocale} command does exactly this.
+  The declaration \isakeyword{sublocale} $l_1
+  \subseteq l_2$ causes locale $l_2$ to be \emph{interpreted} in the
+  context of $l_1$.  This means that all conclusions of $l_2$ are made
+  available in $l_1$.
+
+  Of course, the change of hierarchy must be supported by a theorem
+  that reflects, in our example, that total orders are indeed
+  lattices.  Therefore the \isakeyword{sublocale} command generates a
+  goal, which must be discharged by the user.  This is illustrated in
+  the following paragraphs.  First the sublocale relation is stated.
+*}
 
   sublocale %visible total_order \<subseteq> lattice
 
-txt {* This enters the context of locale @{text total_order}, in
-  which the goal @{subgoals [display]} must be shown.  First, the
-  locale predicate needs to be unfolded --- for example using its
+txt {* \normalsize
+  This enters the context of locale @{text total_order}, in
+  which the goal @{subgoals [display]} must be shown.
+  Now the
+  locale predicate needs to be unfolded --- for example, using its
   definition or by introduction rules
-  provided by the locale package.  The methods @{text intro_locales}
-  and @{text unfold_locales} automate this.  They are aware of the
+  provided by the locale package.  For automation, the locale package
+  provides the methods @{text intro_locales} and @{text
+  unfold_locales}.  They are aware of the
   current context and dependencies between locales and automatically
   discharge goals implied by these.  While @{text unfold_locales}
   always unfolds locale predicates to assumptions, @{text
@@ -618,22 +699,26 @@
   is smaller.
 
   For the current goal, we would like to get hold of
-  the assumptions of @{text lattice}, hence @{text unfold_locales}
-  is appropriate. *}
+  the assumptions of @{text lattice}, which need to be shown, hence
+  @{text unfold_locales} is appropriate. *}
 
   proof unfold_locales
 
-txt {* Since both @{text lattice} and @{text total_order}
-  inherit @{text partial_order}, the assumptions of the latter are
-  discharged, and the only subgoals that remain are the assumptions
-  introduced in @{text lattice} @{subgoals [display]}
-  The proof for the first subgoal is *}
+txt {* \normalsize
+  Since the fact that both lattices and total orders are partial
+  orders is already reflected in the locale hierarchy, the assumptions
+  of @{text partial_order} are discharged automatically, and only the
+  assumptions introduced in @{text lattice} remain as subgoals
+  @{subgoals [display]}
+  The proof for the first subgoal is obtained by constructing an
+  infimum, whose existence is implied by totality. *}
 
     fix x y
     from total have "is_inf x y (if x \<sqsubseteq> y then x else y)"
       by (auto simp: is_inf_def)
     then show "\<exists>inf. is_inf x y inf" ..
-txt {* The proof for the second subgoal is analogous and not
+txt {* \normalsize
+   The proof for the second subgoal is analogous and not
   reproduced here. *}
   next %invisible
     fix x y
@@ -641,31 +726,45 @@
       by (auto simp: is_sup_def)
     then show "\<exists>sup. is_sup x y sup" .. qed %visible
 
-text {* Similarly, total orders are distributive lattices. *}
+text {* Similarly, we may establish that total orders are distributive
+  lattices with a second \isakeyword{sublocale} statement. *}
 
   sublocale total_order \<subseteq> distrib_lattice
-  proof unfold_locales
+    proof unfold_locales
     fix %"proof" x y z
     show "x \<sqinter> (y \<squnion> z) = x \<sqinter> y \<squnion> x \<sqinter> z" (is "?l = ?r")
       txt {* Jacobson I, p.\ 462 *}
     proof -
       { assume c: "y \<sqsubseteq> x" "z \<sqsubseteq> x"
-	from c have "?l = y \<squnion> z"
-	  by (metis c join_connection2 join_related2 meet_related2 total)
-	also from c have "... = ?r" by (metis meet_related2)
-	finally have "?l = ?r" . }
+        from c have "?l = y \<squnion> z"
+          by (metis c join_connection2 join_related2 meet_related2 total)
+        also from c have "... = ?r" by (metis meet_related2)
+        finally have "?l = ?r" . }
       moreover
       { assume c: "x \<sqsubseteq> y \<or> x \<sqsubseteq> z"
-	from c have "?l = x"
-	  by (metis join_connection2 join_related2 meet_connection total trans)
-	also from c have "... = ?r"
-	  by (metis join_commute join_related2 meet_connection meet_related2 total)
-	finally have "?l = ?r" . }
+        from c have "?l = x"
+          by (metis join_connection2 join_related2 meet_connection total trans)
+        also from c have "... = ?r"
+          by (metis join_commute join_related2 meet_connection meet_related2 total)
+        finally have "?l = ?r" . }
       moreover note total
       ultimately show ?thesis by blast
     qed
   qed
 
-text {* The locale hierarchy is now as shown in Figure~\ref{fig:lattices}(c). *}
+text {* The locale hierarchy is now as shown in
+  Figure~\ref{fig:lattices}(c). *}
+
+text {*
+  Locale interpretation is \emph{dynamic}.  The statement
+  \isakeyword{sublocale} $l_1 \subseteq l_2$ will not just add the
+  current conclusions of $l_2$ to $l_1$.  Rather the dependency is
+  stored, and conclusions that will be
+  added to $l_2$ in future are automatically propagated to $l_1$.
+  The sublocale relation is transitive --- that is, propagation takes
+  effect along chains of sublocales.  Even cycles in the sublocale relation are
+  supported, as long as these cycles do not lead to infinite chains.
+  Details are discussed in the technical report \cite{Ballarin2006a}.
+  See also Section~\ref{sec:infinite-chains} of this tutorial.  *}
 
 end
--- a/doc-src/Locales/Locales/Examples1.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/Locales/Locales/Examples1.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -1,36 +1,28 @@
 theory Examples1
 imports Examples
 begin
-
-section {* Use of Locales in Theories and Proofs *}
-
-text {* Locales enable to prove theorems abstractly, relative to
-  sets of assumptions.  These theorems can then be used in other
-  contexts where the assumptions themselves, or
-  instances of the assumptions, are theorems.  This form of theorem
-  reuse is called \emph{interpretation}.
+text {* \vspace{-5ex} *}
+section {* Use of Locales in Theories and Proofs
+  \label{sec:interpretation} *}
 
-  The changes of the locale
-  hierarchy from the previous sections are examples of
-  interpretations.  The command \isakeyword{sublocale} $l_1
-  \subseteq l_2$ is said to \emph{interpret} locale $l_2$ in the
-  context of $l_1$.  It causes all theorems of $l_2$ to be made
-  available in $l_1$.  The interpretation is \emph{dynamic}: not only
-  theorems already present in $l_2$ are available in $l_1$.  Theorems
-  that will be added to $l_2$ in future will automatically be
-  propagated to $l_1$.
+text {*
+  Locales can be interpreted in the contexts of theories and
+  structured proofs.  These interpretations are dynamic, too.
+  Conclusions of locales will be propagated to the current theory or
+  the current proof context.%
+\footnote{Strictly speaking, only interpretation in theories is
+  dynamic since it is not possible to change locales or the locale
+  hierarchy from within a proof.}
+  The focus of this section is on
+  interpretation in theories, but we will also encounter
+  interpretations in proofs, in
+  Section~\ref{sec:local-interpretation}.
 
-  Locales can also be interpreted in the contexts of theories and
-  structured proofs.  These interpretations are dynamic, too.
-  Theorems added to locales will be propagated to theories.
-  In this section the interpretation in
-  theories is illustrated; interpretation in proofs is analogous.
-
-  As an example, consider the type of natural numbers @{typ nat}.  The
-  relation @{text \<le>} is a total order over @{typ nat},
-  divisibility @{text dvd} is a distributive lattice.  We start with the
-  interpretation that @{text \<le>} is a partial order.  The facilities of
-  the interpretation command are explored in three versions.
+  As an example, consider the type of integers @{typ int}.  The
+  relation @{term "op \<le>"} is a total order over @{typ int}.  We start
+  with the interpretation that @{term "op \<le>"} is a partial order.  The
+  facilities of the interpretation command are explored gradually in
+  three versions.
   *}
 
 
@@ -38,45 +30,60 @@
   \label{sec:po-first} *}
 
 text {*
-  In the most basic form, interpretation just replaces the locale
-  parameters by terms.  The following command interprets the locale
-  @{text partial_order} in the global context of the theory.  The
-  parameter @{term le} is replaced by @{term "op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool"}. *} 
+  The command \isakeyword{interpretation} is for the interpretation of
+  locale in theories.  In the following example, the parameter of locale
+  @{text partial_order} is replaced by @{term "op \<le> :: int \<Rightarrow> int \<Rightarrow>
+  bool"} and the locale instance is interpreted in the current
+  theory. *}
 
-  interpretation %visible nat: partial_order "op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool"
-txt {* The locale name is succeeded by a \emph{parameter
-  instantiation}.  This is a list of terms, which refer to
-  the parameters in the order of declaration in the locale.  The
-  locale name is preceded by an optional \emph{interpretation
-  qualifier}, here @{text nat}.
+  interpretation %visible int: partial_order "op \<le> :: int \<Rightarrow> int \<Rightarrow> bool"
+txt {* \normalsize
+  The argument of the command is a simple \emph{locale expression}
+  consisting of the name of the interpreted locale, which is
+  preceded by the qualifier @{text "int:"} and succeeded by a
+  white-space-separated list of terms, which provide a full
+  instantiation of the locale parameters.  The parameters are referred
+  to by order of declaration, which is also the order in which
+  \isakeyword{print\_locale} outputs them.  The locale has only a
+  single parameter, hence the list of instantiation terms is a
+  singleton.
 
-  The command creates the goal%
-\footnote{Note that @{text op} binds tighter than functions
-  application: parentheses around @{text "op \<le>"} are not necessary.}
+  The command creates the goal
   @{subgoals [display]} which can be shown easily:
  *}
     by unfold_locales auto
 
-text {*  Now theorems from the locale are available in the theory,
-  interpreted for natural numbers, for example @{thm [source]
-  nat.trans}: @{thm [display, indent=2] nat.trans}
-
-  The interpretation qualifier, @{text nat} in the example, is applied
-  to all names processed by the interpretation.  If a qualifer is
-  given in the \isakeyword{interpretation} command, its use is
-  mandatory when referencing the name.  For example, the above theorem
-  cannot be referred to simply by @{text trans}.  This prevents
-  unwanted hiding of theorems. *}
+text {*  The effect of the command is that instances of all
+  conclusions of the locale are available in the theory, where names
+  are prefixed by the qualifier.  For example, transitivity for @{typ
+  int} is named @{thm [source] int.trans} and is the following
+  theorem:
+  @{thm [display, indent=2] int.trans}
+  It is not possible to reference this theorem simply as @{text
+  trans}.  This prevents unwanted hiding of existing theorems of the
+  theory by an interpretation. *}
 
 
 subsection {* Second Version: Replacement of Definitions *}
 
-text {* The above interpretation also creates the theorem
-  @{thm [source] nat.less_le_trans}: @{thm [display, indent=2]
-  nat.less_le_trans}
-  Here, @{term "partial_order.less (op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool)"}
-  represents the strict order, although @{text "<"} is the natural
-  strict order for @{typ nat}.  Interpretation allows to map concepts
-  introduced by definitions in locales to the corresponding
-  concepts of the theory.  *}
+text {* Not only does the above interpretation qualify theorem names.
+  The prefix @{text int} is applied to all names introduced in locale
+  conclusions including names introduced in definitions.  The
+  qualified name @{text int.less} is short for
+  the interpretation of the definition, which is @{term int.less}.
+  Qualified name and expanded form may be used almost
+  interchangeably.%
+\footnote{Since @{term "op \<le>"} is polymorphic, for @{term int.less} a
+  more general type will be inferred than for @{text int.less} which
+  is over type @{typ int}.}
+  The latter is preferred on output, as for example in the theorem
+  @{thm [source] int.less_le_trans}: @{thm [display, indent=2]
+  int.less_le_trans}
+  Both notations for the strict order are not satisfactory.  The
+  constant @{term "op <"} is the strict order for @{typ int}.
+  In order to allow for the desired replacement, interpretation
+  accepts \emph{equations} in addition to the parameter instantiation.
+  These follow the locale expression and are indicated with the
+  keyword \isakeyword{where}.  This is the revised interpretation:
+  *}
 end
--- a/doc-src/Locales/Locales/Examples2.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/Locales/Locales/Examples2.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -1,27 +1,24 @@
 theory Examples2
 imports Examples
 begin
-text {* This is achieved by unfolding suitable equations during
-  interpretation.  These equations are given after the keyword
-  \isakeyword{where} and require proofs.  The revised command
-  that replaces @{text "\<sqsubset>"} by @{text "<"} is: *}
+text {* \vspace{-5ex} *}
+  interpretation %visible int: partial_order "op \<le> :: [int, int] \<Rightarrow> bool"
+    where "partial_order.less op \<le> (x::int) y = (x < y)"
+  proof -
+    txt {* \normalsize The goals are now:
+      @{subgoals [display]}
+      The proof that~@{text \<le>} is a partial order is as above. *}
+    show "partial_order (op \<le> :: int \<Rightarrow> int \<Rightarrow> bool)"
+      by unfold_locales auto
+    txt {* \normalsize The second goal is shown by unfolding the
+      definition of @{term "partial_order.less"}. *}
+    show "partial_order.less op \<le> (x::int) y = (x < y)"
+      unfolding partial_order.less_def [OF `partial_order op \<le>`]
+      by auto
+  qed
 
-interpretation %visible nat: partial_order "op \<le> :: [nat, nat] \<Rightarrow> bool"
-  where "partial_order.less op \<le> (x::nat) y = (x < y)"
-proof -
-  txt {* The goals are @{subgoals [display]}
-    The proof that @{text \<le>} is a partial order is as above. *}
-  show "partial_order (op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool)"
-    by unfold_locales auto
-  txt {* The second goal is shown by unfolding the
-    definition of @{term "partial_order.less"}. *}
-  show "partial_order.less op \<le> (x::nat) y = (x < y)"
-    unfolding partial_order.less_def [OF `partial_order op \<le>`]
-    by auto
-qed
-
-text {* Note that the above proof is not in the context of a locale.
-  Hence, the correct interpretation of @{text
-  "partial_order.less_def"} is obtained manually with @{text OF}.
+text {* Note that the above proof is not in the context of the
+  interpreted locale.  Hence, the premise of @{text
+  "partial_order.less_def"} is discharged manually with @{text OF}.
   *}
 end
--- a/doc-src/Locales/Locales/Examples3.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/Locales/Locales/Examples3.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -1,309 +1,234 @@
 theory Examples3
 imports Examples
 begin
-subsection {* Third Version: Local Interpretation *}
+text {* \vspace{-5ex} *}
+subsection {* Third Version: Local Interpretation
+  \label{sec:local-interpretation} *}
 
-text {* In the above example, the fact that @{text \<le>} is a partial
-  order for the natural numbers was used in the proof of the
-  second goal.  In general, proofs of the equations may involve
-  theorems implied by the fact the assumptions of the instantiated
-  locale hold for the instantiating structure.  If these theorems have
-  been shown abstractly in the locale they can be made available
-  conveniently in the context through an auxiliary local interpretation (keyword
-  \isakeyword{interpret}).  This interpretation is inside the proof of the global
-  interpretation.  The third revision of the example illustrates this.  *}
+text {* In the above example, the fact that @{term "op \<le>"} is a partial
+  order for the integers was used in the second goal to
+  discharge the premise in the definition of @{text "op \<sqsubset>"}.  In
+  general, proofs of the equations not only may involve definitions
+  from the interpreted locale but arbitrarily complex arguments in the
+  context of the locale.  Therefore is would be convenient to have the
+  interpreted locale conclusions temporary available in the proof.
+  This can be achieved by a locale interpretation in the proof body.
+  The command for local interpretations is \isakeyword{interpret}.  We
+  repeat the example from the previous section to illustrate this. *}
 
-interpretation %visible nat: partial_order "op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool"
-  where "partial_order.less op \<le> (x::nat) y = (x < y)"
-proof -
-  show "partial_order (op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool)"
-    by unfold_locales auto
-  then interpret nat: partial_order "op \<le> :: [nat, nat] \<Rightarrow> bool" .
-  show "partial_order.less op \<le> (x::nat) y = (x < y)"
-    unfolding nat.less_def by auto
-qed
+  interpretation %visible int: partial_order "op \<le> :: int \<Rightarrow> int \<Rightarrow> bool"
+    where "partial_order.less op \<le> (x::int) y = (x < y)"
+  proof -
+    show "partial_order (op \<le> :: int \<Rightarrow> int \<Rightarrow> bool)"
+      by unfold_locales auto
+    then interpret int: partial_order "op \<le> :: [int, int] \<Rightarrow> bool" .
+    show "partial_order.less op \<le> (x::int) y = (x < y)"
+      unfolding int.less_def by auto
+  qed
 
-text {* The inner interpretation does not require an elaborate new
-  proof, it is immediate from the preceding fact and proved with
-  ``.''.  It enriches the local proof context by the very theorems
+text {* The inner interpretation is immediate from the preceding fact
+  and proved by assumption (Isar short hand ``.'').  It enriches the
+  local proof context by the theorems
   also obtained in the interpretation from Section~\ref{sec:po-first},
-  and @{text nat.less_def} may directly be used to unfold the
+  and @{text int.less_def} may directly be used to unfold the
   definition.  Theorems from the local interpretation disappear after
-  leaving the proof context --- that is, after the closing
-  \isakeyword{qed} --- and are then replaced by those with the desired
-  substitutions of the strict order.  *}
+  leaving the proof context --- that is, after the succeeding
+  \isakeyword{next} or \isakeyword{qed} statement. *}
 
 
 subsection {* Further Interpretations *}
 
-text {* Further interpretations are necessary to reuse theorems from
-  the other locales.  In @{text lattice} the operations @{text \<sqinter>} and
-  @{text \<squnion>} are substituted by @{term "min :: nat \<Rightarrow> nat \<Rightarrow> nat"} and
-  @{term "max :: nat \<Rightarrow> nat \<Rightarrow> nat"}.  The entire proof for the
-  interpretation is reproduced in order to give an example of a more
-  elaborate interpretation proof.  *}
+text {* Further interpretations are necessary for
+  the other locales.  In @{text lattice} the operations~@{text \<sqinter>}
+  and~@{text \<squnion>} are substituted by @{term "min :: int \<Rightarrow> int \<Rightarrow> int"}
+  and @{term "max :: int \<Rightarrow> int \<Rightarrow> int"}.  The entire proof for the
+  interpretation is reproduced to give an example of a more
+  elaborate interpretation proof.  Note that the equations are named
+  so they can be used in a later example.  *}
 
-interpretation %visible nat: lattice "op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool"
-  where "lattice.meet op \<le> (x::nat) y = min x y"
-    and "lattice.join op \<le> (x::nat) y = max x y"
-proof -
-  show "lattice (op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool)"
-    txt {* We have already shown that this is a partial order, *}
-    apply unfold_locales
-    txt {* hence only the lattice axioms remain to be shown: @{subgoals
-      [display]}  After unfolding @{text is_inf} and @{text is_sup}, *}
-    apply (unfold nat.is_inf_def nat.is_sup_def)
-    txt {* the goals become @{subgoals [display]} which can be solved
-      by Presburger arithmetic. *}
-    by arith+
-  txt {* In order to show the equations, we put ourselves in a
-    situation where the lattice theorems can be used in a convenient way. *}
-  then interpret nat: lattice "op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool" .
-  show "lattice.meet op \<le> (x::nat) y = min x y"
-    by (bestsimp simp: nat.meet_def nat.is_inf_def)
-  show "lattice.join op \<le> (x::nat) y = max x y"
-    by (bestsimp simp: nat.join_def nat.is_sup_def)
-qed
+  interpretation %visible int: lattice "op \<le> :: int \<Rightarrow> int \<Rightarrow> bool"
+    where int_min_eq: "lattice.meet op \<le> (x::int) y = min x y"
+      and int_max_eq: "lattice.join op \<le> (x::int) y = max x y"
+  proof -
+    show "lattice (op \<le> :: int \<Rightarrow> int \<Rightarrow> bool)"
+      txt {* \normalsize We have already shown that this is a partial
+	order, *}
+      apply unfold_locales
+      txt {* \normalsize hence only the lattice axioms remain to be
+	shown.
+        @{subgoals [display]}
+	By @{text is_inf} and @{text is_sup}, *}
+      apply (unfold int.is_inf_def int.is_sup_def)
+      txt {* \normalsize the goals are transformed to these
+	statements:
+	@{subgoals [display]}
+	This is Presburger arithmetic, which can be solved by the
+	method @{text arith}. *}
+      by arith+
+    txt {* \normalsize In order to show the equations, we put ourselves
+      in a situation where the lattice theorems can be used in a
+      convenient way. *}
+    then interpret int: lattice "op \<le> :: int \<Rightarrow> int \<Rightarrow> bool" .
+    show "lattice.meet op \<le> (x::int) y = min x y"
+      by (bestsimp simp: int.meet_def int.is_inf_def)
+    show "lattice.join op \<le> (x::int) y = max x y"
+      by (bestsimp simp: int.join_def int.is_sup_def)
+  qed
 
-text {* Next follows that @{text \<le>} is a total order. *}
+text {* Next follows that @{text "op \<le>"} is a total order, again for
+  the integers. *}
 
-interpretation %visible nat: total_order "op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool"
-  by unfold_locales arith
+  interpretation %visible int: total_order "op \<le> :: int \<Rightarrow> int \<Rightarrow> bool"
+    by unfold_locales arith
 
 text {* Theorems that are available in the theory at this point are shown in
-  Table~\ref{tab:nat-lattice}.
+  Table~\ref{tab:int-lattice}.  Two points are worth noting:
 
 \begin{table}
 \hrule
 \vspace{2ex}
 \begin{center}
 \begin{tabular}{l}
-  @{thm [source] nat.less_def} from locale @{text partial_order}: \\
-  \quad @{thm nat.less_def} \\
-  @{thm [source] nat.meet_left} from locale @{text lattice}: \\
-  \quad @{thm nat.meet_left} \\
-  @{thm [source] nat.join_distr} from locale @{text distrib_lattice}: \\
-  \quad @{thm nat.join_distr} \\
-  @{thm [source] nat.less_total} from locale @{text total_order}: \\
-  \quad @{thm nat.less_total}
+  @{thm [source] int.less_def} from locale @{text partial_order}: \\
+  \quad @{thm int.less_def} \\
+  @{thm [source] int.meet_left} from locale @{text lattice}: \\
+  \quad @{thm int.meet_left} \\
+  @{thm [source] int.join_distr} from locale @{text distrib_lattice}: \\
+  \quad @{thm int.join_distr} \\
+  @{thm [source] int.less_total} from locale @{text total_order}: \\
+  \quad @{thm int.less_total}
 \end{tabular}
 \end{center}
 \hrule
-\caption{Interpreted theorems for @{text \<le>} on the natural numbers.}
-\label{tab:nat-lattice}
+\caption{Interpreted theorems for~@{text \<le>} on the integers.}
+\label{tab:int-lattice}
 \end{table}
 
-  Note that since the locale hierarchy reflects that total orders are
-  distributive lattices, an explicit interpretation of distributive
-  lattices for the order relation on natural numbers is not neccessary.
-
-  Why not push this idea further and just give the last interpretation
-  as a single interpretation instead of the sequence of three?  The
-  reasons for this are twofold:
 \begin{itemize}
 \item
-  Often it is easier to work in an incremental fashion, because later
-  interpretations require theorems provided by earlier
-  interpretations.
+  Locale @{text distrib_lattice} was also interpreted.  Since the
+  locale hierarchy reflects that total orders are distributive
+  lattices, the interpretation of the latter was inserted
+  automatically with the interpretation of the former.  In general,
+  interpretation traverses the locale hierarchy upwards and interprets
+  all encountered locales, regardless whether imported or proved via
+  the \isakeyword{sublocale} command.  Existing interpretations are
+  skipped avoiding duplicate work.
 \item
-  Assume that a definition is made in some locale $l_1$, and that $l_2$
-  imports $l_1$.  Let an equation for the definition be
-  proved in an interpretation of $l_2$.  The equation will be unfolded
-  in interpretations of theorems added to $l_2$ or below in the import
-  hierarchy, but not for theorems added above $l_2$.
-  Hence, an equation interpreting a definition should always be given in
-  an interpretation of the locale where the definition is made, not in
-  an interpretation of a locale further down the hierarchy.
+  The predicate @{term "op <"} appears in theorem @{thm [source]
+  int.less_total}
+  although an equation for the replacement of @{text "op \<sqsubset>"} was only
+  given in the interpretation of @{text partial_order}.  The
+  interpretation equations are pushed downwards the hierarchy for
+  related interpretations --- that is, for interpretations that share
+  the instances of parameters they have in common.
 \end{itemize}
   *}
 
-
-subsection {* Lattice @{text "dvd"} on @{typ nat} *}
-
-text {* Divisibility on the natural numbers is a distributive lattice
-  but not a total order.  Interpretation again proceeds
-  incrementally. *}
-
-interpretation nat_dvd: partial_order "op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool"
-  where "partial_order.less op dvd (x::nat) y = (x dvd y \<and> x \<noteq> y)"
-proof -
-  show "partial_order (op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool)"
-    by unfold_locales (auto simp: dvd_def)
-  then interpret nat_dvd: partial_order "op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool" .
-  show "partial_order.less op dvd (x::nat) y = (x dvd y \<and> x \<noteq> y)"
-    apply (unfold nat_dvd.less_def)
-    apply auto
-    done
-qed
-
-text {* Note that in Isabelle/HOL there is no symbol for strict
-  divisibility.  Instead, interpretation substitutes @{term "x dvd y \<and>
-  x \<noteq> y"}.  *}
-
-interpretation nat_dvd: lattice "op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool"
-  where nat_dvd_meet_eq: "lattice.meet (op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool) = gcd"
-    and nat_dvd_join_eq: "lattice.join (op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool) = lcm"
-proof -
-  show "lattice (op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool)"
-    apply unfold_locales
-    apply (unfold nat_dvd.is_inf_def nat_dvd.is_sup_def)
-    apply (rule_tac x = "gcd x y" in exI)
-    apply auto [1]
-    apply (rule_tac x = "lcm x y" in exI)
-    apply (auto intro: lcm_least_nat)
-    done
-  then interpret nat_dvd: lattice "op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool" .
-  show "lattice.meet (op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool) = gcd"
-    apply (auto simp add: expand_fun_eq)
-    apply (unfold nat_dvd.meet_def)
-    apply (rule the_equality)
-    apply (unfold nat_dvd.is_inf_def)
-    by auto
-  show "lattice.join (op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool) = lcm"
-    apply (auto simp add: expand_fun_eq)
-    apply (unfold nat_dvd.join_def)
-    apply (rule the_equality)
-    apply (unfold nat_dvd.is_sup_def)
-    apply (auto intro: lcm_least_nat iff: lcm_unique_nat)
-    done
-qed
-
-text {* Equations @{thm [source] nat_dvd_meet_eq} and @{thm [source]
-  nat_dvd_join_eq} are used in the main part the subsequent
-  interpretation. *}
-
-(*
-definition
-  is_lcm :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
-  "is_lcm p m n \<longleftrightarrow> m dvd p \<and> n dvd p \<and>
-    (\<forall>d. m dvd d \<longrightarrow> n dvd d \<longrightarrow> p dvd d)"
-
-lemma is_gcd: "is_lcm (lcm (m, n)) m n"
-  by (simp add: is_lcm_def lcm_least)
-
-lemma gcd_lcm_distr_lemma:
-  "[| is_gcd g1 x l1; is_lcm l1 y z; is_gcd g2 x y; is_gcd g3 x z |] ==> is_lcm g1 g2 g3"
-apply (unfold is_gcd_def is_lcm_def dvd_def)
-apply (clarsimp simp: mult_ac)
-apply (blast intro: mult_is_0)
-thm mult_is_0 [THEN iffD1]
-*)
-
-lemma %invisible gcd_lcm_distr:
-  "gcd x (lcm y z) = lcm (gcd x y) (gcd x z)" sorry
-
-interpretation %visible nat_dvd:
-  distrib_lattice "op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool"
-  apply unfold_locales
-  txt {* @{subgoals [display]} *}
-  apply (unfold nat_dvd_meet_eq nat_dvd_join_eq)
-  txt {* @{subgoals [display]} *}
-  apply (rule gcd_lcm_distr) done
-
-
-text {* Theorems that are available in the theory after these
-  interpretations are shown in Table~\ref{tab:nat-dvd-lattice}.
-
-\begin{table}
-\hrule
-\vspace{2ex}
-\begin{center}
-\begin{tabular}{l}
-  @{thm [source] nat_dvd.less_def} from locale @{text partial_order}: \\
-  \quad @{thm nat_dvd.less_def} \\
-  @{thm [source] nat_dvd.meet_left} from locale @{text lattice}: \\
-  \quad @{thm nat_dvd.meet_left} \\
-  @{thm [source] nat_dvd.join_distr} from locale @{text distrib_lattice}: \\
-  \quad @{thm nat_dvd.join_distr} \\
-\end{tabular}
-\end{center}
-\hrule
-\caption{Interpreted theorems for @{text dvd} on the natural numbers.}
-\label{tab:nat-dvd-lattice}
-\end{table}
-  *}
-
-text {*
-  The syntax of the interpretation commands is shown in
-  Table~\ref{tab:commands}.  The grammar refers to
-  \textit{expression}, which stands for a \emph{locale} expression.
-  Locale expressions are discussed in the following section.
-  *}
+text {* The interpretations for a locale $n$ within the current
+  theory may be inspected with \isakeyword{print\_interps}~$n$.  This
+  prints the list of instances of $n$, for which interpretations exist.
+  For example, \isakeyword{print\_interps} @{term partial_order}
+  outputs the following:
+\begin{small}
+\begin{alltt}
+  int! : partial_order "op \(\le\)"
+\end{alltt}
+\end{small}
+  Of course, there is only one interpretation.
+  The interpretation qualifier on the left is decorated with an
+  exclamation point.  This means that it is mandatory.  Qualifiers
+  can either be \emph{mandatory} or \emph{optional}, designated by
+  ``!'' or ``?'' respectively.  Mandatory qualifiers must occur in a
+  name reference while optional ones need not.  Mandatory qualifiers
+  prevent accidental hiding of names, while optional qualifiers can be
+  more convenient to use.  For \isakeyword{interpretation}, the
+  default is ``!''.
+*}
 
 
 section {* Locale Expressions \label{sec:expressions} *}
 
 text {*
-  A map @{term \<phi>} between partial orders @{text \<sqsubseteq>} and @{text \<preceq>}
+  A map~@{term \<phi>} between partial orders~@{text \<sqsubseteq>} and~@{text \<preceq>}
   is called order preserving if @{text "x \<sqsubseteq> y"} implies @{text "\<phi> x \<preceq>
   \<phi> y"}.  This situation is more complex than those encountered so
   far: it involves two partial orders, and it is desirable to use the
   existing locale for both.
 
-  Inspecting the grammar of locale commands in
-  Table~\ref{tab:commands} reveals that the import of a locale can be
-  more than just a single locale.  In general, the import is a
-  \emph{locale expression}, which enables to combine locales
-  and instantiate parameters.  A locale expression is a sequence of
-  locale \emph{instances} followed by an optional \isakeyword{for}
-  clause.  Each instance consists of a locale reference, which may be
-  preceded by a qualifer and succeeded by instantiations of the
-  parameters of that locale.  Instantiations may be either positional
-  or through explicit mappings of parameters to arguments.
+  A locale for order preserving maps requires three parameters: @{text
+  le}~(\isakeyword{infixl}~@{text \<sqsubseteq>}) and @{text
+  le'}~(\isakeyword{infixl}~@{text \<preceq>}) for the orders and~@{text \<phi>}
+  for the map.
+
+  In order to reuse the existing locale for partial orders, which has
+  the single parameter~@{text le}, it must be imported twice, once
+  mapping its parameter to~@{text le} from the new locale and once
+  to~@{text le'}.  This can be achieved with a compound locale
+  expression.
 
-  Using a locale expression, a locale for order
-  preserving maps can be declared in the following way.  *}
+  In general, a locale expression is a sequence of \emph{locale instances}
+  separated by~``$\textbf{+}$'' and followed by a \isakeyword{for}
+  clause.
+  An instance has the following format:
+\begin{quote}
+  \textit{qualifier} \textbf{:} \textit{locale-name}
+  \textit{parameter-instantiation}
+\end{quote}
+  We have already seen locale instances as arguments to
+  \isakeyword{interpretation} in Section~\ref{sec:interpretation}.
+  As before, the qualifier serves to disambiguate names from
+  different instances of the same locale.  While in
+  \isakeyword{interpretation} qualifiers default to mandatory, in
+  import and in the \isakeyword{sublocale} command, they default to
+  optional.
+
+  Since the parameters~@{text le} and~@{text le'} are to be partial
+  orders, our locale for order preserving maps will import the these
+  instances:
+\begin{small}
+\begin{alltt}
+  le: partial_order le
+  le': partial_order le'
+\end{alltt}
+\end{small}
+  For matter of convenience we choose to name parameter names and
+  qualifiers alike.  This is an arbitrary decision.  Technically, qualifiers
+  and parameters are unrelated.
+
+  Having determined the instances, let us turn to the \isakeyword{for}
+  clause.  It serves to declare locale parameters in the same way as
+  the context element \isakeyword{fixes} does.  Context elements can
+  only occur after the import section, and therefore the parameters
+  referred to in the instances must be declared in the \isakeyword{for}
+  clause.  The \isakeyword{for} clause is also where the syntax of these
+  parameters is declared.
+
+  Two context elements for the map parameter~@{text \<phi>} and the
+  assumptions that it is order preserving complete the locale
+  declaration. *}
 
   locale order_preserving =
     le: partial_order le + le': partial_order le'
       for le (infixl "\<sqsubseteq>" 50) and le' (infixl "\<preceq>" 50) +
-    fixes \<phi> :: "'a \<Rightarrow> 'b"
+    fixes \<phi>
     assumes hom_le: "x \<sqsubseteq> y \<Longrightarrow> \<phi> x \<preceq> \<phi> y"
 
-text {* The second and third line contain the expression --- two
-  instances of the partial order locale where the parameter is
-  instantiated to @{text le}
-  and @{text le'}, respectively.  The \isakeyword{for} clause consists
-  of parameter declarations and is similar to the context element
-  \isakeyword{fixes}.  The notable difference is that the
-  \isakeyword{for} clause is part of the expression, and only
-  parameters defined in the expression may occur in its instances.
+text (in order_preserving) {* Here are examples of theorems that are
+  available in the locale:
 
-  Instances define \emph{morphisms} on locales.  Their effect on the
-  parameters is lifted to terms, propositions and theorems in the
-  canonical way,
-  and thus to the assumptions and conclusions of a locale.  The
-  assumption of a locale expression is the conjunction of the
-  assumptions of the instances.  The conclusions of a sequence of
-  instances are obtained by appending the conclusions of the
-  instances in the order of the sequence.
+  \hspace*{1em}@{thm [source] hom_le}: @{thm hom_le}
+
+  \hspace*{1em}@{thm [source] le.less_le_trans}: @{thm le.less_le_trans}
 
-  The qualifiers in the expression are already a familiar concept from
-  the \isakeyword{interpretation} command
-  (Section~\ref{sec:po-first}).  Here, they serve to distinguish names
-  (in particular theorem names) for the two partial orders within the
-  locale.  Qualifiers in the \isakeyword{locale} command (and in
-  \isakeyword{sublocale}) default to optional --- that is, they need
-  not occur in references to the qualified names.  Here are examples
-  of theorems in locale @{text order_preserving}: *}
-
-context %invisible order_preserving begin
-
-text {*
-  @{thm [source] le.less_le_trans}: @{thm le.less_le_trans}
-
-  @{thm [source] hom_le}: @{thm hom_le}
-  *}
-
-text {* The theorems for the partial order @{text \<preceq>}
-  are qualified by @{text le'}.  For example, @{thm [source]
-  le'.less_le_trans}: @{thm [display, indent=2] le'.less_le_trans} *}
-
-end %invisible
-
-text {* This example reveals that there is no infix syntax for the
-  strict operation associated with @{text \<preceq>}.  This can be declared
-  through an abbreviation.  *}
+  \hspace*{1em}@{thm [source] le'.less_le_trans}:
+  @{thm [display, indent=4] le'.less_le_trans}
+  While there is infix syntax for the strict operation associated to
+  @{term "op \<sqsubseteq>"}, there is none for the strict version of @{term
+  "op \<preceq>"}.  The abbreviation @{text less} with its infix syntax is only
+  available for the original instance it was declared for.  We may
+  introduce the abbreviation @{text less'} with infix syntax~@{text \<prec>}
+  with the following declaration: *}
 
   abbreviation (in order_preserving)
     less' (infixl "\<prec>" 50) where "less' \<equiv> partial_order.less le'"
@@ -312,33 +237,56 @@
   @{thm [source]  le'.less_le_trans}:
   @{thm [display, indent=2] le'.less_le_trans} *}
 
-text (in order_preserving)  {* Qualifiers not only apply to theorem names, but also to names
-  introduced by definitions and abbreviations.  For example, in @{text
-  partial_order} the name @{text less} abbreviates @{term
-  "partial_order.less le"}.  Therefore, in @{text order_preserving}
-  the qualified name @{text le.less} abbreviates @{term
-  "partial_order.less le"} and @{text le'.less} abbreviates @{term
-  "partial_order.less le'"}.  Hence, the equation in the abbreviation
-  above could have been written more concisely as @{text "less' \<equiv>
-  le'.less"}. *}
+text {* There are short notations for locale expressions.  These are
+  discussed in the following. *}
+
+
+subsection {* Default Instantiations *}
+
+text {* 
+  It is possible to omit parameter instantiations.  The
+  instantiation then defaults to the name of
+  the parameter itself.  For example, the locale expression @{text
+  partial_order} is short for @{text "partial_order le"}, since the
+  locale's single parameter is~@{text le}.  We took advantage of this
+  in the \isakeyword{sublocale} declarations of
+  Section~\ref{sec:changing-the-hierarchy}. *}
+
+
+subsection {* Implicit Parameters \label{sec:implicit-parameters} *}
+
+text {* In a locale expression that occurs within a locale
+  declaration, omitted parameters additionally extend the (possibly
+  empty) \isakeyword{for} clause.
 
-text {* Readers may find the declaration of locale @{text
-  order_preserving} a little awkward, because the declaration and
-  concrete syntax for @{text le} from @{text partial_order} are
-  repeated in the declaration of @{text order_preserving}.  Locale
-  expressions provide a convenient short hand for this.  A parameter
-  in an instance is \emph{untouched} if no instantiation term is
-  provided for it.  In positional instantiations, a parameter position
-  may be skipped with an underscore, and it is allowed to give fewer
-  instantiation terms than the instantiated locale's number of
-  parameters.  In named instantiations, instantiation pairs for
-  certain parameters may simply be omitted.  Untouched parameters are
-  implicitly declared by the locale expression and with their concrete
-  syntax.  In the sequence of parameters, they appear before the
-  parameters from the \isakeyword{for} clause.
+  The \isakeyword{for} clause is a general construct of Isabelle/Isar
+  to mark names occurring in the preceding declaration as ``arbitrary
+  but fixed''.  This is necessary for example, if the name is already
+  bound in a surrounding context.  In a locale expression, names
+  occurring in parameter instantiations should be bound by a
+  \isakeyword{for} clause whenever these names are not introduced
+  elsewhere in the context --- for example, on the left hand side of a
+  \isakeyword{sublocale} declaration.
 
-  The following locales illustrate this.  A map @{text \<phi>} is a
-  lattice homomorphism if it preserves meet and join. *}
+  There is an exception to this rule in locale declarations, where the
+  \isakeyword{for} clause serves to declare locale parameters.  Here,
+  locale parameters for which no parameter instantiation is given are
+  implicitly added, with their mixfix syntax, at the beginning of the
+  \isakeyword{for} clause.  For example, in a locale declaration, the
+  expression @{text partial_order} is short for
+\begin{small}
+\begin{alltt}
+  partial_order le \isakeyword{for} le (\isakeyword{infixl} "\(\sqsubseteq\)" 50)\textrm{.}
+\end{alltt}
+\end{small}
+  This short hand was used in the locale declarations throughout
+  Section~\ref{sec:import}.
+ *}
+
+text{*
+  The following locale declarations provide more examples.  A
+  map~@{text \<phi>} is a lattice homomorphism if it preserves meet and
+  join. *}
 
   locale lattice_hom =
     le: lattice + le': lattice le' for le' (infixl "\<preceq>" 50) +
@@ -346,29 +294,43 @@
     assumes hom_meet: "\<phi> (x \<sqinter> y) = le'.meet (\<phi> x) (\<phi> y)"
       and hom_join: "\<phi> (x \<squnion> y) = le'.join (\<phi> x) (\<phi> y)"
 
-  abbreviation (in lattice_hom)
-    meet' (infixl "\<sqinter>''" 50) where "meet' \<equiv> le'.meet"
-  abbreviation (in lattice_hom)
-    join' (infixl "\<squnion>''" 50) where "join' \<equiv> le'.join"
+text {* The parameter instantiation in the first instance of @{term
+  lattice} is omitted.  This causes the parameter~@{text le} to be
+  added to the \isakeyword{for} clause, and the locale has
+  parameters~@{text le},~@{text le'} and, of course,~@{text \<phi>}.
 
-text {* A homomorphism is an endomorphism if both orders coincide. *}
+  Before turning to the second example, we complete the locale by
+  providing infix syntax for the meet and join operations of the
+  second lattice.
+*}
+
+  context lattice_hom begin
+  abbreviation meet' (infixl "\<sqinter>''" 50) where "meet' \<equiv> le'.meet"
+  abbreviation join' (infixl "\<squnion>''" 50) where "join' \<equiv> le'.join"
+  end
+
+text {* The next example makes radical use of the short hand
+  facilities.  A homomorphism is an endomorphism if both orders
+  coincide. *}
 
   locale lattice_end = lattice_hom _ le
 
-text {* In this declaration, the first parameter of @{text
-  lattice_hom}, @{text le}, is untouched and is then used to instantiate
-  the second parameter.  Its concrete syntax is preserved. *}
-
+text {* The notation~@{text _} enables to omit a parameter in a
+  positional instantiation.  The omitted parameter,~@{text le} becomes
+  the parameter of the declared locale and is, in the following
+  position, used to instantiate the second parameter of @{text
+  lattice_hom}.  The effect is that of identifying the first in second
+  parameter of the homomorphism locale. *}
 
 text {* The inheritance diagram of the situation we have now is shown
   in Figure~\ref{fig:hom}, where the dashed line depicts an
-  interpretation which is introduced below.  Renamings are
-  indicated by $\sqsubseteq \mapsto \preceq$ etc.  The expression
-  imported by @{text lattice_end} identifies the first and second
-  parameter of @{text lattice_hom}.  By looking at the inheritance diagram it would seem
+  interpretation which is introduced below.  Parameter instantiations
+  are indicated by $\sqsubseteq \mapsto \preceq$ etc.  By looking at
+  the inheritance diagram it would seem
   that two identical copies of each of the locales @{text
-  partial_order} and @{text lattice} are imported.  This is not the
-  case!  Inheritance paths with identical morphisms are detected and
+  partial_order} and @{text lattice} are imported by @{text
+  lattice_end}.  This is not the case!  Inheritance paths with
+  identical morphisms are automatically detected and
   the conclusions of the respective locales appear only once.
 
 \begin{figure}
@@ -421,21 +383,130 @@
   lattice_hom}, for example
   @{thm [source] hom_le}:
   @{thm [display, indent=2] hom_le}
+  This theorem will be useful in the following section.
   *}
 
 
+section {* Conditional Interpretation *}
+
+text {* There are situations where an interpretation is not possible
+  in the general case since the desired property is only valid if
+  certain conditions are fulfilled.  Take, for example, the function
+  @{text "\<lambda>i. n * i"} that scales its argument by a constant factor.
+  This function is order preserving (and even a lattice endomorphism)
+  with respect to @{term "op \<le>"} provided @{text "n \<ge> 0"}.
+
+  It is not possible to express this using a global interpretation,
+  because it is in general unspecified whether~@{term n} is
+  non-negative, but one may make an interpretation in an inner context
+  of a proof where full information is available.
+  This is not fully satisfactory either, since potentially
+  interpretations may be required to make interpretations in many
+  contexts.  What is
+  required is an interpretation that depends on the condition --- and
+  this can be done with the \isakeyword{sublocale} command.  For this
+  purpose, we introduce a locale for the condition. *}
+
+  locale non_negative =
+    fixes n :: int
+    assumes non_neg: "0 \<le> n"
+
+text {* It is again convenient to make the interpretation in an
+  incremental fashion, first for order preserving maps, the for
+  lattice endomorphisms. *}
+
+  sublocale non_negative \<subseteq>
+      order_preserving "op \<le>" "op \<le>" "\<lambda>i. n * i"
+    using non_neg by unfold_locales (rule mult_left_mono)
+
+text {* While the proof of the previous interpretation
+  is straightforward from monotonicity lemmas for~@{term "op *"}, the
+  second proof follows a useful pattern. *}
+
+  sublocale %visible non_negative \<subseteq> lattice_end "op \<le>" "\<lambda>i. n * i"
+  proof (unfold_locales, unfold int_min_eq int_max_eq)
+    txt {* \normalsize Unfolding the locale predicates \emph{and} the
+      interpretation equations immediately yields two subgoals that
+      reflect the core conjecture.
+      @{subgoals [display]}
+      It is now necessary to show, in the context of @{term
+      non_negative}, that multiplication by~@{term n} commutes with
+      @{term min} and @{term max}. *}
+  qed (auto simp: hom_le)
+
+text (in order_preserving) {* The lemma @{thm [source] hom_le}
+  simplifies a proof that would have otherwise been lengthy and we may
+  consider making it a default rule for the simplifier: *}
+
+  lemmas (in order_preserving) hom_le [simp]
+
+
+subsection {* Avoiding Infinite Chains of Interpretations
+  \label{sec:infinite-chains} *}
+
+text {* Similar situations arise frequently in formalisations of
+  abstract algebra where it is desirable to express that certain
+  constructions preserve certain properties.  For example, polynomials
+  over rings are rings, or --- an example from the domain where the
+  illustrations of this tutorial are taken from --- a partial order
+  may be obtained for a function space by point-wise lifting of the
+  partial order of the co-domain.  This corresponds to the following
+  interpretation: *}
+
+  sublocale %visible partial_order \<subseteq> f: partial_order "\<lambda>f g. \<forall>x. f x \<sqsubseteq> g x"
+    oops
+
+text {* Unfortunately this is a cyclic interpretation that leads to an
+  infinite chain, namely
+  @{text [display, indent=2] "partial_order \<subseteq> partial_order (\<lambda>f g. \<forall>x. f x \<sqsubseteq> g x) \<subseteq>
+  partial_order (\<lambda>f g. \<forall>x y. f x y \<sqsubseteq> g x y) \<subseteq>  \<dots>"}
+  and the interpretation is rejected.
+
+  Instead it is necessary to declare a locale that is logically
+  equivalent to @{term partial_order} but serves to collect facts
+  about functions spaces where the co-domain is a partial order, and
+  to make the interpretation in its context: *}
+
+  locale fun_partial_order = partial_order
+
+  sublocale fun_partial_order \<subseteq>
+      f: partial_order "\<lambda>f g. \<forall>x. f x \<sqsubseteq> g x"
+    by unfold_locales (fast,rule,fast,blast intro: trans)
+
+text {* It is quite common in abstract algebra that such a construction
+  maps a hierarchy of algebraic structures (or specifications) to a
+  related hierarchy.  By means of the same lifting, a function space
+  is a lattice if its co-domain is a lattice: *}
+
+  locale fun_lattice = fun_partial_order + lattice
+
+  sublocale fun_lattice \<subseteq> f: lattice "\<lambda>f g. \<forall>x. f x \<sqsubseteq> g x"
+    proof unfold_locales
+    fix f g
+    have "partial_order.is_inf (\<lambda>f g. \<forall>x. f x \<sqsubseteq> g x) f g (\<lambda>x. f x \<sqinter> g x)"
+      apply (rule is_infI) apply rule+ apply (drule spec, assumption)+ done
+    then show "\<exists>inf. partial_order.is_inf (\<lambda>f g. \<forall>x. f x \<sqsubseteq> g x) f g inf"
+      by fast
+  next
+    fix f g
+    have "partial_order.is_sup (\<lambda>f g. \<forall>x. f x \<sqsubseteq> g x) f g (\<lambda>x. f x \<squnion> g x)"
+      apply (rule is_supI) apply rule+ apply (drule spec, assumption)+ done
+    then show "\<exists>sup. partial_order.is_sup (\<lambda>f g. \<forall>x. f x \<sqsubseteq> g x) f g sup"
+      by fast
+  qed
+
 
 section {* Further Reading *}
 
 text {* More information on locales and their interpretation is
   available.  For the locale hierarchy of import and interpretation
-  dependencies see \cite{Ballarin2006a}; interpretations in theories
-  and proofs are covered in \cite{Ballarin2006b}.  In the latter, we
+  dependencies see~\cite{Ballarin2006a}; interpretations in theories
+  and proofs are covered in~\cite{Ballarin2006b}.  In the latter, I
   show how interpretation in proofs enables to reason about families
   of algebraic structures, which cannot be expressed with locales
   directly.
 
-  Haftmann and Wenzel \cite{HaftmannWenzel2007} overcome a restriction
+  Haftmann and Wenzel~\cite{HaftmannWenzel2007} overcome a restriction
   of axiomatic type classes through a combination with locale
   interpretation.  The result is a Haskell-style class system with a
   facility to generate ML and Haskell code.  Classes are sufficient for
@@ -444,10 +515,21 @@
   category.  Order preserving maps, homomorphisms and vector spaces,
   on the other hand, do not.
 
-  The original work of Kamm\"uller on locales \cite{KammullerEtAl1999}
-  may be of interest from a historical perspective.  The mathematical
-  background on orders and lattices is taken from Jacobson's textbook
-  on algebra \cite[Chapter~8]{Jacobson1985}.
+  The locales reimplementation for Isabelle 2009 provides, among other
+  improvements, a clean integration with Isabelle/Isar's local theory
+  mechanisms, which are described in another paper by Haftmann and
+  Wenzel~\cite{HaftmannWenzel2009}.
+
+  The original work of Kamm\"uller on locales~\cite{KammullerEtAl1999}
+  may be of interest from a historical perspective.  My previous
+  report on locales and locale expressions~\cite{Ballarin2004a}
+  describes a simpler form of expressions than available now and is
+  outdated. The mathematical background on orders and lattices is
+  taken from Jacobson's textbook on algebra~\cite[Chapter~8]{Jacobson1985}.
+
+  The sources of this tutorial, which include all proofs, are
+  available with the Isabelle distribution at
+  \url{http://isabelle.in.tum.de}.
   *}
 
 text {*
@@ -542,8 +624,9 @@
   \multicolumn{3}{l}{Diagnostics} \\
 
   \textit{toplevel} & ::=
-  & \textbf{print\_locale} [ ``\textbf{!}'' ] \textit{locale} \\
-  & | & \textbf{print\_locales} 
+  & \textbf{print\_locales} \\
+  & | & \textbf{print\_locale} [ ``\textbf{!}'' ] \textit{locale} \\
+  & | & \textbf{print\_interps} \textit{locale}
 \end{tabular}
 \end{center}
 \hrule
@@ -552,8 +635,22 @@
 \end{table}
   *}
 
+text {* \textbf{Revision History.}  For the present third revision of
+  the tutorial, much of the explanatory text
+  was rewritten.  Inheritance of interpretation equations is
+  available with the forthcoming release of Isabelle, which at the
+  time of editing these notes is expected for the end of 2009.
+  The second revision accommodates changes introduced by the locales
+  reimplementation for Isabelle 2009.  Most notably locale expressions
+  have been generalised from renaming to instantiation.  *}
+
 text {* \textbf{Acknowledgements.}  Alexander Krauss, Tobias Nipkow,
-  Christian Sternagel and Makarius Wenzel have made useful comments on
-  a draft of this document. *}
+  Randy Pollack, Christian Sternagel and Makarius Wenzel have made
+  useful comments on earlier versions of this document.  The section
+  on conditional interpretation was inspired by a number of e-mail
+  enquiries the author received from locale users, and which suggested
+  that this use case is important enough to deserve explicit
+  explanation.  The term \emph{conditional interpretation} is due to
+  Larry Paulson. *}
 
 end
--- a/doc-src/Locales/Locales/ROOT.ML	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/Locales/Locales/ROOT.ML	Wed Oct 21 16:57:57 2009 +0200
@@ -1,4 +1,3 @@
-no_document use_thy "GCD";
 use_thy "Examples1";
 use_thy "Examples2";
-setmp_noncritical quick_and_dirty true use_thy "Examples3";
+use_thy "Examples3";
--- a/doc-src/Locales/Locales/document/Examples.tex	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/Locales/Locales/document/Examples.tex	Wed Oct 21 16:57:57 2009 +0200
@@ -9,7 +9,7 @@
 \isatagtheory
 \isacommand{theory}\isamarkupfalse%
 \ Examples\isanewline
-\isakeyword{imports}\ Main\ GCD\isanewline
+\isakeyword{imports}\ Main\isanewline
 \isakeyword{begin}%
 \endisatagtheory
 {\isafoldtheory}%
@@ -46,19 +46,20 @@
 \[
   \isa{{\isasymAnd}x\isactrlsub {\isadigit{1}}{\isasymdots}x\isactrlsub n{\isachardot}\ {\isasymlbrakk}\ A\isactrlsub {\isadigit{1}}{\isacharsemicolon}\ {\isasymdots}\ {\isacharsemicolon}A\isactrlsub m\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isasymdots}}
 \]
-  where variables \isa{x\isactrlsub {\isadigit{1}}}, \ldots, \isa{x\isactrlsub n} are called
-  \emph{parameters} and the premises $\isa{A\isactrlsub {\isadigit{1}}}, \ldots,
-  \isa{A\isactrlsub m}$ \emph{assumptions}.  A formula \isa{C}
+  where variables~\isa{x\isactrlsub {\isadigit{1}}}, \ldots,~\isa{x\isactrlsub n} are called
+  \emph{parameters} and the premises $\isa{A\isactrlsub {\isadigit{1}}}, \ldots,~\isa{A\isactrlsub m}$ \emph{assumptions}.  A formula~\isa{C}
   is a \emph{theorem} in the context if it is a conclusion
 \[
-%\label{eq-fact-in-context}
   \isa{{\isasymAnd}x\isactrlsub {\isadigit{1}}{\isasymdots}x\isactrlsub n{\isachardot}\ {\isasymlbrakk}\ A\isactrlsub {\isadigit{1}}{\isacharsemicolon}\ {\isasymdots}\ {\isacharsemicolon}A\isactrlsub m\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ C}.
 \]
   Isabelle/Isar's notion of context goes beyond this logical view.
   Its contexts record, in a consecutive order, proved
-  conclusions along with attributes, which
-  may control proof procedures.  Contexts also contain syntax information
-  for parameters and for terms depending on them.%
+  conclusions along with \emph{attributes}, which can provide context
+  specific configuration information for proof procedures and concrete
+  syntax.  From a logical perspective, locales are just contexts that
+  have been made persistent.  To the user, though, they provide
+  powerful means for declaring and combining contexts, and for the
+  reuse of theorems proved in these contexts.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -67,7 +68,7 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-Locales can be seen as persistent contexts.  In its simplest form, a
+In its simplest form, a
   \emph{locale declaration} consists of a sequence of context elements
   declaring parameters (keyword \isakeyword{fixes}) and assumptions
   (keyword \isakeyword{assumes}).  The following is the specification of
@@ -81,27 +82,79 @@
 \ \ \ \ \ \ \isakeyword{and}\ anti{\isacharunderscore}sym\ {\isacharbrackleft}intro{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymlbrakk}\ x\ {\isasymsqsubseteq}\ y{\isacharsemicolon}\ y\ {\isasymsqsubseteq}\ x\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ x\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline
 \ \ \ \ \ \ \isakeyword{and}\ trans\ {\isacharbrackleft}trans{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymlbrakk}\ x\ {\isasymsqsubseteq}\ y{\isacharsemicolon}\ y\ {\isasymsqsubseteq}\ z\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ x\ {\isasymsqsubseteq}\ z{\isachardoublequoteclose}%
 \begin{isamarkuptext}%
-The parameter of this locale is \isa{le}, with infix syntax
-  \isa{{\isasymsqsubseteq}}.  There is an implicit type parameter \isa{{\isacharprime}a}.  It
-  is not necessary to declare parameter types: most general types will
-  be inferred from the context elements for all parameters.
+The parameter of this locale is~\isa{le},
+  which is a binary predicate with infix syntax~\isa{{\isasymsqsubseteq}}.  The
+  parameter syntax is available in the subsequent
+  assumptions, which are the familiar partial order axioms.
+
+  Isabelle recognises unbound names as free variables.  In locale
+  assumptions, these are implicitly universally quantified.  That is,
+  \isa{{\isasymlbrakk}x\ {\isasymsqsubseteq}\ y{\isacharsemicolon}\ y\ {\isasymsqsubseteq}\ z{\isasymrbrakk}\ {\isasymLongrightarrow}\ x\ {\isasymsqsubseteq}\ z} in fact means
+  \isa{{\isasymAnd}x\ y\ z{\isachardot}\ {\isasymlbrakk}x\ {\isasymsqsubseteq}\ y{\isacharsemicolon}\ y\ {\isasymsqsubseteq}\ z{\isasymrbrakk}\ {\isasymLongrightarrow}\ x\ {\isasymsqsubseteq}\ z}.
 
-  The above declaration not only introduces the locale, it also
-  defines the \emph{locale predicate} \isa{partial{\isacharunderscore}order} with
-  definition \isa{partial{\isacharunderscore}order{\isacharunderscore}def}:
+  Two commands are provided to inspect locales:
+  \isakeyword{print\_locales} lists the names of all locales of the
+  current theory; \isakeyword{print\_locale}~$n$ prints the parameters
+  and assumptions of locale $n$; the variation \isakeyword{print\_locale!}~$n$
+  additionally outputs the conclusions that are stored in the locale.
+  We may inspect the new locale
+  by issuing \isakeyword{print\_locale!} \isa{partial{\isacharunderscore}order}.  The output
+  is the following list of context elements.
+\begin{small}
+\begin{alltt}
+  \isakeyword{fixes} le :: "'a \(\Rightarrow\) 'a \(\Rightarrow\)  bool" (\isakeyword{infixl} "\(\sqsubseteq\)" 50)
+  \isakeyword{assumes} "partial_order op \(\sqsubseteq\)"
+  \isakeyword{notes} assumption
+    refl [intro, simp] = `?x \(\sqsubseteq\) ?x`
+    \isakeyword{and}
+    anti_sym [intro] = `\(\isasymlbrakk\)?x \(\sqsubseteq\) ?y; ?y \(\sqsubseteq\) ?x\(\isasymrbrakk\) \(\Longrightarrow\) ?x = ?y`
+    \isakeyword{and}
+    trans [trans] = `\(\isasymlbrakk\)?x \(\sqsubseteq\) ?y; ?y \(\sqsubseteq\) ?z\(\isasymrbrakk\) \(\Longrightarrow\) ?x \(\sqsubseteq\) ?z`
+\end{alltt}
+\end{small}
+  The keyword \isakeyword{notes} denotes a conclusion element.  There
+  is one conclusion, which was added automatically.  Instead, there is
+  only one assumption, namely \isa{partial{\isacharunderscore}order\ op\ {\isasymsqsubseteq}}.  The locale
+  declaration has introduced the predicate \isa{partial{\isacharunderscore}order} to the theory.  This predicate is the
+  \emph{locale predicate}.  Its definition may be inspected by
+  issuing \isakeyword{thm} \isa{partial{\isacharunderscore}order{\isacharunderscore}def}.
   \begin{isabelle}%
 \ \ partial{\isacharunderscore}order\ {\isacharquery}le\ {\isasymequiv}\isanewline
 \isaindent{\ \ }{\isacharparenleft}{\isasymforall}x{\isachardot}\ {\isacharquery}le\ x\ x{\isacharparenright}\ {\isasymand}\isanewline
 \isaindent{\ \ }{\isacharparenleft}{\isasymforall}x\ y{\isachardot}\ {\isacharquery}le\ x\ y\ {\isasymlongrightarrow}\ {\isacharquery}le\ y\ x\ {\isasymlongrightarrow}\ x\ {\isacharequal}\ y{\isacharparenright}\ {\isasymand}\isanewline
 \isaindent{\ \ }{\isacharparenleft}{\isasymforall}x\ y\ z{\isachardot}\ {\isacharquery}le\ x\ y\ {\isasymlongrightarrow}\ {\isacharquery}le\ y\ z\ {\isasymlongrightarrow}\ {\isacharquery}le\ x\ z{\isacharparenright}%
 \end{isabelle}
+  In our example, this is a unary predicate over the parameter of the
+  locale.  It is equivalent to the original assumptions, which have
+  been turned into conclusions and are
+  available as theorems in the context of the locale.  The names and
+  attributes from the locale declaration are associated to these
+  theorems and are effective in the context of the locale.
 
-  The specification of a locale is fixed, but its list of conclusions
+  Each conclusion has a \emph{foundational theorem} as counterpart
+  in the theory.  Technically, this is simply the theorem composed
+  of context and conclusion.  For the transitivity theorem, this is
+  \isa{partial{\isacharunderscore}order{\isachardot}trans}:
+  \begin{isabelle}%
+\ \ partial{\isacharunderscore}order\ {\isacharquery}le\ {\isasymequiv}\isanewline
+\isaindent{\ \ }{\isacharparenleft}{\isasymforall}x{\isachardot}\ {\isacharquery}le\ x\ x{\isacharparenright}\ {\isasymand}\isanewline
+\isaindent{\ \ }{\isacharparenleft}{\isasymforall}x\ y{\isachardot}\ {\isacharquery}le\ x\ y\ {\isasymlongrightarrow}\ {\isacharquery}le\ y\ x\ {\isasymlongrightarrow}\ x\ {\isacharequal}\ y{\isacharparenright}\ {\isasymand}\isanewline
+\isaindent{\ \ }{\isacharparenleft}{\isasymforall}x\ y\ z{\isachardot}\ {\isacharquery}le\ x\ y\ {\isasymlongrightarrow}\ {\isacharquery}le\ y\ z\ {\isasymlongrightarrow}\ {\isacharquery}le\ x\ z{\isacharparenright}%
+\end{isabelle}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsubsection{Targets: Extending Locales%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+The specification of a locale is fixed, but its list of conclusions
   may be extended through Isar commands that take a \emph{target} argument.
   In the following, \isakeyword{definition} and 
   \isakeyword{theorem} are illustrated.
   Table~\ref{tab:commands-with-target} lists Isar commands that accept
-  a target.  There are various ways of specifying the target.  A
+  a target.  Isar provides various ways of specifying the target.  A
   target for a single command may be indicated with keyword
   \isakeyword{in} in the following way:
 
@@ -131,24 +184,21 @@
 \ \ \ \ less\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqsubset}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\isanewline
 \ \ \ \ \isakeyword{where}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymsqsubset}\ y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isasymsqsubseteq}\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}{\isachardoublequoteclose}%
 \begin{isamarkuptext}%
-A definition in a locale depends on the locale parameters.
-  Here, a global constant \isa{partial{\isacharunderscore}order{\isachardot}less} is declared, which is lifted over the
-  locale parameter \isa{le}.  Its definition is the global theorem
-  \isa{partial{\isacharunderscore}order{\isachardot}less{\isacharunderscore}def}:
+The strict order \isa{less} with infix
+  syntax~\isa{{\isasymsqsubset}} is
+  defined in terms of the locale parameter~\isa{le} and the general
+  equality of the object logic we work in.  The definition generates a
+  \emph{foundational constant}
+  \isa{partial{\isacharunderscore}order{\isachardot}less} with definition \isa{partial{\isacharunderscore}order{\isachardot}less{\isacharunderscore}def}:
   \begin{isabelle}%
 \ \ partial{\isacharunderscore}order\ {\isacharquery}le\ {\isasymLongrightarrow}\isanewline
 \isaindent{\ \ }partial{\isacharunderscore}order{\isachardot}less\ {\isacharquery}le\ {\isacharquery}x\ {\isacharquery}y\ {\isacharequal}\ {\isacharparenleft}{\isacharquery}le\ {\isacharquery}x\ {\isacharquery}y\ {\isasymand}\ {\isacharquery}x\ {\isasymnoteq}\ {\isacharquery}y{\isacharparenright}%
 \end{isabelle}
   At the same time, the locale is extended by syntax transformations
-  hiding this construction in the context of the locale.  That is,
-  \isa{partial{\isacharunderscore}order{\isachardot}less\ le} is printed and parsed as infix
-  \isa{{\isasymsqsubset}}.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-Finally, the conclusion of the definition
-  is added to the locale, \isa{less{\isacharunderscore}def}:
+  hiding this construction in the context of the locale.  Here, the
+  abbreviation \isa{less} is available for
+  \isa{partial{\isacharunderscore}order{\isachardot}less\ le}, and it is printed
+  and parsed as infix~\isa{{\isasymsqsubset}}.  Finally, the conclusion \isa{less{\isacharunderscore}def} is added to the locale:
   \begin{isabelle}%
 \ \ {\isacharparenleft}{\isacharquery}x\ {\isasymsqsubset}\ {\isacharquery}y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isacharquery}x\ {\isasymsqsubseteq}\ {\isacharquery}y\ {\isasymand}\ {\isacharquery}x\ {\isasymnoteq}\ {\isacharquery}y{\isacharparenright}%
 \end{isabelle}%
@@ -156,8 +206,9 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-As an example of a theorem statement in the locale, here is the
-  derivation of a transitivity law.%
+The treatment of theorem statements is more straightforward.
+  As an example, here is the derivation of a transitivity law for the
+  strict order relation.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 \ \ \isacommand{lemma}\isamarkupfalse%
@@ -180,13 +231,17 @@
 \endisadelimvisible
 %
 \begin{isamarkuptext}%
-In the context of the proof, assumptions and theorems of the
-  locale may be used.  Attributes are effective: \isa{anti{\isacharunderscore}sym} was
+In the context of the proof, conclusions of the
+  locale may be used like theorems.  Attributes are effective: \isa{anti{\isacharunderscore}sym} was
   declared as introduction rule, hence it is in the context's set of
   rules used by the classical reasoner by default.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
+\isamarkupsubsection{Context Blocks%
+}
+\isamarkuptrue%
+%
 \begin{isamarkuptext}%
 When working with locales, sequences of commands with the same
   target are frequent.  A block of commands, delimited by
@@ -199,7 +254,7 @@
 
   This style of working is illustrated in the block below, where
   notions of infimum and supremum for partial orders are introduced,
-  together with theorems.%
+  together with theorems about their uniqueness.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 \ \ \isacommand{context}\isamarkupfalse%
@@ -250,7 +305,7 @@
 \ is{\isacharunderscore}inf{\isacharunderscore}uniq{\isacharcolon}\ {\isachardoublequoteopen}{\isasymlbrakk}is{\isacharunderscore}inf\ x\ y\ i{\isacharsemicolon}\ is{\isacharunderscore}inf\ x\ y\ i{\isacharprime}{\isasymrbrakk}\ {\isasymLongrightarrow}\ i\ {\isacharequal}\ i{\isacharprime}{\isachardoublequoteclose}\isanewline
 %
 \isadelimproof
-\ \ %
+\ \ \ \ %
 \endisadelimproof
 %
 \isatagproof
@@ -367,7 +422,7 @@
 \ is{\isacharunderscore}sup{\isacharunderscore}uniq{\isacharcolon}\ {\isachardoublequoteopen}{\isasymlbrakk}is{\isacharunderscore}sup\ x\ y\ s{\isacharsemicolon}\ is{\isacharunderscore}sup\ x\ y\ s{\isacharprime}{\isasymrbrakk}\ {\isasymLongrightarrow}\ s\ {\isacharequal}\ s{\isacharprime}{\isachardoublequoteclose}\isanewline
 %
 \isadelimproof
-\ \ %
+\ \ \ \ %
 \endisadelimproof
 %
 \isatagproof
@@ -466,17 +521,12 @@
 \ \ \isacommand{end}\isamarkupfalse%
 %
 \begin{isamarkuptext}%
-Two commands are provided to inspect locales:
-  \isakeyword{print\_locales} lists the names of all locales of the
-  current theory; \isakeyword{print\_locale}~$n$ prints the parameters
-  and assumptions of locale $n$; \isakeyword{print\_locale!}~$n$
-  additionally outputs the conclusions.
-
-  The syntax of the locale commands discussed in this tutorial is
-  shown in Table~\ref{tab:commands}.  The grammer is complete with the
-  exception of additional context elements not discussed here.  See the
-  Isabelle/Isar Reference Manual~\cite{IsarRef}
-  for full documentation.%
+The syntax of the locale commands discussed in this tutorial is
+  shown in Table~\ref{tab:commands}.  The grammar is complete with the
+  exception of the context elements \isakeyword{constrains} and
+  \isakeyword{defines}, which are provided for backward
+  compatibility.  See the Isabelle/Isar Reference
+  Manual~\cite{IsarRef} for full documentation.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -488,11 +538,15 @@
 Algebraic structures are commonly defined by adding operations and
   properties to existing structures.  For example, partial orders
   are extended to lattices and total orders.  Lattices are extended to
-  distributive lattices.
-
-  With locales, this inheritance is achieved through \emph{import} of a
-  locale.  Import is a separate entity in the locale declaration.  If
-  present, it precedes the context elements.%
+  distributive lattices.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+With locales, this kind of inheritance is achieved through
+  \emph{import} of locales.  The import part of a locale declaration,
+  if present, precedes the context elements.  Here is an example,
+  where partial orders are extended to lattices.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 \ \ \isacommand{locale}\isamarkupfalse%
@@ -502,14 +556,13 @@
 \ \ \isakeyword{begin}%
 \begin{isamarkuptext}%
 These assumptions refer to the predicates for infimum
-  and supremum defined in \isa{partial{\isacharunderscore}order}.  We may now introduce
-  the notions of meet and join.%
+  and supremum defined for \isa{partial{\isacharunderscore}order} in the previous
+  section.  We now introduce the notions of meet and join.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 \ \ \isacommand{definition}\isamarkupfalse%
 \isanewline
 \ \ \ \ meet\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqinter}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\ {\isachardoublequoteopen}x\ {\isasymsqinter}\ y\ {\isacharequal}\ {\isacharparenleft}THE\ inf{\isachardot}\ is{\isacharunderscore}inf\ x\ y\ inf{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\isanewline
 \ \ \isacommand{definition}\isamarkupfalse%
 \isanewline
 \ \ \ \ join\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqunion}{\isachardoublequoteclose}\ {\isadigit{6}}{\isadigit{5}}{\isacharparenright}\ \isakeyword{where}\ {\isachardoublequoteopen}x\ {\isasymsqunion}\ y\ {\isacharequal}\ {\isacharparenleft}THE\ sup{\isachardot}\ is{\isacharunderscore}sup\ x\ y\ sup{\isacharparenright}{\isachardoublequoteclose}\isanewline
@@ -1071,8 +1124,10 @@
 \ \ \isacommand{end}\isamarkupfalse%
 %
 \begin{isamarkuptext}%
-Locales for total orders and distributive lattices follow.
-  Each comes with an example theorem.%
+Locales for total orders and distributive lattices follow to
+  establish a sufficiently rich landscape of locales for
+  further examples in this tutorial.  Each comes with an example
+  theorem.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 \ \ \isacommand{locale}\isamarkupfalse%
@@ -1147,12 +1202,13 @@
 \endisadelimproof
 %
 \begin{isamarkuptext}%
-The locale hierachy obtained through these declarations is shown in Figure~\ref{fig:lattices}(a).
+The locale hierarchy obtained through these declarations is shown in
+  Figure~\ref{fig:lattices}(a).
 
 \begin{figure}
 \hrule \vspace{2ex}
 \begin{center}
-\subfigure[Declared hierachy]{
+\subfigure[Declared hierarchy]{
 \begin{tikzpicture}
   \node (po) at (0,0) {\isa{partial{\isacharunderscore}order}};
   \node (lat) at (-1.5,-1) {\isa{lattice}};
@@ -1201,12 +1257,37 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-Total orders are lattices.  Hence, by deriving the lattice
-  axioms for total orders, the hierarchy may be changed
-  and \isa{lattice} be placed between \isa{partial{\isacharunderscore}order}
-  and \isa{total{\isacharunderscore}order}, as shown in Figure~\ref{fig:lattices}(b).
-  Changes to the locale hierarchy may be declared
-  with the \isakeyword{sublocale} command.%
+Locales enable to prove theorems abstractly, relative to
+  sets of assumptions.  These theorems can then be used in other
+  contexts where the assumptions themselves, or
+  instances of the assumptions, are theorems.  This form of theorem
+  reuse is called \emph{interpretation}.  Locales generalise
+  interpretation from theorems to conclusions, enabling the reuse of
+  definitions and other constructs that are not part of the
+  specifications of the locales.
+
+  The first from of interpretation we will consider in this tutorial
+  is provided by the \isakeyword{sublocale} command.  It enables to
+  modify the import hierarchy to reflect the \emph{logical} relation
+  between locales.
+
+  Consider the locale hierarchy from Figure~\ref{fig:lattices}(a).
+  Total orders are lattices, although this is not reflected here, and
+  definitions, theorems and other conclusions
+  from \isa{lattice} are not available in \isa{total{\isacharunderscore}order}.  To
+  obtain the situation in Figure~\ref{fig:lattices}(b), it is
+  sufficient to add the conclusions of the latter locale to the former.
+  The \isakeyword{sublocale} command does exactly this.
+  The declaration \isakeyword{sublocale} $l_1
+  \subseteq l_2$ causes locale $l_2$ to be \emph{interpreted} in the
+  context of $l_1$.  This means that all conclusions of $l_2$ are made
+  available in $l_1$.
+
+  Of course, the change of hierarchy must be supported by a theorem
+  that reflects, in our example, that total orders are indeed
+  lattices.  Therefore the \isakeyword{sublocale} command generates a
+  goal, which must be discharged by the user.  This is illustrated in
+  the following paragraphs.  First the sublocale relation is stated.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -1218,14 +1299,16 @@
 \isacommand{sublocale}\isamarkupfalse%
 \ total{\isacharunderscore}order\ {\isasymsubseteq}\ lattice%
 \begin{isamarkuptxt}%
-This enters the context of locale \isa{total{\isacharunderscore}order}, in
+\normalsize
+  This enters the context of locale \isa{total{\isacharunderscore}order}, in
   which the goal \begin{isabelle}%
 \ {\isadigit{1}}{\isachardot}\ lattice\ op\ {\isasymsqsubseteq}%
-\end{isabelle} must be shown.  First, the
-  locale predicate needs to be unfolded --- for example using its
+\end{isabelle} must be shown.
+  Now the
+  locale predicate needs to be unfolded --- for example, using its
   definition or by introduction rules
-  provided by the locale package.  The methods \isa{intro{\isacharunderscore}locales}
-  and \isa{unfold{\isacharunderscore}locales} automate this.  They are aware of the
+  provided by the locale package.  For automation, the locale package
+  provides the methods \isa{intro{\isacharunderscore}locales} and \isa{unfold{\isacharunderscore}locales}.  They are aware of the
   current context and dependencies between locales and automatically
   discharge goals implied by these.  While \isa{unfold{\isacharunderscore}locales}
   always unfolds locale predicates to assumptions, \isa{intro{\isacharunderscore}locales} only unfolds definitions along the locale
@@ -1234,21 +1317,24 @@
   is smaller.
 
   For the current goal, we would like to get hold of
-  the assumptions of \isa{lattice}, hence \isa{unfold{\isacharunderscore}locales}
-  is appropriate.%
+  the assumptions of \isa{lattice}, which need to be shown, hence
+  \isa{unfold{\isacharunderscore}locales} is appropriate.%
 \end{isamarkuptxt}%
 \isamarkuptrue%
 \ \ \isacommand{proof}\isamarkupfalse%
 \ unfold{\isacharunderscore}locales%
 \begin{isamarkuptxt}%
-Since both \isa{lattice} and \isa{total{\isacharunderscore}order}
-  inherit \isa{partial{\isacharunderscore}order}, the assumptions of the latter are
-  discharged, and the only subgoals that remain are the assumptions
-  introduced in \isa{lattice} \begin{isabelle}%
+\normalsize
+  Since the fact that both lattices and total orders are partial
+  orders is already reflected in the locale hierarchy, the assumptions
+  of \isa{partial{\isacharunderscore}order} are discharged automatically, and only the
+  assumptions introduced in \isa{lattice} remain as subgoals
+  \begin{isabelle}%
 \ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ y{\isachardot}\ {\isasymexists}inf{\isachardot}\ is{\isacharunderscore}inf\ x\ y\ inf\isanewline
 \ {\isadigit{2}}{\isachardot}\ {\isasymAnd}x\ y{\isachardot}\ {\isasymexists}sup{\isachardot}\ is{\isacharunderscore}sup\ x\ y\ sup%
 \end{isabelle}
-  The proof for the first subgoal is%
+  The proof for the first subgoal is obtained by constructing an
+  infimum, whose existence is implied by totality.%
 \end{isamarkuptxt}%
 \isamarkuptrue%
 \ \ \ \ \isacommand{fix}\isamarkupfalse%
@@ -1263,7 +1349,8 @@
 \ {\isachardoublequoteopen}{\isasymexists}inf{\isachardot}\ is{\isacharunderscore}inf\ x\ y\ inf{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 %
 \begin{isamarkuptxt}%
-The proof for the second subgoal is analogous and not
+\normalsize
+   The proof for the second subgoal is analogous and not
   reproduced here.%
 \end{isamarkuptxt}%
 \isamarkuptrue%
@@ -1315,14 +1402,15 @@
 \endisadelimvisible
 %
 \begin{isamarkuptext}%
-Similarly, total orders are distributive lattices.%
+Similarly, we may establish that total orders are distributive
+  lattices with a second \isakeyword{sublocale} statement.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 \ \ \isacommand{sublocale}\isamarkupfalse%
 \ total{\isacharunderscore}order\ {\isasymsubseteq}\ distrib{\isacharunderscore}lattice\isanewline
 %
 \isadelimproof
-\ \ %
+\ \ \ \ %
 \endisadelimproof
 %
 \isatagproof
@@ -1396,7 +1484,22 @@
 \endisadelimproof
 %
 \begin{isamarkuptext}%
-The locale hierarchy is now as shown in Figure~\ref{fig:lattices}(c).%
+The locale hierarchy is now as shown in
+  Figure~\ref{fig:lattices}(c).%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+Locale interpretation is \emph{dynamic}.  The statement
+  \isakeyword{sublocale} $l_1 \subseteq l_2$ will not just add the
+  current conclusions of $l_2$ to $l_1$.  Rather the dependency is
+  stored, and conclusions that will be
+  added to $l_2$ in future are automatically propagated to $l_1$.
+  The sublocale relation is transitive --- that is, propagation takes
+  effect along chains of sublocales.  Even cycles in the sublocale relation are
+  supported, as long as these cycles do not lead to infinite chains.
+  Details are discussed in the technical report \cite{Ballarin2006a}.
+  See also Section~\ref{sec:infinite-chains} of this tutorial.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
--- a/doc-src/Locales/Locales/document/Examples1.tex	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/Locales/Locales/document/Examples1.tex	Wed Oct 21 16:57:57 2009 +0200
@@ -18,38 +18,34 @@
 %
 \endisadelimtheory
 %
-\isamarkupsection{Use of Locales in Theories and Proofs%
+\begin{isamarkuptext}%
+\vspace{-5ex}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsection{Use of Locales in Theories and Proofs
+  \label{sec:interpretation}%
 }
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-Locales enable to prove theorems abstractly, relative to
-  sets of assumptions.  These theorems can then be used in other
-  contexts where the assumptions themselves, or
-  instances of the assumptions, are theorems.  This form of theorem
-  reuse is called \emph{interpretation}.
+Locales can be interpreted in the contexts of theories and
+  structured proofs.  These interpretations are dynamic, too.
+  Conclusions of locales will be propagated to the current theory or
+  the current proof context.%
+\footnote{Strictly speaking, only interpretation in theories is
+  dynamic since it is not possible to change locales or the locale
+  hierarchy from within a proof.}
+  The focus of this section is on
+  interpretation in theories, but we will also encounter
+  interpretations in proofs, in
+  Section~\ref{sec:local-interpretation}.
 
-  The changes of the locale
-  hierarchy from the previous sections are examples of
-  interpretations.  The command \isakeyword{sublocale} $l_1
-  \subseteq l_2$ is said to \emph{interpret} locale $l_2$ in the
-  context of $l_1$.  It causes all theorems of $l_2$ to be made
-  available in $l_1$.  The interpretation is \emph{dynamic}: not only
-  theorems already present in $l_2$ are available in $l_1$.  Theorems
-  that will be added to $l_2$ in future will automatically be
-  propagated to $l_1$.
-
-  Locales can also be interpreted in the contexts of theories and
-  structured proofs.  These interpretations are dynamic, too.
-  Theorems added to locales will be propagated to theories.
-  In this section the interpretation in
-  theories is illustrated; interpretation in proofs is analogous.
-
-  As an example, consider the type of natural numbers \isa{nat}.  The
-  relation \isa{{\isasymle}} is a total order over \isa{nat},
-  divisibility \isa{dvd} is a distributive lattice.  We start with the
-  interpretation that \isa{{\isasymle}} is a partial order.  The facilities of
-  the interpretation command are explored in three versions.%
+  As an example, consider the type of integers \isa{int}.  The
+  relation \isa{op\ {\isasymle}} is a total order over \isa{int}.  We start
+  with the interpretation that \isa{op\ {\isasymle}} is a partial order.  The
+  facilities of the interpretation command are explored gradually in
+  three versions.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -59,10 +55,10 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-In the most basic form, interpretation just replaces the locale
-  parameters by terms.  The following command interprets the locale
-  \isa{partial{\isacharunderscore}order} in the global context of the theory.  The
-  parameter \isa{le} is replaced by \isa{op\ {\isasymle}}.%
+The command \isakeyword{interpretation} is for the interpretation of
+  locale in theories.  In the following example, the parameter of locale
+  \isa{partial{\isacharunderscore}order} is replaced by \isa{op\ {\isasymle}} and the locale instance is interpreted in the current
+  theory.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -72,17 +68,20 @@
 %
 \isatagvisible
 \isacommand{interpretation}\isamarkupfalse%
-\ nat{\isacharcolon}\ partial{\isacharunderscore}order\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}%
+\ int{\isacharcolon}\ partial{\isacharunderscore}order\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ int\ {\isasymRightarrow}\ int\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}%
 \begin{isamarkuptxt}%
-The locale name is succeeded by a \emph{parameter
-  instantiation}.  This is a list of terms, which refer to
-  the parameters in the order of declaration in the locale.  The
-  locale name is preceded by an optional \emph{interpretation
-  qualifier}, here \isa{nat}.
+\normalsize
+  The argument of the command is a simple \emph{locale expression}
+  consisting of the name of the interpreted locale, which is
+  preceded by the qualifier \isa{int{\isacharcolon}} and succeeded by a
+  white-space-separated list of terms, which provide a full
+  instantiation of the locale parameters.  The parameters are referred
+  to by order of declaration, which is also the order in which
+  \isakeyword{print\_locale} outputs them.  The locale has only a
+  single parameter, hence the list of instantiation terms is a
+  singleton.
 
-  The command creates the goal%
-\footnote{Note that \isa{op} binds tighter than functions
-  application: parentheses around \isa{op\ {\isasymle}} are not necessary.}
+  The command creates the goal
   \begin{isabelle}%
 \ {\isadigit{1}}{\isachardot}\ partial{\isacharunderscore}order\ op\ {\isasymle}%
 \end{isabelle} which can be shown easily:%
@@ -98,17 +97,15 @@
 \endisadelimvisible
 %
 \begin{isamarkuptext}%
-Now theorems from the locale are available in the theory,
-  interpreted for natural numbers, for example \isa{nat{\isachardot}trans}: \begin{isabelle}%
+The effect of the command is that instances of all
+  conclusions of the locale are available in the theory, where names
+  are prefixed by the qualifier.  For example, transitivity for \isa{int} is named \isa{int{\isachardot}trans} and is the following
+  theorem:
+  \begin{isabelle}%
 \ \ {\isasymlbrakk}{\isacharquery}x\ {\isasymle}\ {\isacharquery}y{\isacharsemicolon}\ {\isacharquery}y\ {\isasymle}\ {\isacharquery}z{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}x\ {\isasymle}\ {\isacharquery}z%
 \end{isabelle}
-
-  The interpretation qualifier, \isa{nat} in the example, is applied
-  to all names processed by the interpretation.  If a qualifer is
-  given in the \isakeyword{interpretation} command, its use is
-  mandatory when referencing the name.  For example, the above theorem
-  cannot be referred to simply by \isa{trans}.  This prevents
-  unwanted hiding of theorems.%
+  It is not possible to reference this theorem simply as \isa{trans}.  This prevents unwanted hiding of existing theorems of the
+  theory by an interpretation.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -117,16 +114,27 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-The above interpretation also creates the theorem
-  \isa{nat{\isachardot}less{\isacharunderscore}le{\isacharunderscore}trans}: \begin{isabelle}%
+Not only does the above interpretation qualify theorem names.
+  The prefix \isa{int} is applied to all names introduced in locale
+  conclusions including names introduced in definitions.  The
+  qualified name \isa{int{\isachardot}less} is short for
+  the interpretation of the definition, which is \isa{partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}}.
+  Qualified name and expanded form may be used almost
+  interchangeably.%
+\footnote{Since \isa{op\ {\isasymle}} is polymorphic, for \isa{partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}} a
+  more general type will be inferred than for \isa{int{\isachardot}less} which
+  is over type \isa{int}.}
+  The latter is preferred on output, as for example in the theorem
+  \isa{int{\isachardot}less{\isacharunderscore}le{\isacharunderscore}trans}: \begin{isabelle}%
 \ \ {\isasymlbrakk}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharquery}x\ {\isacharquery}y{\isacharsemicolon}\ {\isacharquery}y\ {\isasymle}\ {\isacharquery}z{\isasymrbrakk}\isanewline
 \isaindent{\ \ }{\isasymLongrightarrow}\ partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharquery}x\ {\isacharquery}z%
 \end{isabelle}
-  Here, \isa{partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}}
-  represents the strict order, although \isa{{\isacharless}} is the natural
-  strict order for \isa{nat}.  Interpretation allows to map concepts
-  introduced by definitions in locales to the corresponding
-  concepts of the theory.%
+  Both notations for the strict order are not satisfactory.  The
+  constant \isa{op\ {\isacharless}} is the strict order for \isa{int}.
+  In order to allow for the desired replacement, interpretation
+  accepts \emph{equations} in addition to the parameter instantiation.
+  These follow the locale expression and are indicated with the
+  keyword \isakeyword{where}.  This is the revised interpretation:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
--- a/doc-src/Locales/Locales/document/Examples2.tex	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/Locales/Locales/document/Examples2.tex	Wed Oct 21 16:57:57 2009 +0200
@@ -19,47 +19,45 @@
 \endisadelimtheory
 %
 \begin{isamarkuptext}%
-This is achieved by unfolding suitable equations during
-  interpretation.  These equations are given after the keyword
-  \isakeyword{where} and require proofs.  The revised command
-  that replaces \isa{{\isasymsqsubset}} by \isa{{\isacharless}} is:%
+\vspace{-5ex}%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
 \isadelimvisible
-%
+\ \ %
 \endisadelimvisible
 %
 \isatagvisible
 \isacommand{interpretation}\isamarkupfalse%
-\ nat{\isacharcolon}\ partial{\isacharunderscore}order\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ {\isacharbrackleft}nat{\isacharcomma}\ nat{\isacharbrackright}\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
-\ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\isacommand{proof}\isamarkupfalse%
+\ int{\isacharcolon}\ partial{\isacharunderscore}order\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ {\isacharbrackleft}int{\isacharcomma}\ int{\isacharbrackright}\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}int{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{proof}\isamarkupfalse%
 \ {\isacharminus}%
 \begin{isamarkuptxt}%
-The goals are \begin{isabelle}%
+\normalsize The goals are now:
+      \begin{isabelle}%
 \ {\isadigit{1}}{\isachardot}\ partial{\isacharunderscore}order\ op\ {\isasymle}\isanewline
 \ {\isadigit{2}}{\isachardot}\ partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ x\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}%
 \end{isabelle}
-    The proof that \isa{{\isasymle}} is a partial order is as above.%
+      The proof that~\isa{{\isasymle}} is a partial order is as above.%
 \end{isamarkuptxt}%
 \isamarkuptrue%
-\ \ \isacommand{show}\isamarkupfalse%
-\ {\isachardoublequoteopen}partial{\isacharunderscore}order\ {\isacharparenleft}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{by}\isamarkupfalse%
+\ \ \ \ \isacommand{show}\isamarkupfalse%
+\ {\isachardoublequoteopen}partial{\isacharunderscore}order\ {\isacharparenleft}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ int\ {\isasymRightarrow}\ int\ {\isasymRightarrow}\ bool{\isacharparenright}{\isachardoublequoteclose}\isanewline
+\ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
 \ unfold{\isacharunderscore}locales\ auto%
 \begin{isamarkuptxt}%
-The second goal is shown by unfolding the
-    definition of \isa{partial{\isacharunderscore}order{\isachardot}less}.%
+\normalsize The second goal is shown by unfolding the
+      definition of \isa{partial{\isacharunderscore}order{\isachardot}less}.%
 \end{isamarkuptxt}%
 \isamarkuptrue%
-\ \ \isacommand{show}\isamarkupfalse%
-\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{unfolding}\isamarkupfalse%
+\ \ \ \ \isacommand{show}\isamarkupfalse%
+\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}int{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
+\ \ \ \ \ \ \isacommand{unfolding}\isamarkupfalse%
 \ partial{\isacharunderscore}order{\isachardot}less{\isacharunderscore}def\ {\isacharbrackleft}OF\ {\isacharbackquoteopen}partial{\isacharunderscore}order\ op\ {\isasymle}{\isacharbackquoteclose}{\isacharbrackright}\isanewline
-\ \ \ \ \isacommand{by}\isamarkupfalse%
+\ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
 \ auto\isanewline
-\isacommand{qed}\isamarkupfalse%
+\ \ \isacommand{qed}\isamarkupfalse%
 %
 \endisatagvisible
 {\isafoldvisible}%
@@ -69,8 +67,8 @@
 \endisadelimvisible
 %
 \begin{isamarkuptext}%
-Note that the above proof is not in the context of a locale.
-  Hence, the correct interpretation of \isa{partial{\isacharunderscore}order{\isachardot}less{\isacharunderscore}def} is obtained manually with \isa{OF}.%
+Note that the above proof is not in the context of the
+  interpreted locale.  Hence, the premise of \isa{partial{\isacharunderscore}order{\isachardot}less{\isacharunderscore}def} is discharged manually with \isa{OF}.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
--- a/doc-src/Locales/Locales/document/Examples3.tex	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/Locales/Locales/document/Examples3.tex	Wed Oct 21 16:57:57 2009 +0200
@@ -18,47 +18,54 @@
 %
 \endisadelimtheory
 %
-\isamarkupsubsection{Third Version: Local Interpretation%
+\begin{isamarkuptext}%
+\vspace{-5ex}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsubsection{Third Version: Local Interpretation
+  \label{sec:local-interpretation}%
 }
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-In the above example, the fact that \isa{{\isasymle}} is a partial
-  order for the natural numbers was used in the proof of the
-  second goal.  In general, proofs of the equations may involve
-  theorems implied by the fact the assumptions of the instantiated
-  locale hold for the instantiating structure.  If these theorems have
-  been shown abstractly in the locale they can be made available
-  conveniently in the context through an auxiliary local interpretation (keyword
-  \isakeyword{interpret}).  This interpretation is inside the proof of the global
-  interpretation.  The third revision of the example illustrates this.%
+In the above example, the fact that \isa{op\ {\isasymle}} is a partial
+  order for the integers was used in the second goal to
+  discharge the premise in the definition of \isa{op\ {\isasymsqsubset}}.  In
+  general, proofs of the equations not only may involve definitions
+  from the interpreted locale but arbitrarily complex arguments in the
+  context of the locale.  Therefore is would be convenient to have the
+  interpreted locale conclusions temporary available in the proof.
+  This can be achieved by a locale interpretation in the proof body.
+  The command for local interpretations is \isakeyword{interpret}.  We
+  repeat the example from the previous section to illustrate this.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
 \isadelimvisible
-%
+\ \ %
 \endisadelimvisible
 %
 \isatagvisible
 \isacommand{interpretation}\isamarkupfalse%
-\ nat{\isacharcolon}\ partial{\isacharunderscore}order\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
-\ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\isacommand{proof}\isamarkupfalse%
+\ int{\isacharcolon}\ partial{\isacharunderscore}order\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ int\ {\isasymRightarrow}\ int\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}int{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{proof}\isamarkupfalse%
 \ {\isacharminus}\isanewline
-\ \ \isacommand{show}\isamarkupfalse%
-\ {\isachardoublequoteopen}partial{\isacharunderscore}order\ {\isacharparenleft}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{by}\isamarkupfalse%
+\ \ \ \ \isacommand{show}\isamarkupfalse%
+\ {\isachardoublequoteopen}partial{\isacharunderscore}order\ {\isacharparenleft}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ int\ {\isasymRightarrow}\ int\ {\isasymRightarrow}\ bool{\isacharparenright}{\isachardoublequoteclose}\isanewline
+\ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
 \ unfold{\isacharunderscore}locales\ auto\isanewline
-\ \ \isacommand{then}\isamarkupfalse%
+\ \ \ \ \isacommand{then}\isamarkupfalse%
 \ \isacommand{interpret}\isamarkupfalse%
-\ nat{\isacharcolon}\ partial{\isacharunderscore}order\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ {\isacharbrackleft}nat{\isacharcomma}\ nat{\isacharbrackright}\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse%
+\ int{\isacharcolon}\ partial{\isacharunderscore}order\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ {\isacharbrackleft}int{\isacharcomma}\ int{\isacharbrackright}\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse%
 \isanewline
-\ \ \isacommand{show}\isamarkupfalse%
-\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{unfolding}\isamarkupfalse%
-\ nat{\isachardot}less{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
+\ \ \ \ \isacommand{show}\isamarkupfalse%
+\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}int{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
+\ \ \ \ \ \ \isacommand{unfolding}\isamarkupfalse%
+\ int{\isachardot}less{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
 \ auto\isanewline
-\isacommand{qed}\isamarkupfalse%
+\ \ \isacommand{qed}\isamarkupfalse%
 %
 \endisatagvisible
 {\isafoldvisible}%
@@ -68,15 +75,14 @@
 \endisadelimvisible
 %
 \begin{isamarkuptext}%
-The inner interpretation does not require an elaborate new
-  proof, it is immediate from the preceding fact and proved with
-  ``.''.  It enriches the local proof context by the very theorems
+The inner interpretation is immediate from the preceding fact
+  and proved by assumption (Isar short hand ``.'').  It enriches the
+  local proof context by the theorems
   also obtained in the interpretation from Section~\ref{sec:po-first},
-  and \isa{nat{\isachardot}less{\isacharunderscore}def} may directly be used to unfold the
+  and \isa{int{\isachardot}less{\isacharunderscore}def} may directly be used to unfold the
   definition.  Theorems from the local interpretation disappear after
-  leaving the proof context --- that is, after the closing
-  \isakeyword{qed} --- and are then replaced by those with the desired
-  substitutions of the strict order.%
+  leaving the proof context --- that is, after the succeeding
+  \isakeyword{next} or \isakeyword{qed} statement.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -85,71 +91,80 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-Further interpretations are necessary to reuse theorems from
-  the other locales.  In \isa{lattice} the operations \isa{{\isasymsqinter}} and
-  \isa{{\isasymsqunion}} are substituted by \isa{min} and
-  \isa{max}.  The entire proof for the
-  interpretation is reproduced in order to give an example of a more
-  elaborate interpretation proof.%
+Further interpretations are necessary for
+  the other locales.  In \isa{lattice} the operations~\isa{{\isasymsqinter}}
+  and~\isa{{\isasymsqunion}} are substituted by \isa{min}
+  and \isa{max}.  The entire proof for the
+  interpretation is reproduced to give an example of a more
+  elaborate interpretation proof.  Note that the equations are named
+  so they can be used in a later example.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
 \isadelimvisible
-%
+\ \ %
 \endisadelimvisible
 %
 \isatagvisible
 \isacommand{interpretation}\isamarkupfalse%
-\ nat{\isacharcolon}\ lattice\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
-\ \ \isakeyword{where}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ min\ x\ y{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}lattice{\isachardot}join\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ max\ x\ y{\isachardoublequoteclose}\isanewline
-\isacommand{proof}\isamarkupfalse%
+\ int{\isacharcolon}\ lattice\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ int\ {\isasymRightarrow}\ int\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isakeyword{where}\ int{\isacharunderscore}min{\isacharunderscore}eq{\isacharcolon}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}int{\isacharparenright}\ y\ {\isacharequal}\ min\ x\ y{\isachardoublequoteclose}\isanewline
+\ \ \ \ \ \ \isakeyword{and}\ int{\isacharunderscore}max{\isacharunderscore}eq{\isacharcolon}\ {\isachardoublequoteopen}lattice{\isachardot}join\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}int{\isacharparenright}\ y\ {\isacharequal}\ max\ x\ y{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{proof}\isamarkupfalse%
 \ {\isacharminus}\isanewline
-\ \ \isacommand{show}\isamarkupfalse%
-\ {\isachardoublequoteopen}lattice\ {\isacharparenleft}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}{\isachardoublequoteclose}%
+\ \ \ \ \isacommand{show}\isamarkupfalse%
+\ {\isachardoublequoteopen}lattice\ {\isacharparenleft}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ int\ {\isasymRightarrow}\ int\ {\isasymRightarrow}\ bool{\isacharparenright}{\isachardoublequoteclose}%
 \begin{isamarkuptxt}%
-We have already shown that this is a partial order,%
+\normalsize We have already shown that this is a partial
+	order,%
 \end{isamarkuptxt}%
 \isamarkuptrue%
-\ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
 \ unfold{\isacharunderscore}locales%
 \begin{isamarkuptxt}%
-hence only the lattice axioms remain to be shown: \begin{isabelle}%
+\normalsize hence only the lattice axioms remain to be
+	shown.
+        \begin{isabelle}%
 \ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ y{\isachardot}\ {\isasymexists}inf{\isachardot}\ partial{\isacharunderscore}order{\isachardot}is{\isacharunderscore}inf\ op\ {\isasymle}\ x\ y\ inf\isanewline
 \ {\isadigit{2}}{\isachardot}\ {\isasymAnd}x\ y{\isachardot}\ {\isasymexists}sup{\isachardot}\ partial{\isacharunderscore}order{\isachardot}is{\isacharunderscore}sup\ op\ {\isasymle}\ x\ y\ sup%
-\end{isabelle}  After unfolding \isa{is{\isacharunderscore}inf} and \isa{is{\isacharunderscore}sup},%
+\end{isabelle}
+	By \isa{is{\isacharunderscore}inf} and \isa{is{\isacharunderscore}sup},%
 \end{isamarkuptxt}%
 \isamarkuptrue%
-\ \ \ \ \isacommand{apply}\isamarkupfalse%
-\ {\isacharparenleft}unfold\ nat{\isachardot}is{\isacharunderscore}inf{\isacharunderscore}def\ nat{\isachardot}is{\isacharunderscore}sup{\isacharunderscore}def{\isacharparenright}%
+\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}unfold\ int{\isachardot}is{\isacharunderscore}inf{\isacharunderscore}def\ int{\isachardot}is{\isacharunderscore}sup{\isacharunderscore}def{\isacharparenright}%
 \begin{isamarkuptxt}%
-the goals become \begin{isabelle}%
+\normalsize the goals are transformed to these
+	statements:
+	\begin{isabelle}%
 \ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ y{\isachardot}\ {\isasymexists}inf{\isasymle}x{\isachardot}\ inf\ {\isasymle}\ y\ {\isasymand}\ {\isacharparenleft}{\isasymforall}z{\isachardot}\ z\ {\isasymle}\ x\ {\isasymand}\ z\ {\isasymle}\ y\ {\isasymlongrightarrow}\ z\ {\isasymle}\ inf{\isacharparenright}\isanewline
 \ {\isadigit{2}}{\isachardot}\ {\isasymAnd}x\ y{\isachardot}\ {\isasymexists}sup{\isasymge}x{\isachardot}\ y\ {\isasymle}\ sup\ {\isasymand}\ {\isacharparenleft}{\isasymforall}z{\isachardot}\ x\ {\isasymle}\ z\ {\isasymand}\ y\ {\isasymle}\ z\ {\isasymlongrightarrow}\ sup\ {\isasymle}\ z{\isacharparenright}%
-\end{isabelle} which can be solved
-      by Presburger arithmetic.%
+\end{isabelle}
+	This is Presburger arithmetic, which can be solved by the
+	method \isa{arith}.%
 \end{isamarkuptxt}%
 \isamarkuptrue%
-\ \ \ \ \isacommand{by}\isamarkupfalse%
+\ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
 \ arith{\isacharplus}%
 \begin{isamarkuptxt}%
-In order to show the equations, we put ourselves in a
-    situation where the lattice theorems can be used in a convenient way.%
+\normalsize In order to show the equations, we put ourselves
+      in a situation where the lattice theorems can be used in a
+      convenient way.%
 \end{isamarkuptxt}%
 \isamarkuptrue%
-\ \ \isacommand{then}\isamarkupfalse%
+\ \ \ \ \isacommand{then}\isamarkupfalse%
 \ \isacommand{interpret}\isamarkupfalse%
-\ nat{\isacharcolon}\ lattice\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse%
+\ int{\isacharcolon}\ lattice\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ int\ {\isasymRightarrow}\ int\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse%
 \isanewline
-\ \ \isacommand{show}\isamarkupfalse%
-\ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ min\ x\ y{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{by}\isamarkupfalse%
-\ {\isacharparenleft}bestsimp\ simp{\isacharcolon}\ nat{\isachardot}meet{\isacharunderscore}def\ nat{\isachardot}is{\isacharunderscore}inf{\isacharunderscore}def{\isacharparenright}\isanewline
-\ \ \isacommand{show}\isamarkupfalse%
-\ {\isachardoublequoteopen}lattice{\isachardot}join\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ max\ x\ y{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{by}\isamarkupfalse%
-\ {\isacharparenleft}bestsimp\ simp{\isacharcolon}\ nat{\isachardot}join{\isacharunderscore}def\ nat{\isachardot}is{\isacharunderscore}sup{\isacharunderscore}def{\isacharparenright}\isanewline
-\isacommand{qed}\isamarkupfalse%
+\ \ \ \ \isacommand{show}\isamarkupfalse%
+\ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}int{\isacharparenright}\ y\ {\isacharequal}\ min\ x\ y{\isachardoublequoteclose}\isanewline
+\ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}bestsimp\ simp{\isacharcolon}\ int{\isachardot}meet{\isacharunderscore}def\ int{\isachardot}is{\isacharunderscore}inf{\isacharunderscore}def{\isacharparenright}\isanewline
+\ \ \ \ \isacommand{show}\isamarkupfalse%
+\ {\isachardoublequoteopen}lattice{\isachardot}join\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}int{\isacharparenright}\ y\ {\isacharequal}\ max\ x\ y{\isachardoublequoteclose}\isanewline
+\ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}bestsimp\ simp{\isacharcolon}\ int{\isachardot}join{\isacharunderscore}def\ int{\isachardot}is{\isacharunderscore}sup{\isacharunderscore}def{\isacharparenright}\isanewline
+\ \ \isacommand{qed}\isamarkupfalse%
 %
 \endisatagvisible
 {\isafoldvisible}%
@@ -159,18 +174,19 @@
 \endisadelimvisible
 %
 \begin{isamarkuptext}%
-Next follows that \isa{{\isasymle}} is a total order.%
+Next follows that \isa{op\ {\isasymle}} is a total order, again for
+  the integers.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
 \isadelimvisible
-%
+\ \ %
 \endisadelimvisible
 %
 \isatagvisible
 \isacommand{interpretation}\isamarkupfalse%
-\ nat{\isacharcolon}\ total{\isacharunderscore}order\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
-\ \ \isacommand{by}\isamarkupfalse%
+\ int{\isacharcolon}\ total{\isacharunderscore}order\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ int\ {\isasymRightarrow}\ int\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{by}\isamarkupfalse%
 \ unfold{\isacharunderscore}locales\ arith%
 \endisatagvisible
 {\isafoldvisible}%
@@ -181,261 +197,69 @@
 %
 \begin{isamarkuptext}%
 Theorems that are available in the theory at this point are shown in
-  Table~\ref{tab:nat-lattice}.
+  Table~\ref{tab:int-lattice}.  Two points are worth noting:
 
 \begin{table}
 \hrule
 \vspace{2ex}
 \begin{center}
 \begin{tabular}{l}
-  \isa{nat{\isachardot}less{\isacharunderscore}def} from locale \isa{partial{\isacharunderscore}order}: \\
+  \isa{int{\isachardot}less{\isacharunderscore}def} from locale \isa{partial{\isacharunderscore}order}: \\
   \quad \isa{{\isacharparenleft}{\isacharquery}x\ {\isacharless}\ {\isacharquery}y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isacharquery}x\ {\isasymle}\ {\isacharquery}y\ {\isasymand}\ {\isacharquery}x\ {\isasymnoteq}\ {\isacharquery}y{\isacharparenright}} \\
-  \isa{nat{\isachardot}meet{\isacharunderscore}left} from locale \isa{lattice}: \\
+  \isa{int{\isachardot}meet{\isacharunderscore}left} from locale \isa{lattice}: \\
   \quad \isa{min\ {\isacharquery}x\ {\isacharquery}y\ {\isasymle}\ {\isacharquery}x} \\
-  \isa{nat{\isachardot}join{\isacharunderscore}distr} from locale \isa{distrib{\isacharunderscore}lattice}: \\
+  \isa{int{\isachardot}join{\isacharunderscore}distr} from locale \isa{distrib{\isacharunderscore}lattice}: \\
   \quad \isa{max\ {\isacharquery}x\ {\isacharparenleft}min\ {\isacharquery}y\ {\isacharquery}z{\isacharparenright}\ {\isacharequal}\ min\ {\isacharparenleft}max\ {\isacharquery}x\ {\isacharquery}y{\isacharparenright}\ {\isacharparenleft}max\ {\isacharquery}x\ {\isacharquery}z{\isacharparenright}} \\
-  \isa{nat{\isachardot}less{\isacharunderscore}total} from locale \isa{total{\isacharunderscore}order}: \\
+  \isa{int{\isachardot}less{\isacharunderscore}total} from locale \isa{total{\isacharunderscore}order}: \\
   \quad \isa{{\isacharquery}x\ {\isacharless}\ {\isacharquery}y\ {\isasymor}\ {\isacharquery}x\ {\isacharequal}\ {\isacharquery}y\ {\isasymor}\ {\isacharquery}y\ {\isacharless}\ {\isacharquery}x}
 \end{tabular}
 \end{center}
 \hrule
-\caption{Interpreted theorems for \isa{{\isasymle}} on the natural numbers.}
-\label{tab:nat-lattice}
+\caption{Interpreted theorems for~\isa{{\isasymle}} on the integers.}
+\label{tab:int-lattice}
 \end{table}
 
-  Note that since the locale hierarchy reflects that total orders are
-  distributive lattices, an explicit interpretation of distributive
-  lattices for the order relation on natural numbers is not neccessary.
-
-  Why not push this idea further and just give the last interpretation
-  as a single interpretation instead of the sequence of three?  The
-  reasons for this are twofold:
 \begin{itemize}
 \item
-  Often it is easier to work in an incremental fashion, because later
-  interpretations require theorems provided by earlier
-  interpretations.
+  Locale \isa{distrib{\isacharunderscore}lattice} was also interpreted.  Since the
+  locale hierarchy reflects that total orders are distributive
+  lattices, the interpretation of the latter was inserted
+  automatically with the interpretation of the former.  In general,
+  interpretation traverses the locale hierarchy upwards and interprets
+  all encountered locales, regardless whether imported or proved via
+  the \isakeyword{sublocale} command.  Existing interpretations are
+  skipped avoiding duplicate work.
 \item
-  Assume that a definition is made in some locale $l_1$, and that $l_2$
-  imports $l_1$.  Let an equation for the definition be
-  proved in an interpretation of $l_2$.  The equation will be unfolded
-  in interpretations of theorems added to $l_2$ or below in the import
-  hierarchy, but not for theorems added above $l_2$.
-  Hence, an equation interpreting a definition should always be given in
-  an interpretation of the locale where the definition is made, not in
-  an interpretation of a locale further down the hierarchy.
+  The predicate \isa{op\ {\isacharless}} appears in theorem \isa{int{\isachardot}less{\isacharunderscore}total}
+  although an equation for the replacement of \isa{op\ {\isasymsqsubset}} was only
+  given in the interpretation of \isa{partial{\isacharunderscore}order}.  The
+  interpretation equations are pushed downwards the hierarchy for
+  related interpretations --- that is, for interpretations that share
+  the instances of parameters they have in common.
 \end{itemize}%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
-\isamarkupsubsection{Lattice \isa{dvd} on \isa{nat}%
-}
-\isamarkuptrue%
-%
 \begin{isamarkuptext}%
-Divisibility on the natural numbers is a distributive lattice
-  but not a total order.  Interpretation again proceeds
-  incrementally.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{interpretation}\isamarkupfalse%
-\ nat{\isacharunderscore}dvd{\isacharcolon}\ partial{\isacharunderscore}order\ {\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
-\ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ dvd\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ dvd\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isatagproof
-\isacommand{proof}\isamarkupfalse%
-\ {\isacharminus}\isanewline
-\ \ \isacommand{show}\isamarkupfalse%
-\ {\isachardoublequoteopen}partial{\isacharunderscore}order\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{by}\isamarkupfalse%
-\ unfold{\isacharunderscore}locales\ {\isacharparenleft}auto\ simp{\isacharcolon}\ dvd{\isacharunderscore}def{\isacharparenright}\isanewline
-\ \ \isacommand{then}\isamarkupfalse%
-\ \isacommand{interpret}\isamarkupfalse%
-\ nat{\isacharunderscore}dvd{\isacharcolon}\ partial{\isacharunderscore}order\ {\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse%
-\isanewline
-\ \ \isacommand{show}\isamarkupfalse%
-\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ dvd\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ dvd\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{apply}\isamarkupfalse%
-\ {\isacharparenleft}unfold\ nat{\isacharunderscore}dvd{\isachardot}less{\isacharunderscore}def{\isacharparenright}\isanewline
-\ \ \ \ \isacommand{apply}\isamarkupfalse%
-\ auto\isanewline
-\ \ \ \ \isacommand{done}\isamarkupfalse%
-\isanewline
-\isacommand{qed}\isamarkupfalse%
-%
-\endisatagproof
-{\isafoldproof}%
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\begin{isamarkuptext}%
-Note that in Isabelle/HOL there is no symbol for strict
-  divisibility.  Instead, interpretation substitutes \isa{x\ dvd\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y}.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{interpretation}\isamarkupfalse%
-\ nat{\isacharunderscore}dvd{\isacharcolon}\ lattice\ {\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
-\ \ \isakeyword{where}\ nat{\isacharunderscore}dvd{\isacharunderscore}meet{\isacharunderscore}eq{\isacharcolon}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isacharequal}\ gcd{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isakeyword{and}\ nat{\isacharunderscore}dvd{\isacharunderscore}join{\isacharunderscore}eq{\isacharcolon}\ {\isachardoublequoteopen}lattice{\isachardot}join\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isacharequal}\ lcm{\isachardoublequoteclose}\isanewline
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isatagproof
-\isacommand{proof}\isamarkupfalse%
-\ {\isacharminus}\isanewline
-\ \ \isacommand{show}\isamarkupfalse%
-\ {\isachardoublequoteopen}lattice\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{apply}\isamarkupfalse%
-\ unfold{\isacharunderscore}locales\isanewline
-\ \ \ \ \isacommand{apply}\isamarkupfalse%
-\ {\isacharparenleft}unfold\ nat{\isacharunderscore}dvd{\isachardot}is{\isacharunderscore}inf{\isacharunderscore}def\ nat{\isacharunderscore}dvd{\isachardot}is{\isacharunderscore}sup{\isacharunderscore}def{\isacharparenright}\isanewline
-\ \ \ \ \isacommand{apply}\isamarkupfalse%
-\ {\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}gcd\ x\ y{\isachardoublequoteclose}\ \isakeyword{in}\ exI{\isacharparenright}\isanewline
-\ \ \ \ \isacommand{apply}\isamarkupfalse%
-\ auto\ {\isacharbrackleft}{\isadigit{1}}{\isacharbrackright}\isanewline
-\ \ \ \ \isacommand{apply}\isamarkupfalse%
-\ {\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}lcm\ x\ y{\isachardoublequoteclose}\ \isakeyword{in}\ exI{\isacharparenright}\isanewline
-\ \ \ \ \isacommand{apply}\isamarkupfalse%
-\ {\isacharparenleft}auto\ intro{\isacharcolon}\ lcm{\isacharunderscore}least{\isacharunderscore}nat{\isacharparenright}\isanewline
-\ \ \ \ \isacommand{done}\isamarkupfalse%
-\isanewline
-\ \ \isacommand{then}\isamarkupfalse%
-\ \isacommand{interpret}\isamarkupfalse%
-\ nat{\isacharunderscore}dvd{\isacharcolon}\ lattice\ {\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse%
-\isanewline
-\ \ \isacommand{show}\isamarkupfalse%
-\ {\isachardoublequoteopen}lattice{\isachardot}meet\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isacharequal}\ gcd{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{apply}\isamarkupfalse%
-\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ expand{\isacharunderscore}fun{\isacharunderscore}eq{\isacharparenright}\isanewline
-\ \ \ \ \isacommand{apply}\isamarkupfalse%
-\ {\isacharparenleft}unfold\ nat{\isacharunderscore}dvd{\isachardot}meet{\isacharunderscore}def{\isacharparenright}\isanewline
-\ \ \ \ \isacommand{apply}\isamarkupfalse%
-\ {\isacharparenleft}rule\ the{\isacharunderscore}equality{\isacharparenright}\isanewline
-\ \ \ \ \isacommand{apply}\isamarkupfalse%
-\ {\isacharparenleft}unfold\ nat{\isacharunderscore}dvd{\isachardot}is{\isacharunderscore}inf{\isacharunderscore}def{\isacharparenright}\isanewline
-\ \ \ \ \isacommand{by}\isamarkupfalse%
-\ auto\isanewline
-\ \ \isacommand{show}\isamarkupfalse%
-\ {\isachardoublequoteopen}lattice{\isachardot}join\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isacharequal}\ lcm{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{apply}\isamarkupfalse%
-\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ expand{\isacharunderscore}fun{\isacharunderscore}eq{\isacharparenright}\isanewline
-\ \ \ \ \isacommand{apply}\isamarkupfalse%
-\ {\isacharparenleft}unfold\ nat{\isacharunderscore}dvd{\isachardot}join{\isacharunderscore}def{\isacharparenright}\isanewline
-\ \ \ \ \isacommand{apply}\isamarkupfalse%
-\ {\isacharparenleft}rule\ the{\isacharunderscore}equality{\isacharparenright}\isanewline
-\ \ \ \ \isacommand{apply}\isamarkupfalse%
-\ {\isacharparenleft}unfold\ nat{\isacharunderscore}dvd{\isachardot}is{\isacharunderscore}sup{\isacharunderscore}def{\isacharparenright}\isanewline
-\ \ \ \ \isacommand{apply}\isamarkupfalse%
-\ {\isacharparenleft}auto\ intro{\isacharcolon}\ lcm{\isacharunderscore}least{\isacharunderscore}nat\ iff{\isacharcolon}\ lcm{\isacharunderscore}unique{\isacharunderscore}nat{\isacharparenright}\isanewline
-\ \ \ \ \isacommand{done}\isamarkupfalse%
-\isanewline
-\isacommand{qed}\isamarkupfalse%
-%
-\endisatagproof
-{\isafoldproof}%
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\begin{isamarkuptext}%
-Equations \isa{nat{\isacharunderscore}dvd{\isacharunderscore}meet{\isacharunderscore}eq} and \isa{nat{\isacharunderscore}dvd{\isacharunderscore}join{\isacharunderscore}eq} are used in the main part the subsequent
-  interpretation.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadeliminvisible
-%
-\endisadeliminvisible
-%
-\isataginvisible
-\isacommand{lemma}\isamarkupfalse%
-\ gcd{\isacharunderscore}lcm{\isacharunderscore}distr{\isacharcolon}\isanewline
-\ \ {\isachardoublequoteopen}gcd\ x\ {\isacharparenleft}lcm\ y\ z{\isacharparenright}\ {\isacharequal}\ lcm\ {\isacharparenleft}gcd\ x\ y{\isacharparenright}\ {\isacharparenleft}gcd\ x\ z{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{sorry}\isamarkupfalse%
-%
-\endisataginvisible
-{\isafoldinvisible}%
-%
-\isadeliminvisible
-%
-\endisadeliminvisible
-\isanewline
-%
-\isadelimvisible
-\isanewline
-%
-\endisadelimvisible
-%
-\isatagvisible
-\isacommand{interpretation}\isamarkupfalse%
-\ nat{\isacharunderscore}dvd{\isacharcolon}\isanewline
-\ \ distrib{\isacharunderscore}lattice\ {\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
-\ \ \isacommand{apply}\isamarkupfalse%
-\ unfold{\isacharunderscore}locales%
-\begin{isamarkuptxt}%
-\begin{isabelle}%
-\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ y\ z{\isachardot}\isanewline
-\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }lattice{\isachardot}meet\ op\ dvd\ x\ {\isacharparenleft}lattice{\isachardot}join\ op\ dvd\ y\ z{\isacharparenright}\ {\isacharequal}\isanewline
-\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }lattice{\isachardot}join\ op\ dvd\ {\isacharparenleft}lattice{\isachardot}meet\ op\ dvd\ x\ y{\isacharparenright}\isanewline
-\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ \ }{\isacharparenleft}lattice{\isachardot}meet\ op\ dvd\ x\ z{\isacharparenright}%
-\end{isabelle}%
-\end{isamarkuptxt}%
-\isamarkuptrue%
-\ \ \isacommand{apply}\isamarkupfalse%
-\ {\isacharparenleft}unfold\ nat{\isacharunderscore}dvd{\isacharunderscore}meet{\isacharunderscore}eq\ nat{\isacharunderscore}dvd{\isacharunderscore}join{\isacharunderscore}eq{\isacharparenright}%
-\begin{isamarkuptxt}%
-\begin{isabelle}%
-\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ y\ z{\isachardot}\ gcd\ x\ {\isacharparenleft}lcm\ y\ z{\isacharparenright}\ {\isacharequal}\ lcm\ {\isacharparenleft}gcd\ x\ y{\isacharparenright}\ {\isacharparenleft}gcd\ x\ z{\isacharparenright}%
-\end{isabelle}%
-\end{isamarkuptxt}%
-\isamarkuptrue%
-\ \ \isacommand{apply}\isamarkupfalse%
-\ {\isacharparenleft}rule\ gcd{\isacharunderscore}lcm{\isacharunderscore}distr{\isacharparenright}\ \isacommand{done}\isamarkupfalse%
-%
-\endisatagvisible
-{\isafoldvisible}%
-%
-\isadelimvisible
-%
-\endisadelimvisible
-%
-\begin{isamarkuptext}%
-Theorems that are available in the theory after these
-  interpretations are shown in Table~\ref{tab:nat-dvd-lattice}.
-
-\begin{table}
-\hrule
-\vspace{2ex}
-\begin{center}
-\begin{tabular}{l}
-  \isa{nat{\isacharunderscore}dvd{\isachardot}less{\isacharunderscore}def} from locale \isa{partial{\isacharunderscore}order}: \\
-  \quad \isa{{\isacharparenleft}{\isacharquery}x\ dvd\ {\isacharquery}y\ {\isasymand}\ {\isacharquery}x\ {\isasymnoteq}\ {\isacharquery}y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isacharquery}x\ dvd\ {\isacharquery}y\ {\isasymand}\ {\isacharquery}x\ {\isasymnoteq}\ {\isacharquery}y{\isacharparenright}} \\
-  \isa{nat{\isacharunderscore}dvd{\isachardot}meet{\isacharunderscore}left} from locale \isa{lattice}: \\
-  \quad \isa{gcd\ {\isacharquery}x\ {\isacharquery}y\ dvd\ {\isacharquery}x} \\
-  \isa{nat{\isacharunderscore}dvd{\isachardot}join{\isacharunderscore}distr} from locale \isa{distrib{\isacharunderscore}lattice}: \\
-  \quad \isa{lcm\ {\isacharquery}x\ {\isacharparenleft}gcd\ {\isacharquery}y\ {\isacharquery}z{\isacharparenright}\ {\isacharequal}\ gcd\ {\isacharparenleft}lcm\ {\isacharquery}x\ {\isacharquery}y{\isacharparenright}\ {\isacharparenleft}lcm\ {\isacharquery}x\ {\isacharquery}z{\isacharparenright}} \\
-\end{tabular}
-\end{center}
-\hrule
-\caption{Interpreted theorems for \isa{dvd} on the natural numbers.}
-\label{tab:nat-dvd-lattice}
-\end{table}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-The syntax of the interpretation commands is shown in
-  Table~\ref{tab:commands}.  The grammar refers to
-  \textit{expression}, which stands for a \emph{locale} expression.
-  Locale expressions are discussed in the following section.%
+The interpretations for a locale $n$ within the current
+  theory may be inspected with \isakeyword{print\_interps}~$n$.  This
+  prints the list of instances of $n$, for which interpretations exist.
+  For example, \isakeyword{print\_interps} \isa{partial{\isacharunderscore}order}
+  outputs the following:
+\begin{small}
+\begin{alltt}
+  int! : partial_order "op \(\le\)"
+\end{alltt}
+\end{small}
+  Of course, there is only one interpretation.
+  The interpretation qualifier on the left is decorated with an
+  exclamation point.  This means that it is mandatory.  Qualifiers
+  can either be \emph{mandatory} or \emph{optional}, designated by
+  ``!'' or ``?'' respectively.  Mandatory qualifiers must occur in a
+  name reference while optional ones need not.  Mandatory qualifiers
+  prevent accidental hiding of names, while optional qualifiers can be
+  more convenient to use.  For \isakeyword{interpretation}, the
+  default is ``!''.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -444,110 +268,86 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-A map \isa{{\isasymphi}} between partial orders \isa{{\isasymsqsubseteq}} and \isa{{\isasympreceq}}
+A map~\isa{{\isasymphi}} between partial orders~\isa{{\isasymsqsubseteq}} and~\isa{{\isasympreceq}}
   is called order preserving if \isa{x\ {\isasymsqsubseteq}\ y} implies \isa{{\isasymphi}\ x\ {\isasympreceq}\ {\isasymphi}\ y}.  This situation is more complex than those encountered so
   far: it involves two partial orders, and it is desirable to use the
   existing locale for both.
 
-  Inspecting the grammar of locale commands in
-  Table~\ref{tab:commands} reveals that the import of a locale can be
-  more than just a single locale.  In general, the import is a
-  \emph{locale expression}, which enables to combine locales
-  and instantiate parameters.  A locale expression is a sequence of
-  locale \emph{instances} followed by an optional \isakeyword{for}
-  clause.  Each instance consists of a locale reference, which may be
-  preceded by a qualifer and succeeded by instantiations of the
-  parameters of that locale.  Instantiations may be either positional
-  or through explicit mappings of parameters to arguments.
+  A locale for order preserving maps requires three parameters: \isa{le}~(\isakeyword{infixl}~\isa{{\isasymsqsubseteq}}) and \isa{le{\isacharprime}}~(\isakeyword{infixl}~\isa{{\isasympreceq}}) for the orders and~\isa{{\isasymphi}}
+  for the map.
+
+  In order to reuse the existing locale for partial orders, which has
+  the single parameter~\isa{le}, it must be imported twice, once
+  mapping its parameter to~\isa{le} from the new locale and once
+  to~\isa{le{\isacharprime}}.  This can be achieved with a compound locale
+  expression.
+
+  In general, a locale expression is a sequence of \emph{locale instances}
+  separated by~``$\textbf{+}$'' and followed by a \isakeyword{for}
+  clause.
+  An instance has the following format:
+\begin{quote}
+  \textit{qualifier} \textbf{:} \textit{locale-name}
+  \textit{parameter-instantiation}
+\end{quote}
+  We have already seen locale instances as arguments to
+  \isakeyword{interpretation} in Section~\ref{sec:interpretation}.
+  As before, the qualifier serves to disambiguate names from
+  different instances of the same locale.  While in
+  \isakeyword{interpretation} qualifiers default to mandatory, in
+  import and in the \isakeyword{sublocale} command, they default to
+  optional.
 
-  Using a locale expression, a locale for order
-  preserving maps can be declared in the following way.%
+  Since the parameters~\isa{le} and~\isa{le{\isacharprime}} are to be partial
+  orders, our locale for order preserving maps will import the these
+  instances:
+\begin{small}
+\begin{alltt}
+  le: partial_order le
+  le': partial_order le'
+\end{alltt}
+\end{small}
+  For matter of convenience we choose to name parameter names and
+  qualifiers alike.  This is an arbitrary decision.  Technically, qualifiers
+  and parameters are unrelated.
+
+  Having determined the instances, let us turn to the \isakeyword{for}
+  clause.  It serves to declare locale parameters in the same way as
+  the context element \isakeyword{fixes} does.  Context elements can
+  only occur after the import section, and therefore the parameters
+  referred to in the instances must be declared in the \isakeyword{for}
+  clause.  The \isakeyword{for} clause is also where the syntax of these
+  parameters is declared.
+
+  Two context elements for the map parameter~\isa{{\isasymphi}} and the
+  assumptions that it is order preserving complete the locale
+  declaration.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 \ \ \isacommand{locale}\isamarkupfalse%
 \ order{\isacharunderscore}preserving\ {\isacharequal}\isanewline
 \ \ \ \ le{\isacharcolon}\ partial{\isacharunderscore}order\ le\ {\isacharplus}\ le{\isacharprime}{\isacharcolon}\ partial{\isacharunderscore}order\ le{\isacharprime}\isanewline
 \ \ \ \ \ \ \isakeyword{for}\ le\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqsubseteq}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{and}\ le{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasympreceq}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ {\isacharplus}\isanewline
-\ \ \ \ \isakeyword{fixes}\ {\isasymphi}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isakeyword{fixes}\ {\isasymphi}\isanewline
 \ \ \ \ \isakeyword{assumes}\ hom{\isacharunderscore}le{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y\ {\isasymLongrightarrow}\ {\isasymphi}\ x\ {\isasympreceq}\ {\isasymphi}\ y{\isachardoublequoteclose}%
 \begin{isamarkuptext}%
-The second and third line contain the expression --- two
-  instances of the partial order locale where the parameter is
-  instantiated to \isa{le}
-  and \isa{le{\isacharprime}}, respectively.  The \isakeyword{for} clause consists
-  of parameter declarations and is similar to the context element
-  \isakeyword{fixes}.  The notable difference is that the
-  \isakeyword{for} clause is part of the expression, and only
-  parameters defined in the expression may occur in its instances.
+Here are examples of theorems that are
+  available in the locale:
 
-  Instances define \emph{morphisms} on locales.  Their effect on the
-  parameters is lifted to terms, propositions and theorems in the
-  canonical way,
-  and thus to the assumptions and conclusions of a locale.  The
-  assumption of a locale expression is the conjunction of the
-  assumptions of the instances.  The conclusions of a sequence of
-  instances are obtained by appending the conclusions of the
-  instances in the order of the sequence.
+  \hspace*{1em}\isa{hom{\isacharunderscore}le}: \isa{{\isacharquery}x\ {\isasymsqsubseteq}\ {\isacharquery}y\ {\isasymLongrightarrow}\ {\isasymphi}\ {\isacharquery}x\ {\isasympreceq}\ {\isasymphi}\ {\isacharquery}y}
+
+  \hspace*{1em}\isa{le{\isachardot}less{\isacharunderscore}le{\isacharunderscore}trans}: \isa{{\isasymlbrakk}{\isacharquery}x\ {\isasymsqsubset}\ {\isacharquery}y{\isacharsemicolon}\ {\isacharquery}y\ {\isasymsqsubseteq}\ {\isacharquery}z{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}x\ {\isasymsqsubset}\ {\isacharquery}z}
 
-  The qualifiers in the expression are already a familiar concept from
-  the \isakeyword{interpretation} command
-  (Section~\ref{sec:po-first}).  Here, they serve to distinguish names
-  (in particular theorem names) for the two partial orders within the
-  locale.  Qualifiers in the \isakeyword{locale} command (and in
-  \isakeyword{sublocale}) default to optional --- that is, they need
-  not occur in references to the qualified names.  Here are examples
-  of theorems in locale \isa{order{\isacharunderscore}preserving}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadeliminvisible
-%
-\endisadeliminvisible
-%
-\isataginvisible
-\isacommand{context}\isamarkupfalse%
-\ order{\isacharunderscore}preserving\ \isakeyword{begin}%
-\endisataginvisible
-{\isafoldinvisible}%
-%
-\isadeliminvisible
-%
-\endisadeliminvisible
-%
-\begin{isamarkuptext}%
-\isa{le{\isachardot}less{\isacharunderscore}le{\isacharunderscore}trans}: \isa{{\isasymlbrakk}{\isacharquery}x\ {\isasymsqsubset}\ {\isacharquery}y{\isacharsemicolon}\ {\isacharquery}y\ {\isasymsqsubseteq}\ {\isacharquery}z{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}x\ {\isasymsqsubset}\ {\isacharquery}z}
-
-  \isa{hom{\isacharunderscore}le}: \isa{{\isacharquery}x\ {\isasymsqsubseteq}\ {\isacharquery}y\ {\isasymLongrightarrow}\ {\isasymphi}\ {\isacharquery}x\ {\isasympreceq}\ {\isasymphi}\ {\isacharquery}y}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-The theorems for the partial order \isa{{\isasympreceq}}
-  are qualified by \isa{le{\isacharprime}}.  For example, \isa{le{\isacharprime}{\isachardot}less{\isacharunderscore}le{\isacharunderscore}trans}: \begin{isabelle}%
-\ \ {\isasymlbrakk}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasympreceq}\ {\isacharquery}x\ {\isacharquery}y{\isacharsemicolon}\ {\isacharquery}y\ {\isasympreceq}\ {\isacharquery}z{\isasymrbrakk}\isanewline
-\isaindent{\ \ }{\isasymLongrightarrow}\ partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasympreceq}\ {\isacharquery}x\ {\isacharquery}z%
-\end{isabelle}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadeliminvisible
-%
-\endisadeliminvisible
-%
-\isataginvisible
-\isacommand{end}\isamarkupfalse%
-%
-\endisataginvisible
-{\isafoldinvisible}%
-%
-\isadeliminvisible
-%
-\endisadeliminvisible
-%
-\begin{isamarkuptext}%
-This example reveals that there is no infix syntax for the
-  strict operation associated with \isa{{\isasympreceq}}.  This can be declared
-  through an abbreviation.%
+  \hspace*{1em}\isa{le{\isacharprime}{\isachardot}less{\isacharunderscore}le{\isacharunderscore}trans}:
+  \begin{isabelle}%
+\ \ \ \ {\isasymlbrakk}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasympreceq}\ {\isacharquery}x\ {\isacharquery}y{\isacharsemicolon}\ {\isacharquery}y\ {\isasympreceq}\ {\isacharquery}z{\isasymrbrakk}\isanewline
+\isaindent{\ \ \ \ }{\isasymLongrightarrow}\ partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasympreceq}\ {\isacharquery}x\ {\isacharquery}z%
+\end{isabelle}
+  While there is infix syntax for the strict operation associated to
+  \isa{op\ {\isasymsqsubseteq}}, there is none for the strict version of \isa{op\ {\isasympreceq}}.  The abbreviation \isa{less} with its infix syntax is only
+  available for the original instance it was declared for.  We may
+  introduce the abbreviation \isa{less{\isacharprime}} with infix syntax~\isa{{\isasymprec}}
+  with the following declaration:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 \ \ \isacommand{abbreviation}\isamarkupfalse%
@@ -563,30 +363,63 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-Qualifiers not only apply to theorem names, but also to names
-  introduced by definitions and abbreviations.  For example, in \isa{partial{\isacharunderscore}order} the name \isa{less} abbreviates \isa{op\ {\isasymsqsubset}}.  Therefore, in \isa{order{\isacharunderscore}preserving}
-  the qualified name \isa{le{\isachardot}less} abbreviates \isa{op\ {\isasymsqsubset}} and \isa{le{\isacharprime}{\isachardot}less} abbreviates \isa{op\ {\isasymprec}}.  Hence, the equation in the abbreviation
-  above could have been written more concisely as \isa{less{\isacharprime}\ {\isasymequiv}\ le{\isacharprime}{\isachardot}less}.%
+There are short notations for locale expressions.  These are
+  discussed in the following.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsubsection{Default Instantiations%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+It is possible to omit parameter instantiations.  The
+  instantiation then defaults to the name of
+  the parameter itself.  For example, the locale expression \isa{partial{\isacharunderscore}order} is short for \isa{partial{\isacharunderscore}order\ le}, since the
+  locale's single parameter is~\isa{le}.  We took advantage of this
+  in the \isakeyword{sublocale} declarations of
+  Section~\ref{sec:changing-the-hierarchy}.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsubsection{Implicit Parameters \label{sec:implicit-parameters}%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+In a locale expression that occurs within a locale
+  declaration, omitted parameters additionally extend the (possibly
+  empty) \isakeyword{for} clause.
+
+  The \isakeyword{for} clause is a general construct of Isabelle/Isar
+  to mark names occurring in the preceding declaration as ``arbitrary
+  but fixed''.  This is necessary for example, if the name is already
+  bound in a surrounding context.  In a locale expression, names
+  occurring in parameter instantiations should be bound by a
+  \isakeyword{for} clause whenever these names are not introduced
+  elsewhere in the context --- for example, on the left hand side of a
+  \isakeyword{sublocale} declaration.
+
+  There is an exception to this rule in locale declarations, where the
+  \isakeyword{for} clause serves to declare locale parameters.  Here,
+  locale parameters for which no parameter instantiation is given are
+  implicitly added, with their mixfix syntax, at the beginning of the
+  \isakeyword{for} clause.  For example, in a locale declaration, the
+  expression \isa{partial{\isacharunderscore}order} is short for
+\begin{small}
+\begin{alltt}
+  partial_order le \isakeyword{for} le (\isakeyword{infixl} "\(\sqsubseteq\)" 50)\textrm{.}
+\end{alltt}
+\end{small}
+  This short hand was used in the locale declarations throughout
+  Section~\ref{sec:import}.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-Readers may find the declaration of locale \isa{order{\isacharunderscore}preserving} a little awkward, because the declaration and
-  concrete syntax for \isa{le} from \isa{partial{\isacharunderscore}order} are
-  repeated in the declaration of \isa{order{\isacharunderscore}preserving}.  Locale
-  expressions provide a convenient short hand for this.  A parameter
-  in an instance is \emph{untouched} if no instantiation term is
-  provided for it.  In positional instantiations, a parameter position
-  may be skipped with an underscore, and it is allowed to give fewer
-  instantiation terms than the instantiated locale's number of
-  parameters.  In named instantiations, instantiation pairs for
-  certain parameters may simply be omitted.  Untouched parameters are
-  implicitly declared by the locale expression and with their concrete
-  syntax.  In the sequence of parameters, they appear before the
-  parameters from the \isakeyword{for} clause.
-
-  The following locales illustrate this.  A map \isa{{\isasymphi}} is a
-  lattice homomorphism if it preserves meet and join.%
+The following locale declarations provide more examples.  A
+  map~\isa{{\isasymphi}} is a lattice homomorphism if it preserves meet and
+  join.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 \ \ \isacommand{locale}\isamarkupfalse%
@@ -594,35 +427,50 @@
 \ \ \ \ le{\isacharcolon}\ lattice\ {\isacharplus}\ le{\isacharprime}{\isacharcolon}\ lattice\ le{\isacharprime}\ \isakeyword{for}\ le{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasympreceq}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ {\isacharplus}\isanewline
 \ \ \ \ \isakeyword{fixes}\ {\isasymphi}\isanewline
 \ \ \ \ \isakeyword{assumes}\ hom{\isacharunderscore}meet{\isacharcolon}\ {\isachardoublequoteopen}{\isasymphi}\ {\isacharparenleft}x\ {\isasymsqinter}\ y{\isacharparenright}\ {\isacharequal}\ le{\isacharprime}{\isachardot}meet\ {\isacharparenleft}{\isasymphi}\ x{\isacharparenright}\ {\isacharparenleft}{\isasymphi}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ \ \ \ \ \isakeyword{and}\ hom{\isacharunderscore}join{\isacharcolon}\ {\isachardoublequoteopen}{\isasymphi}\ {\isacharparenleft}x\ {\isasymsqunion}\ y{\isacharparenright}\ {\isacharequal}\ le{\isacharprime}{\isachardot}join\ {\isacharparenleft}{\isasymphi}\ x{\isacharparenright}\ {\isacharparenleft}{\isasymphi}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\isanewline
+\ \ \ \ \ \ \isakeyword{and}\ hom{\isacharunderscore}join{\isacharcolon}\ {\isachardoublequoteopen}{\isasymphi}\ {\isacharparenleft}x\ {\isasymsqunion}\ y{\isacharparenright}\ {\isacharequal}\ le{\isacharprime}{\isachardot}join\ {\isacharparenleft}{\isasymphi}\ x{\isacharparenright}\ {\isacharparenleft}{\isasymphi}\ y{\isacharparenright}{\isachardoublequoteclose}%
+\begin{isamarkuptext}%
+The parameter instantiation in the first instance of \isa{lattice} is omitted.  This causes the parameter~\isa{le} to be
+  added to the \isakeyword{for} clause, and the locale has
+  parameters~\isa{le},~\isa{le{\isacharprime}} and, of course,~\isa{{\isasymphi}}.
+
+  Before turning to the second example, we complete the locale by
+  providing infix syntax for the meet and join operations of the
+  second lattice.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+\ \ \isacommand{context}\isamarkupfalse%
+\ lattice{\isacharunderscore}hom\ \isakeyword{begin}\isanewline
 \ \ \isacommand{abbreviation}\isamarkupfalse%
-\ {\isacharparenleft}\isakeyword{in}\ lattice{\isacharunderscore}hom{\isacharparenright}\isanewline
-\ \ \ \ meet{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqinter}{\isacharprime}{\isacharprime}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\ {\isachardoublequoteopen}meet{\isacharprime}\ {\isasymequiv}\ le{\isacharprime}{\isachardot}meet{\isachardoublequoteclose}\isanewline
+\ meet{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqinter}{\isacharprime}{\isacharprime}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\ {\isachardoublequoteopen}meet{\isacharprime}\ {\isasymequiv}\ le{\isacharprime}{\isachardot}meet{\isachardoublequoteclose}\isanewline
 \ \ \isacommand{abbreviation}\isamarkupfalse%
-\ {\isacharparenleft}\isakeyword{in}\ lattice{\isacharunderscore}hom{\isacharparenright}\isanewline
-\ \ \ \ join{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqunion}{\isacharprime}{\isacharprime}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\ {\isachardoublequoteopen}join{\isacharprime}\ {\isasymequiv}\ le{\isacharprime}{\isachardot}join{\isachardoublequoteclose}%
+\ join{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqunion}{\isacharprime}{\isacharprime}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\ {\isachardoublequoteopen}join{\isacharprime}\ {\isasymequiv}\ le{\isacharprime}{\isachardot}join{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{end}\isamarkupfalse%
+%
 \begin{isamarkuptext}%
-A homomorphism is an endomorphism if both orders coincide.%
+The next example makes radical use of the short hand
+  facilities.  A homomorphism is an endomorphism if both orders
+  coincide.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 \ \ \isacommand{locale}\isamarkupfalse%
 \ lattice{\isacharunderscore}end\ {\isacharequal}\ lattice{\isacharunderscore}hom\ {\isacharunderscore}\ le%
 \begin{isamarkuptext}%
-In this declaration, the first parameter of \isa{lattice{\isacharunderscore}hom}, \isa{le}, is untouched and is then used to instantiate
-  the second parameter.  Its concrete syntax is preserved.%
+The notation~\isa{{\isacharunderscore}} enables to omit a parameter in a
+  positional instantiation.  The omitted parameter,~\isa{le} becomes
+  the parameter of the declared locale and is, in the following
+  position, used to instantiate the second parameter of \isa{lattice{\isacharunderscore}hom}.  The effect is that of identifying the first in second
+  parameter of the homomorphism locale.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
 The inheritance diagram of the situation we have now is shown
   in Figure~\ref{fig:hom}, where the dashed line depicts an
-  interpretation which is introduced below.  Renamings are
-  indicated by $\sqsubseteq \mapsto \preceq$ etc.  The expression
-  imported by \isa{lattice{\isacharunderscore}end} identifies the first and second
-  parameter of \isa{lattice{\isacharunderscore}hom}.  By looking at the inheritance diagram it would seem
-  that two identical copies of each of the locales \isa{partial{\isacharunderscore}order} and \isa{lattice} are imported.  This is not the
-  case!  Inheritance paths with identical morphisms are detected and
+  interpretation which is introduced below.  Parameter instantiations
+  are indicated by $\sqsubseteq \mapsto \preceq$ etc.  By looking at
+  the inheritance diagram it would seem
+  that two identical copies of each of the locales \isa{partial{\isacharunderscore}order} and \isa{lattice} are imported by \isa{lattice{\isacharunderscore}end}.  This is not the case!  Inheritance paths with
+  identical morphisms are automatically detected and
   the conclusions of the respective locales appear only once.
 
 \begin{figure}
@@ -704,9 +552,236 @@
   \isa{hom{\isacharunderscore}le}:
   \begin{isabelle}%
 \ \ {\isacharquery}x\ {\isasymsqsubseteq}\ {\isacharquery}y\ {\isasymLongrightarrow}\ {\isasymphi}\ {\isacharquery}x\ {\isasympreceq}\ {\isasymphi}\ {\isacharquery}y%
-\end{isabelle}%
+\end{isabelle}
+  This theorem will be useful in the following section.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsection{Conditional Interpretation%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+There are situations where an interpretation is not possible
+  in the general case since the desired property is only valid if
+  certain conditions are fulfilled.  Take, for example, the function
+  \isa{{\isasymlambda}i{\isachardot}\ n\ {\isacharasterisk}\ i} that scales its argument by a constant factor.
+  This function is order preserving (and even a lattice endomorphism)
+  with respect to \isa{op\ {\isasymle}} provided \isa{n\ {\isasymge}\ {\isadigit{0}}}.
+
+  It is not possible to express this using a global interpretation,
+  because it is in general unspecified whether~\isa{n} is
+  non-negative, but one may make an interpretation in an inner context
+  of a proof where full information is available.
+  This is not fully satisfactory either, since potentially
+  interpretations may be required to make interpretations in many
+  contexts.  What is
+  required is an interpretation that depends on the condition --- and
+  this can be done with the \isakeyword{sublocale} command.  For this
+  purpose, we introduce a locale for the condition.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+\ \ \isacommand{locale}\isamarkupfalse%
+\ non{\isacharunderscore}negative\ {\isacharequal}\isanewline
+\ \ \ \ \isakeyword{fixes}\ n\ {\isacharcolon}{\isacharcolon}\ int\isanewline
+\ \ \ \ \isakeyword{assumes}\ non{\isacharunderscore}neg{\isacharcolon}\ {\isachardoublequoteopen}{\isadigit{0}}\ {\isasymle}\ n{\isachardoublequoteclose}%
+\begin{isamarkuptext}%
+It is again convenient to make the interpretation in an
+  incremental fashion, first for order preserving maps, the for
+  lattice endomorphisms.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+\ \ \isacommand{sublocale}\isamarkupfalse%
+\ non{\isacharunderscore}negative\ {\isasymsubseteq}\isanewline
+\ \ \ \ \ \ order{\isacharunderscore}preserving\ {\isachardoublequoteopen}op\ {\isasymle}{\isachardoublequoteclose}\ {\isachardoublequoteopen}op\ {\isasymle}{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isasymlambda}i{\isachardot}\ n\ {\isacharasterisk}\ i{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ \ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ non{\isacharunderscore}neg\ \isacommand{by}\isamarkupfalse%
+\ unfold{\isacharunderscore}locales\ {\isacharparenleft}rule\ mult{\isacharunderscore}left{\isacharunderscore}mono{\isacharparenright}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\begin{isamarkuptext}%
+While the proof of the previous interpretation
+  is straightforward from monotonicity lemmas for~\isa{op\ {\isacharasterisk}}, the
+  second proof follows a useful pattern.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimvisible
+\ \ %
+\endisadelimvisible
+%
+\isatagvisible
+\isacommand{sublocale}\isamarkupfalse%
+\ non{\isacharunderscore}negative\ {\isasymsubseteq}\ lattice{\isacharunderscore}end\ {\isachardoublequoteopen}op\ {\isasymle}{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isasymlambda}i{\isachardot}\ n\ {\isacharasterisk}\ i{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{proof}\isamarkupfalse%
+\ {\isacharparenleft}unfold{\isacharunderscore}locales{\isacharcomma}\ unfold\ int{\isacharunderscore}min{\isacharunderscore}eq\ int{\isacharunderscore}max{\isacharunderscore}eq{\isacharparenright}%
+\begin{isamarkuptxt}%
+\normalsize Unfolding the locale predicates \emph{and} the
+      interpretation equations immediately yields two subgoals that
+      reflect the core conjecture.
+      \begin{isabelle}%
+\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ y{\isachardot}\ n\ {\isacharasterisk}\ min\ x\ y\ {\isacharequal}\ min\ {\isacharparenleft}n\ {\isacharasterisk}\ x{\isacharparenright}\ {\isacharparenleft}n\ {\isacharasterisk}\ y{\isacharparenright}\isanewline
+\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}x\ y{\isachardot}\ n\ {\isacharasterisk}\ max\ x\ y\ {\isacharequal}\ max\ {\isacharparenleft}n\ {\isacharasterisk}\ x{\isacharparenright}\ {\isacharparenleft}n\ {\isacharasterisk}\ y{\isacharparenright}%
+\end{isabelle}
+      It is now necessary to show, in the context of \isa{non{\isacharunderscore}negative}, that multiplication by~\isa{n} commutes with
+      \isa{min} and \isa{max}.%
+\end{isamarkuptxt}%
+\isamarkuptrue%
+\ \ \isacommand{qed}\isamarkupfalse%
+\ {\isacharparenleft}auto\ simp{\isacharcolon}\ hom{\isacharunderscore}le{\isacharparenright}%
+\endisatagvisible
+{\isafoldvisible}%
+%
+\isadelimvisible
+%
+\endisadelimvisible
+%
+\begin{isamarkuptext}%
+The lemma \isa{hom{\isacharunderscore}le}
+  simplifies a proof that would have otherwise been lengthy and we may
+  consider making it a default rule for the simplifier:%
 \end{isamarkuptext}%
 \isamarkuptrue%
+\ \ \isacommand{lemmas}\isamarkupfalse%
+\ {\isacharparenleft}\isakeyword{in}\ order{\isacharunderscore}preserving{\isacharparenright}\ hom{\isacharunderscore}le\ {\isacharbrackleft}simp{\isacharbrackright}%
+\isamarkupsubsection{Avoiding Infinite Chains of Interpretations
+  \label{sec:infinite-chains}%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+Similar situations arise frequently in formalisations of
+  abstract algebra where it is desirable to express that certain
+  constructions preserve certain properties.  For example, polynomials
+  over rings are rings, or --- an example from the domain where the
+  illustrations of this tutorial are taken from --- a partial order
+  may be obtained for a function space by point-wise lifting of the
+  partial order of the co-domain.  This corresponds to the following
+  interpretation:%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimvisible
+\ \ %
+\endisadelimvisible
+%
+\isatagvisible
+\isacommand{sublocale}\isamarkupfalse%
+\ partial{\isacharunderscore}order\ {\isasymsubseteq}\ f{\isacharcolon}\ partial{\isacharunderscore}order\ {\isachardoublequoteopen}{\isasymlambda}f\ g{\isachardot}\ {\isasymforall}x{\isachardot}\ f\ x\ {\isasymsqsubseteq}\ g\ x{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{oops}\isamarkupfalse%
+%
+\endisatagvisible
+{\isafoldvisible}%
+%
+\isadelimvisible
+%
+\endisadelimvisible
+%
+\begin{isamarkuptext}%
+Unfortunately this is a cyclic interpretation that leads to an
+  infinite chain, namely
+  \begin{isabelle}%
+\ \ partial{\isacharunderscore}order\ {\isasymsubseteq}\ partial{\isacharunderscore}order\ {\isacharparenleft}{\isasymlambda}f\ g{\isachardot}\ {\isasymforall}x{\isachardot}\ f\ x\ {\isasymsqsubseteq}\ g\ x{\isacharparenright}\ {\isasymsubseteq}\isanewline
+\isaindent{\ \ }\ \ partial{\isacharunderscore}order\ {\isacharparenleft}{\isasymlambda}f\ g{\isachardot}\ {\isasymforall}x\ y{\isachardot}\ f\ x\ y\ {\isasymsqsubseteq}\ g\ x\ y{\isacharparenright}\ {\isasymsubseteq}\ \ {\isasymdots}%
+\end{isabelle}
+  and the interpretation is rejected.
+
+  Instead it is necessary to declare a locale that is logically
+  equivalent to \isa{partial{\isacharunderscore}order} but serves to collect facts
+  about functions spaces where the co-domain is a partial order, and
+  to make the interpretation in its context:%
+\end{isamarkuptext}%
+\isamarkuptrue%
+\ \ \isacommand{locale}\isamarkupfalse%
+\ fun{\isacharunderscore}partial{\isacharunderscore}order\ {\isacharequal}\ partial{\isacharunderscore}order\isanewline
+\isanewline
+\ \ \isacommand{sublocale}\isamarkupfalse%
+\ fun{\isacharunderscore}partial{\isacharunderscore}order\ {\isasymsubseteq}\isanewline
+\ \ \ \ \ \ f{\isacharcolon}\ partial{\isacharunderscore}order\ {\isachardoublequoteopen}{\isasymlambda}f\ g{\isachardot}\ {\isasymforall}x{\isachardot}\ f\ x\ {\isasymsqsubseteq}\ g\ x{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ \ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{by}\isamarkupfalse%
+\ unfold{\isacharunderscore}locales\ {\isacharparenleft}fast{\isacharcomma}rule{\isacharcomma}fast{\isacharcomma}blast\ intro{\isacharcolon}\ trans{\isacharparenright}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\begin{isamarkuptext}%
+It is quite common in abstract algebra that such a construction
+  maps a hierarchy of algebraic structures (or specifications) to a
+  related hierarchy.  By means of the same lifting, a function space
+  is a lattice if its co-domain is a lattice:%
+\end{isamarkuptext}%
+\isamarkuptrue%
+\ \ \isacommand{locale}\isamarkupfalse%
+\ fun{\isacharunderscore}lattice\ {\isacharequal}\ fun{\isacharunderscore}partial{\isacharunderscore}order\ {\isacharplus}\ lattice\isanewline
+\isanewline
+\ \ \isacommand{sublocale}\isamarkupfalse%
+\ fun{\isacharunderscore}lattice\ {\isasymsubseteq}\ f{\isacharcolon}\ lattice\ {\isachardoublequoteopen}{\isasymlambda}f\ g{\isachardot}\ {\isasymforall}x{\isachardot}\ f\ x\ {\isasymsqsubseteq}\ g\ x{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ \ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{proof}\isamarkupfalse%
+\ unfold{\isacharunderscore}locales\isanewline
+\ \ \ \ \isacommand{fix}\isamarkupfalse%
+\ f\ g\isanewline
+\ \ \ \ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}is{\isacharunderscore}inf\ {\isacharparenleft}{\isasymlambda}f\ g{\isachardot}\ {\isasymforall}x{\isachardot}\ f\ x\ {\isasymsqsubseteq}\ g\ x{\isacharparenright}\ f\ g\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ f\ x\ {\isasymsqinter}\ g\ x{\isacharparenright}{\isachardoublequoteclose}\isanewline
+\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}rule\ is{\isacharunderscore}infI{\isacharparenright}\ \isacommand{apply}\isamarkupfalse%
+\ rule{\isacharplus}\ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}drule\ spec{\isacharcomma}\ assumption{\isacharparenright}{\isacharplus}\ \isacommand{done}\isamarkupfalse%
+\isanewline
+\ \ \ \ \isacommand{then}\isamarkupfalse%
+\ \isacommand{show}\isamarkupfalse%
+\ {\isachardoublequoteopen}{\isasymexists}inf{\isachardot}\ partial{\isacharunderscore}order{\isachardot}is{\isacharunderscore}inf\ {\isacharparenleft}{\isasymlambda}f\ g{\isachardot}\ {\isasymforall}x{\isachardot}\ f\ x\ {\isasymsqsubseteq}\ g\ x{\isacharparenright}\ f\ g\ inf{\isachardoublequoteclose}\isanewline
+\ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
+\ fast\isanewline
+\ \ \isacommand{next}\isamarkupfalse%
+\isanewline
+\ \ \ \ \isacommand{fix}\isamarkupfalse%
+\ f\ g\isanewline
+\ \ \ \ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}is{\isacharunderscore}sup\ {\isacharparenleft}{\isasymlambda}f\ g{\isachardot}\ {\isasymforall}x{\isachardot}\ f\ x\ {\isasymsqsubseteq}\ g\ x{\isacharparenright}\ f\ g\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ f\ x\ {\isasymsqunion}\ g\ x{\isacharparenright}{\isachardoublequoteclose}\isanewline
+\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}rule\ is{\isacharunderscore}supI{\isacharparenright}\ \isacommand{apply}\isamarkupfalse%
+\ rule{\isacharplus}\ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}drule\ spec{\isacharcomma}\ assumption{\isacharparenright}{\isacharplus}\ \isacommand{done}\isamarkupfalse%
+\isanewline
+\ \ \ \ \isacommand{then}\isamarkupfalse%
+\ \isacommand{show}\isamarkupfalse%
+\ {\isachardoublequoteopen}{\isasymexists}sup{\isachardot}\ partial{\isacharunderscore}order{\isachardot}is{\isacharunderscore}sup\ {\isacharparenleft}{\isasymlambda}f\ g{\isachardot}\ {\isasymforall}x{\isachardot}\ f\ x\ {\isasymsqsubseteq}\ g\ x{\isacharparenright}\ f\ g\ sup{\isachardoublequoteclose}\isanewline
+\ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
+\ fast\isanewline
+\ \ \isacommand{qed}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
 %
 \isamarkupsection{Further Reading%
 }
@@ -715,13 +790,13 @@
 \begin{isamarkuptext}%
 More information on locales and their interpretation is
   available.  For the locale hierarchy of import and interpretation
-  dependencies see \cite{Ballarin2006a}; interpretations in theories
-  and proofs are covered in \cite{Ballarin2006b}.  In the latter, we
+  dependencies see~\cite{Ballarin2006a}; interpretations in theories
+  and proofs are covered in~\cite{Ballarin2006b}.  In the latter, I
   show how interpretation in proofs enables to reason about families
   of algebraic structures, which cannot be expressed with locales
   directly.
 
-  Haftmann and Wenzel \cite{HaftmannWenzel2007} overcome a restriction
+  Haftmann and Wenzel~\cite{HaftmannWenzel2007} overcome a restriction
   of axiomatic type classes through a combination with locale
   interpretation.  The result is a Haskell-style class system with a
   facility to generate ML and Haskell code.  Classes are sufficient for
@@ -730,10 +805,21 @@
   category.  Order preserving maps, homomorphisms and vector spaces,
   on the other hand, do not.
 
-  The original work of Kamm\"uller on locales \cite{KammullerEtAl1999}
-  may be of interest from a historical perspective.  The mathematical
-  background on orders and lattices is taken from Jacobson's textbook
-  on algebra \cite[Chapter~8]{Jacobson1985}.%
+  The locales reimplementation for Isabelle 2009 provides, among other
+  improvements, a clean integration with Isabelle/Isar's local theory
+  mechanisms, which are described in another paper by Haftmann and
+  Wenzel~\cite{HaftmannWenzel2009}.
+
+  The original work of Kamm\"uller on locales~\cite{KammullerEtAl1999}
+  may be of interest from a historical perspective.  My previous
+  report on locales and locale expressions~\cite{Ballarin2004a}
+  describes a simpler form of expressions than available now and is
+  outdated. The mathematical background on orders and lattices is
+  taken from Jacobson's textbook on algebra~\cite[Chapter~8]{Jacobson1985}.
+
+  The sources of this tutorial, which include all proofs, are
+  available with the Isabelle distribution at
+  \url{http://isabelle.in.tum.de}.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -829,8 +915,9 @@
   \multicolumn{3}{l}{Diagnostics} \\
 
   \textit{toplevel} & ::=
-  & \textbf{print\_locale} [ ``\textbf{!}'' ] \textit{locale} \\
-  & | & \textbf{print\_locales} 
+  & \textbf{print\_locales} \\
+  & | & \textbf{print\_locale} [ ``\textbf{!}'' ] \textit{locale} \\
+  & | & \textbf{print\_interps} \textit{locale}
 \end{tabular}
 \end{center}
 \hrule
@@ -841,9 +928,26 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
+\textbf{Revision History.}  For the present third revision of
+  the tutorial, much of the explanatory text
+  was rewritten.  Inheritance of interpretation equations is
+  available with the forthcoming release of Isabelle, which at the
+  time of editing these notes is expected for the end of 2009.
+  The second revision accommodates changes introduced by the locales
+  reimplementation for Isabelle 2009.  Most notably locale expressions
+  have been generalised from renaming to instantiation.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
 \textbf{Acknowledgements.}  Alexander Krauss, Tobias Nipkow,
-  Christian Sternagel and Makarius Wenzel have made useful comments on
-  a draft of this document.%
+  Randy Pollack, Christian Sternagel and Makarius Wenzel have made
+  useful comments on earlier versions of this document.  The section
+  on conditional interpretation was inspired by a number of e-mail
+  enquiries the author received from locale users, and which suggested
+  that this use case is important enough to deserve explicit
+  explanation.  The term \emph{conditional interpretation} is due to
+  Larry Paulson.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
--- a/doc-src/Locales/Locales/document/root.bib	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/Locales/Locales/document/root.bib	Wed Oct 21 16:57:57 2009 +0200
@@ -42,6 +42,42 @@
   year = 2006
 }
 
+% TYPES 2003
+
+@inproceedings{Ballarin2004a,
+  author = "Clemens Ballarin",
+  title = "Locales and Locale Expressions in {Isabelle/Isar}",
+  pages = "34--50",
+  crossref = "BerardiEtAl2004"
+}
+
+@proceedings{BerardiEtAl2004,
+  editor = "Stefano Berardi and Mario Coppo and Ferruccio Damiani",
+  title = "Types for Proofs and Programs, TYPES 2003, Torino, Italy",
+  booktitle = "Types for Proofs and Programs, TYPES 2003, Torino, Italy",
+  publisher = "Springer",
+  series = "LNCS 3085",
+  year = 2004
+}
+
+% TYPES 2008
+
+@inproceedings{HaftmannWenzel2009,
+  author = "Florian Haftmann and Makarius Wenzel",
+  title = "Local theory specifications in {Isabelle}/{Isar}",
+  pages = "153--168",
+  crossref = "BerardiEtAl2009"
+}
+
+@proceedings{BerardiEtAl2009,
+  editor = "Stefano Berardi and Ferruccio Damiani and Ugo de Liguoro",
+  title = "Types for Proofs and Programs, TYPES 2008, Torino, Italy",
+  booktitle = "Types for Proofs and Programs, TYPES 2008, Torino, Italy",
+  series = "LNCS 5497",
+  publisher = "Springer",
+  year = 2009
+}
+
 % MKM 2006
 
 @inproceedings{Ballarin2006b,
--- a/doc-src/Locales/Locales/document/root.tex	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/Locales/Locales/document/root.tex	Wed Oct 21 16:57:57 2009 +0200
@@ -6,6 +6,7 @@
 \usepackage{subfigure}
 \usepackage{../../../isabelle,../../../isabellesym}
 \usepackage{verbatim}
+\usepackage{alltt}
 \usepackage{array}
 
 \usepackage{amssymb}
@@ -22,26 +23,26 @@
 
 \begin{document}
 
-\title{Tutorial to Locales and Locale Interpretation \\[1ex]
-  \large 2nd revision, for Isabelle 2009}
+\title{Tutorial to Locales and Locale Interpretation}
 \author{Clemens Ballarin}
 \date{}
 
 \maketitle
 
 \begin{abstract}
-  Locales are Isabelle's mechanism for dealing with parametric theories.
-  We present typical examples of locale specifications,
-  along with interpretations between locales to change their
-  hierarchic dependencies and interpretations to reuse locales in
-  theory contexts and proofs.
+  Locales are Isabelle's approach for dealing with parametric
+  theories.  They have been designed as a module system for a
+  theorem prover that can adequately represent the complex
+  inter-dependencies between structures found in abstract algebra, but
+  have proven fruitful also in other applications --- for example,
+  software verification.
 
-  This tutorial is intended for locale novices; familiarity with
-  Isabelle and Isar is presumed.
-  The second revision accommodates changes introduced by the locales
-  reimplementation for Isabelle 2009.  Most notably, in complex
-  specifications (\emph{locale expressions}) renaming has been
-  generalised to instantiation.
+  Both design and implementation of locales have evolved considerably
+  since Kamm\"uller did his initial experiments.  Today, locales
+  are a simple yet powerful extension of the Isar proof language.
+  The present tutorial covers all major facilities of locales.  It is
+  intended for locale novices; familiarity with Isabelle and Isar is
+  presumed.
 \end{abstract}
 
 \parindent 0pt\parskip 0.5ex
--- a/doc-src/Main/Docs/Main_Doc.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/Main/Docs/Main_Doc.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -164,6 +164,21 @@
 \end{tabular}
 
 
+\section{Hilbert\_Choice}
+
+Hilbert's selection ($\varepsilon$) operator: @{term"SOME x. P"}.
+\smallskip
+
+\begin{tabular}{@ {} l @ {~::~} l @ {}}
+@{const Hilbert_Choice.inv_onto} & @{term_type_only Hilbert_Choice.inv_onto "'a set \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'a)"}
+\end{tabular}
+
+\subsubsection*{Syntax}
+
+\begin{tabular}{@ {} l @ {\quad$\equiv$\quad} l @ {}}
+@{term inv} & @{term[source]"inv_onto UNIV"}
+\end{tabular}
+
 \section{Fixed Points}
 
 Theory: @{theory Inductive}.
--- a/doc-src/Main/Docs/document/Main_Doc.tex	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/Main/Docs/document/Main_Doc.tex	Wed Oct 21 16:57:57 2009 +0200
@@ -175,6 +175,21 @@
 \end{tabular}
 
 
+\section{Hilbert\_Choice}
+
+Hilbert's selection ($\varepsilon$) operator: \isa{SOME\ x{\isachardot}\ P}.
+\smallskip
+
+\begin{tabular}{@ {} l @ {~::~} l @ {}}
+\isa{inv{\isacharunderscore}onto} & \isa{{\isacharprime}a\ set\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}b\ {\isasymRightarrow}\ {\isacharprime}a}
+\end{tabular}
+
+\subsubsection*{Syntax}
+
+\begin{tabular}{@ {} l @ {\quad$\equiv$\quad} l @ {}}
+\isa{inv} & \isa{{\isachardoublequote}inv{\isacharunderscore}onto\ UNIV{\isachardoublequote}}
+\end{tabular}
+
 \section{Fixed Points}
 
 Theory: \isa{Inductive}.
--- a/doc-src/TutorialI/Overview/LNCS/Sets.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/TutorialI/Overview/LNCS/Sets.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -149,7 +149,7 @@
     proof (rule converse_rtrancl_induct)
       show "t \<in> lfp ?F"
       proof (subst lfp_unfold[OF mono_ef])
-	show "t \<in> ?F(lfp ?F)" using tA by blast
+        show "t \<in> ?F(lfp ?F)" using tA by blast
       qed
     next
       fix s s'
@@ -157,7 +157,7 @@
          and IH: "s' \<in> lfp ?F"
       show "s \<in> lfp ?F"
       proof (subst lfp_unfold[OF mono_ef])
-	show "s \<in> ?F(lfp ?F)" using prems by blast
+        show "s \<in> ?F(lfp ?F)" using prems by blast
       qed
     qed
   qed
--- a/doc-src/TutorialI/Protocol/Event.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/TutorialI/Protocol/Event.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -22,7 +22,7 @@
         | Notes agent       msg
        
 consts 
-  bad    :: "agent set"				(*compromised agents*)
+  bad    :: "agent set"                         -- {* compromised agents *}
   knows  :: "agent => event list => msg set"
 
 
@@ -43,19 +43,19 @@
   knows_Cons:
     "knows A (ev # evs) =
        (if A = Spy then 
-	(case ev of
-	   Says A' B X => insert X (knows Spy evs)
-	 | Gets A' X => knows Spy evs
-	 | Notes A' X  => 
-	     if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs)
-	else
-	(case ev of
-	   Says A' B X => 
-	     if A'=A then insert X (knows A evs) else knows A evs
-	 | Gets A' X    => 
-	     if A'=A then insert X (knows A evs) else knows A evs
-	 | Notes A' X    => 
-	     if A'=A then insert X (knows A evs) else knows A evs))"
+        (case ev of
+           Says A' B X => insert X (knows Spy evs)
+         | Gets A' X => knows Spy evs
+         | Notes A' X  => 
+             if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs)
+        else
+        (case ev of
+           Says A' B X => 
+             if A'=A then insert X (knows A evs) else knows A evs
+         | Gets A' X    => 
+             if A'=A then insert X (knows A evs) else knows A evs
+         | Notes A' X    => 
+             if A'=A then insert X (knows A evs) else knows A evs))"
 
 (*
   Case A=Spy on the Gets event
@@ -71,10 +71,10 @@
 primrec
   used_Nil:   "used []         = (UN B. parts (initState B))"
   used_Cons:  "used (ev # evs) =
-		     (case ev of
-			Says A B X => parts {X} \<union> used evs
-		      | Gets A X   => used evs
-		      | Notes A X  => parts {X} \<union> used evs)"
+                     (case ev of
+                        Says A B X => parts {X} \<union> used evs
+                      | Gets A X   => used evs
+                      | Notes A X  => parts {X} \<union> used evs)"
     --{*The case for @{term Gets} seems anomalous, but @{term Gets} always
         follows @{term Says} in real protocols.  Seems difficult to change.
         See @{text Gets_correct} in theory @{text "Guard/Extensions.thy"}. *}
--- a/doc-src/TutorialI/Protocol/Message.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/TutorialI/Protocol/Message.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -61,8 +61,8 @@
      msg = Agent  agent
          | Nonce  nat
          | Key    key
-	 | MPair  msg msg
-	 | Crypt  key msg
+         | MPair  msg msg
+         | Crypt  key msg
 
 text {*
 \noindent
@@ -855,8 +855,8 @@
     (Fake_insert_simp_tac ss 1
      THEN
      IF_UNSOLVED (Blast.depth_tac
-		  (cs addIs [analz_insertI,
-				   impOfSubs analz_subset_parts]) 4 1))
+                  (cs addIs [analz_insertI,
+                                   impOfSubs analz_subset_parts]) 4 1))
 
 fun spy_analz_tac (cs,ss) i =
   DETERM
--- a/doc-src/TutorialI/Protocol/NS_Public.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/TutorialI/Protocol/NS_Public.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -221,8 +221,8 @@
 lemma A_trusts_NS2_lemma [rule_format]:
    "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>
     \<Longrightarrow> Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow>
-	Says A B (Crypt(pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
-	Says B A (Crypt(pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
+        Says A B (Crypt(pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
+        Says B A (Crypt(pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
 apply (erule ns_public.induct, simp_all)
 (*Fake, NS1*)
 apply (blast dest: Spy_not_see_NA)+
@@ -240,8 +240,8 @@
 lemma B_trusts_NS1 [rule_format]:
      "evs \<in> ns_public
       \<Longrightarrow> Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow>
-	  Nonce NA \<notin> analz (knows Spy evs) \<longrightarrow>
-	  Says A B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
+          Nonce NA \<notin> analz (knows Spy evs) \<longrightarrow>
+          Says A B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
 apply (erule ns_public.induct, simp_all)
 (*Fake*)
 apply (blast intro!: analz_insertI)
--- a/doc-src/TutorialI/Protocol/Public.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/TutorialI/Protocol/Public.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -25,11 +25,11 @@
 primrec
         (*Agents know their private key and all public keys*)
   initState_Server:  "initState Server     =    
- 		         insert (Key (priK Server)) (Key ` range pubK)"
+                         insert (Key (priK Server)) (Key ` range pubK)"
   initState_Friend:  "initState (Friend i) =    
- 		         insert (Key (priK (Friend i))) (Key ` range pubK)"
+                         insert (Key (priK (Friend i))) (Key ` range pubK)"
   initState_Spy:     "initState Spy        =    
- 		         (Key`invKey`pubK`bad) Un (Key ` range pubK)"
+                         (Key`invKey`pubK`bad) Un (Key ` range pubK)"
 (*>*)
 
 text {*
--- a/doc-src/TutorialI/Rules/Basic.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/TutorialI/Rules/Basic.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -1,4 +1,3 @@
-(* ID:         $Id$ *)
 theory Basic imports Main begin
 
 lemma conj_rule: "\<lbrakk> P; Q \<rbrakk> \<Longrightarrow> P \<and> (Q \<and> P)"
@@ -149,9 +148,9 @@
 
 lemma "\<lbrakk>\<not>(P\<longrightarrow>Q); \<not>(R\<longrightarrow>Q)\<rbrakk> \<Longrightarrow> R"
 apply (erule_tac Q="R\<longrightarrow>Q" in contrapos_np)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 apply (intro impI)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 by (erule notE)
 
 text {*
@@ -161,11 +160,11 @@
 
 lemma "(P \<or> Q) \<and> R \<Longrightarrow> P \<or> Q \<and> R"
 apply (intro disjCI conjI)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 
 apply (elim conjE disjE)
  apply assumption
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 
 by (erule contrapos_np, rule conjI)
 text{*
@@ -241,18 +240,18 @@
 text{*rename_tac*}
 lemma "x < y \<Longrightarrow> \<forall>x y. P x (f y)"
 apply (intro allI)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 apply (rename_tac v w)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 oops
 
 
 lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (h x); P a\<rbrakk> \<Longrightarrow> P(h (h a))"
 apply (frule spec)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 apply (drule mp, assumption)
 apply (drule spec)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 by (drule mp)
 
 lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (f x); P a\<rbrakk> \<Longrightarrow> P(f (f a))"
@@ -276,11 +275,11 @@
 
 lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (h x); P a\<rbrakk> \<Longrightarrow> P(h (h a))"
 apply (frule spec)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 apply (drule mp, assumption)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 apply (drule_tac x = "h a" in spec)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 by (drule mp)
 
 text {*
@@ -290,15 +289,15 @@
 
 lemma mult_dvd_mono: "\<lbrakk>i dvd m; j dvd n\<rbrakk> \<Longrightarrow> i*j dvd (m*n :: nat)"
 apply (simp add: dvd_def)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 apply (erule exE) 
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 apply (erule exE) 
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 apply (rename_tac l)
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 apply (rule_tac x="k*l" in exI) 
-	--{* @{subgoals[display,indent=0,margin=65]} *}
+        --{* @{subgoals[display,indent=0,margin=65]} *}
 apply simp
 done
 
--- a/doc-src/TutorialI/Rules/rules.tex	Wed Oct 21 16:54:04 2009 +0200
+++ b/doc-src/TutorialI/Rules/rules.tex	Wed Oct 21 16:57:57 2009 +0200
@@ -1357,7 +1357,7 @@
 some $x$ such that $P(x)$ is true, provided one exists.
 Isabelle uses \sdx{SOME} for the Greek letter~$\varepsilon$.
 
-Here is the definition of~\cdx{inv}, which expresses inverses of
+Here is the definition of~\cdx{inv},\footnote{In fact, \isa{inv} is defined via a second constant \isa{inv_onto}, which we ignore here.} which expresses inverses of
 functions:
 \begin{isabelle}
 inv\ f\ \isasymequiv \ \isasymlambda y.\ SOME\ x.\ f\ x\ =\ y%
--- a/etc/proofgeneral-settings.el	Wed Oct 21 16:54:04 2009 +0200
+++ b/etc/proofgeneral-settings.el	Wed Oct 21 16:57:57 2009 +0200
@@ -2,6 +2,8 @@
 
 ;; Examples for sensible settings:
 
+(custom-set-variables '(indent-tabs-mode nil))
+
 ;(custom-set-variables '(isar-eta-contract nil))
 
 ;(custom-set-faces
--- a/src/FOL/FOL.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/FOL/FOL.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -174,13 +174,13 @@
   structure Blast = Blast
   (
     val thy = @{theory}
-    type claset	= Cla.claset
+    type claset = Cla.claset
     val equality_name = @{const_name "op ="}
     val not_name = @{const_name Not}
-    val notE	= @{thm notE}
-    val ccontr	= @{thm ccontr}
+    val notE = @{thm notE}
+    val ccontr = @{thm ccontr}
     val contr_tac = Cla.contr_tac
-    val dup_intr	= Cla.dup_intr
+    val dup_intr = Cla.dup_intr
     val hyp_subst_tac = Hypsubst.blast_hyp_subst_tac
     val rep_cs = Cla.rep_cs
     val cla_modifiers = Cla.cla_modifiers
--- a/src/FOL/ex/Classical.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/FOL/ex/Classical.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -418,7 +418,7 @@
 by fast
 
 text{*Halting problem: Formulation of Li Dafa (AAR Newsletter 27, Oct 1994.)
-	author U. Egly*}
+  author U. Egly*}
 lemma "((\<exists>x. A(x) & (\<forall>y. C(y) --> (\<forall>z. D(x,y,z)))) -->                
    (\<exists>w. C(w) & (\<forall>y. C(y) --> (\<forall>z. D(w,y,z)))))                   
   &                                                                      
--- a/src/FOL/simpdata.ML	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/FOL/simpdata.ML	Wed Oct 21 16:57:57 2009 +0200
@@ -27,7 +27,7 @@
 
 (*Congruence rules for = or <-> (instead of ==)*)
 fun mk_meta_cong rl =
-  standard(mk_meta_eq (mk_meta_prems rl))
+  Drule.standard (mk_meta_eq (mk_meta_prems rl))
   handle THM _ =>
   error("Premises and conclusion of congruence rules must use =-equality or <->");
 
--- a/src/FOLP/simp.ML	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/FOLP/simp.ML	Wed Oct 21 16:57:57 2009 +0200
@@ -519,7 +519,7 @@
 (* Compute Congruence rules for individual constants using the substition
    rules *)
 
-val subst_thms = map standard subst_thms;
+val subst_thms = map Drule.standard subst_thms;
 
 
 fun exp_app(0,t) = t
@@ -543,7 +543,7 @@
 fun find_subst sg T =
 let fun find (thm::thms) =
         let val (Const(_,cT), va, vb) = dest_red(hd(prems_of thm));
-            val [P] = OldTerm.add_term_vars(concl_of thm,[]) \\ [va,vb]
+            val [P] = subtract (op =) [va, vb] (OldTerm.add_term_vars (concl_of thm, []));
             val eqT::_ = binder_types cT
         in if Sign.typ_instance sg (T,eqT) then SOME(thm,va,vb,P)
            else find thms
--- a/src/HOL/Algebra/Bij.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Algebra/Bij.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -31,8 +31,8 @@
 
 subsection {*Bijections Form a Group *}
 
-lemma restrict_Inv_Bij: "f \<in> Bij S \<Longrightarrow> (\<lambda>x \<in> S. (Inv S f) x) \<in> Bij S"
-  by (simp add: Bij_def bij_betw_Inv)
+lemma restrict_inv_onto_Bij: "f \<in> Bij S \<Longrightarrow> (\<lambda>x \<in> S. (inv_onto S f) x) \<in> Bij S"
+  by (simp add: Bij_def bij_betw_inv_onto)
 
 lemma id_Bij: "(\<lambda>x\<in>S. x) \<in> Bij S "
   by (auto simp add: Bij_def bij_betw_def inj_on_def)
@@ -41,8 +41,8 @@
   by (auto simp add: Bij_def bij_betw_compose) 
 
 lemma Bij_compose_restrict_eq:
-     "f \<in> Bij S \<Longrightarrow> compose S (restrict (Inv S f) S) f = (\<lambda>x\<in>S. x)"
-  by (simp add: Bij_def compose_Inv_id)
+     "f \<in> Bij S \<Longrightarrow> compose S (restrict (inv_onto S f) S) f = (\<lambda>x\<in>S. x)"
+  by (simp add: Bij_def compose_inv_onto_id)
 
 theorem group_BijGroup: "group (BijGroup S)"
 apply (simp add: BijGroup_def)
@@ -52,22 +52,22 @@
   apply (simp add: compose_Bij)
   apply (blast intro: compose_assoc [symmetric] dest: Bij_imp_funcset)
  apply (simp add: id_Bij Bij_imp_funcset Bij_imp_extensional, simp)
-apply (blast intro: Bij_compose_restrict_eq restrict_Inv_Bij)
+apply (blast intro: Bij_compose_restrict_eq restrict_inv_onto_Bij)
 done
 
 
 subsection{*Automorphisms Form a Group*}
 
-lemma Bij_Inv_mem: "\<lbrakk> f \<in> Bij S;  x \<in> S\<rbrakk> \<Longrightarrow> Inv S f x \<in> S"
-by (simp add: Bij_def bij_betw_def Inv_mem)
+lemma Bij_inv_onto_mem: "\<lbrakk> f \<in> Bij S;  x \<in> S\<rbrakk> \<Longrightarrow> inv_onto S f x \<in> S"
+by (simp add: Bij_def bij_betw_def inv_onto_into)
 
-lemma Bij_Inv_lemma:
+lemma Bij_inv_onto_lemma:
  assumes eq: "\<And>x y. \<lbrakk>x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> h(g x y) = g (h x) (h y)"
  shows "\<lbrakk>h \<in> Bij S;  g \<in> S \<rightarrow> S \<rightarrow> S;  x \<in> S;  y \<in> S\<rbrakk>
-        \<Longrightarrow> Inv S h (g x y) = g (Inv S h x) (Inv S h y)"
+        \<Longrightarrow> inv_onto S h (g x y) = g (inv_onto S h x) (inv_onto S h y)"
 apply (simp add: Bij_def bij_betw_def)
 apply (subgoal_tac "\<exists>x'\<in>S. \<exists>y'\<in>S. x = h x' & y = h y'", clarify)
- apply (simp add: eq [symmetric] Inv_f_f funcset_mem [THEN funcset_mem], blast)
+ apply (simp add: eq [symmetric] inv_f_f funcset_mem [THEN funcset_mem], blast)
 done
 
 
@@ -84,17 +84,17 @@
 lemma (in group) mult_funcset: "mult G \<in> carrier G \<rightarrow> carrier G \<rightarrow> carrier G"
   by (simp add:  Pi_I group.axioms)
 
-lemma (in group) restrict_Inv_hom:
+lemma (in group) restrict_inv_onto_hom:
       "\<lbrakk>h \<in> hom G G; h \<in> Bij (carrier G)\<rbrakk>
-       \<Longrightarrow> restrict (Inv (carrier G) h) (carrier G) \<in> hom G G"
-  by (simp add: hom_def Bij_Inv_mem restrictI mult_funcset
-                group.axioms Bij_Inv_lemma)
+       \<Longrightarrow> restrict (inv_onto (carrier G) h) (carrier G) \<in> hom G G"
+  by (simp add: hom_def Bij_inv_onto_mem restrictI mult_funcset
+                group.axioms Bij_inv_onto_lemma)
 
 lemma inv_BijGroup:
-     "f \<in> Bij S \<Longrightarrow> m_inv (BijGroup S) f = (\<lambda>x \<in> S. (Inv S f) x)"
+     "f \<in> Bij S \<Longrightarrow> m_inv (BijGroup S) f = (\<lambda>x \<in> S. (inv_onto S f) x)"
 apply (rule group.inv_equality)
 apply (rule group_BijGroup)
-apply (simp_all add: BijGroup_def restrict_Inv_Bij Bij_compose_restrict_eq)
+apply (simp_all add:BijGroup_def restrict_inv_onto_Bij Bij_compose_restrict_eq)
 done
 
 lemma (in group) subgroup_auto:
@@ -115,7 +115,7 @@
   assume "x \<in> auto G" 
   thus "inv\<^bsub>BijGroup (carrier G)\<^esub> x \<in> auto G"
     by (simp del: restrict_apply
-             add: inv_BijGroup auto_def restrict_Inv_Bij restrict_Inv_hom)
+        add: inv_BijGroup auto_def restrict_inv_onto_Bij restrict_inv_onto_hom)
 qed
 
 theorem (in group) AutoGroup: "group (AutoGroup G)"
--- a/src/HOL/Algebra/Coset.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Algebra/Coset.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -609,7 +609,7 @@
     proof (simp add: r_congruent_def sym_def, clarify)
       fix x y
       assume [simp]: "x \<in> carrier G" "y \<in> carrier G" 
-	 and "inv x \<otimes> y \<in> H"
+         and "inv x \<otimes> y \<in> H"
       hence "inv (inv x \<otimes> y) \<in> H" by (simp add: m_inv_closed) 
       thus "inv y \<otimes> x \<in> H" by (simp add: inv_mult_group)
     qed
@@ -618,10 +618,10 @@
     proof (simp add: r_congruent_def trans_def, clarify)
       fix x y z
       assume [simp]: "x \<in> carrier G" "y \<in> carrier G" "z \<in> carrier G"
-	 and "inv x \<otimes> y \<in> H" and "inv y \<otimes> z \<in> H"
+         and "inv x \<otimes> y \<in> H" and "inv y \<otimes> z \<in> H"
       hence "(inv x \<otimes> y) \<otimes> (inv y \<otimes> z) \<in> H" by simp
       hence "inv x \<otimes> (y \<otimes> inv y) \<otimes> z \<in> H"
-	by (simp add: m_assoc del: r_inv Units_r_inv) 
+        by (simp add: m_assoc del: r_inv Units_r_inv) 
       thus "inv x \<otimes> z \<in> H" by simp
     qed
   qed
--- a/src/HOL/Algebra/Divisibility.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Algebra/Divisibility.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -2579,33 +2579,33 @@
            by force
 
       from this obtain p'
-	  where "p' \<in> set (as@bs)"
-	  and pp': "p \<sim> p'" by auto
+          where "p' \<in> set (as@bs)"
+          and pp': "p \<sim> p'" by auto
 
       hence "p' \<in> set as \<or> p' \<in> set bs" by simp
       moreover
       {
-	assume p'elem: "p' \<in> set as"
-	with ascarr have [simp]: "p' \<in> carrier G" by fast
-
-	note pp'
-	also from afac p'elem
-	     have "p' divides a" by (rule factors_dividesI) fact+
-	finally
-	     have "p divides a" by simp
+        assume p'elem: "p' \<in> set as"
+        with ascarr have [simp]: "p' \<in> carrier G" by fast
+
+        note pp'
+        also from afac p'elem
+             have "p' divides a" by (rule factors_dividesI) fact+
+        finally
+             have "p divides a" by simp
       }
       moreover
       {
-	assume p'elem: "p' \<in> set bs"
-	with bscarr have [simp]: "p' \<in> carrier G" by fast
-
-	note pp'
-	also from bfac
-	     have "p' divides b" by (rule factors_dividesI) fact+
-	finally have "p divides b" by simp
+        assume p'elem: "p' \<in> set bs"
+        with bscarr have [simp]: "p' \<in> carrier G" by fast
+
+        note pp'
+        also from bfac
+             have "p' divides b" by (rule factors_dividesI) fact+
+        finally have "p divides b" by simp
       }
       ultimately
-	  show "p divides a \<or> p divides b" by fast
+          show "p divides a \<or> p divides b" by fast
     qed
   qed
 qed
@@ -3176,7 +3176,7 @@
   have "c = c \<otimes> \<one>" by simp
   also from abrelprime[symmetric]
        have "\<dots> \<sim> c \<otimes> somegcd G a b"
-	 by (rule assoc_subst) (simp add: mult_cong_r)+
+         by (rule assoc_subst) (simp add: mult_cong_r)+
   also have "\<dots> \<sim> somegcd G (c \<otimes> a) (c \<otimes> b)" by (rule gcd_mult) fact+
   finally
        have c: "c \<sim> somegcd G (c \<otimes> a) (c \<otimes> b)" by simp
@@ -3188,13 +3188,13 @@
   have "somegcd G a (b \<otimes> c) \<sim> somegcd G a (c \<otimes> b)" by (simp add: m_comm)
   also from a
        have "\<dots> \<sim> somegcd G (somegcd G a (c \<otimes> a)) (c \<otimes> b)"
-	 by (rule assoc_subst) (simp add: gcd_cong_l)+
+         by (rule assoc_subst) (simp add: gcd_cong_l)+
   also from gcd_assoc
        have "\<dots> \<sim> somegcd G a (somegcd G (c \<otimes> a) (c \<otimes> b))"
        by (rule assoc_subst) simp+
   also from c[symmetric]
        have "\<dots> \<sim> somegcd G a c"
-	 by (rule assoc_subst) (simp add: gcd_cong_r)+
+         by (rule assoc_subst) (simp add: gcd_cong_r)+
   also note acrelprime
   finally
        show "somegcd G a (b \<otimes> c) \<sim> \<one>" by simp
--- a/src/HOL/Algebra/FiniteProduct.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Algebra/FiniteProduct.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -423,17 +423,17 @@
       then have "finprod G f A = finprod G f (insert x B)" by simp
       also from insert have "... = f x \<otimes> finprod G f B"
       proof (intro finprod_insert)
-	show "finite B" by fact
+        show "finite B" by fact
       next
-	show "x ~: B" by fact
+        show "x ~: B" by fact
       next
-	assume "x ~: B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i"
-	  "g \<in> insert x B \<rightarrow> carrier G"
-	thus "f \<in> B -> carrier G" by fastsimp
+        assume "x ~: B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i"
+          "g \<in> insert x B \<rightarrow> carrier G"
+        thus "f \<in> B -> carrier G" by fastsimp
       next
-	assume "x ~: B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i"
-	  "g \<in> insert x B \<rightarrow> carrier G"
-	thus "f x \<in> carrier G" by fastsimp
+        assume "x ~: B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i"
+          "g \<in> insert x B \<rightarrow> carrier G"
+        thus "f x \<in> carrier G" by fastsimp
       qed
       also from insert have "... = g x \<otimes> finprod G g B" by fastsimp
       also from insert have "... = finprod G g (insert x B)"
--- a/src/HOL/Algebra/Group.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Algebra/Group.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -553,11 +553,11 @@
 by (simp add: iso_def hom_def inj_on_def bij_betw_def Pi_def)
 
 lemma (in group) iso_sym:
-     "h \<in> G \<cong> H \<Longrightarrow> Inv (carrier G) h \<in> H \<cong> G"
-apply (simp add: iso_def bij_betw_Inv) 
-apply (subgoal_tac "Inv (carrier G) h \<in> carrier H \<rightarrow> carrier G") 
- prefer 2 apply (simp add: bij_betw_imp_funcset [OF bij_betw_Inv]) 
-apply (simp add: hom_def bij_betw_def Inv_f_eq f_Inv_f Pi_def)
+     "h \<in> G \<cong> H \<Longrightarrow> inv_onto (carrier G) h \<in> H \<cong> G"
+apply (simp add: iso_def bij_betw_inv_onto) 
+apply (subgoal_tac "inv_onto (carrier G) h \<in> carrier H \<rightarrow> carrier G") 
+ prefer 2 apply (simp add: bij_betw_imp_funcset [OF bij_betw_inv_onto]) 
+apply (simp add: hom_def bij_betw_def inv_onto_f_eq f_inv_onto_f Pi_def)
 done
 
 lemma (in group) iso_trans: 
@@ -785,16 +785,16 @@
       assume H: "H \<in> A"
       with L have subgroupH: "subgroup H G" by auto
       from subgroupH have groupH: "group (G (| carrier := H |))" (is "group ?H")
-	by (rule subgroup_imp_group)
+        by (rule subgroup_imp_group)
       from groupH have monoidH: "monoid ?H"
-	by (rule group.is_monoid)
+        by (rule group.is_monoid)
       from H have Int_subset: "?Int \<subseteq> H" by fastsimp
       then show "le ?L ?Int H" by simp
     next
       fix H
       assume H: "H \<in> Lower ?L A"
       with L Int_subgroup show "le ?L H ?Int"
-	by (fastsimp simp: Lower_def intro: Inter_greatest)
+        by (fastsimp simp: Lower_def intro: Inter_greatest)
     next
       show "A \<subseteq> carrier ?L" by (rule L)
     next
--- a/src/HOL/Algebra/Lattice.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Algebra/Lattice.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -533,22 +533,22 @@
       assume y: "y \<in> Upper L (insert x A)"
       show "s \<sqsubseteq> y"
       proof (rule least_le [OF least_s], rule Upper_memI)
-	fix z
-	assume z: "z \<in> {a, x}"
-	then show "z \<sqsubseteq> y"
-	proof
+        fix z
+        assume z: "z \<in> {a, x}"
+        then show "z \<sqsubseteq> y"
+        proof
           have y': "y \<in> Upper L A"
             apply (rule subsetD [where A = "Upper L (insert x A)"])
              apply (rule Upper_antimono)
-	     apply blast
-	    apply (rule y)
+             apply blast
+            apply (rule y)
             done
           assume "z = a"
           with y' least_a show ?thesis by (fast dest: least_le)
-	next
-	  assume "z \<in> {x}"  (* FIXME "z = x"; declare specific elim rule for "insert x {}" (!?) *)
+        next
+          assume "z \<in> {x}"  (* FIXME "z = x"; declare specific elim rule for "insert x {}" (!?) *)
           with y L show ?thesis by blast
-	qed
+        qed
       qed (rule Upper_closed [THEN subsetD, OF y])
     next
       from L show "insert x A \<subseteq> carrier L" by simp
@@ -569,9 +569,9 @@
     case True
     with insert show ?thesis
       by simp (simp add: least_cong [OF weak_sup_of_singleton]
-	sup_of_singleton_closed sup_of_singletonI)
-	(* The above step is hairy; least_cong can make simp loop.
-	Would want special version of simp to apply least_cong. *)
+        sup_of_singleton_closed sup_of_singletonI)
+        (* The above step is hairy; least_cong can make simp loop.
+        Would want special version of simp to apply least_cong. *)
   next
     case False
     with insert have "least L (\<Squnion>A) (Upper L A)" by simp
@@ -774,22 +774,22 @@
       assume y: "y \<in> Lower L (insert x A)"
       show "y \<sqsubseteq> i"
       proof (rule greatest_le [OF greatest_i], rule Lower_memI)
-	fix z
-	assume z: "z \<in> {a, x}"
-	then show "y \<sqsubseteq> z"
-	proof
+        fix z
+        assume z: "z \<in> {a, x}"
+        then show "y \<sqsubseteq> z"
+        proof
           have y': "y \<in> Lower L A"
             apply (rule subsetD [where A = "Lower L (insert x A)"])
             apply (rule Lower_antimono)
-	     apply blast
-	    apply (rule y)
+             apply blast
+            apply (rule y)
             done
           assume "z = a"
           with y' greatest_a show ?thesis by (fast dest: greatest_le)
-	next
+        next
           assume "z \<in> {x}"
           with y L show ?thesis by blast
-	qed
+        qed
       qed (rule Lower_closed [THEN subsetD, OF y])
     next
       from L show "insert x A \<subseteq> carrier L" by simp
@@ -809,7 +809,7 @@
     case True
     with insert show ?thesis
       by simp (simp add: greatest_cong [OF weak_inf_of_singleton]
-	inf_of_singleton_closed inf_of_singletonI)
+        inf_of_singleton_closed inf_of_singletonI)
   next
     case False
     from insert show ?thesis
@@ -1291,7 +1291,7 @@
   proof
     from B show "greatest ?L (\<Inter> B \<inter> A) (Lower ?L B)"
       txt {* @{term "\<Inter> B"} is not the infimum of @{term B}:
-	@{term "\<Inter> {} = UNIV"} which is in general bigger than @{term "A"}! *}
+        @{term "\<Inter> {} = UNIV"} which is in general bigger than @{term "A"}! *}
       by (fastsimp intro!: greatest_LowerI simp: Lower_def)
   qed
 qed
--- a/src/HOL/Algebra/UnivPoly.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Algebra/UnivPoly.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -349,19 +349,19 @@
       fix nn assume Succ: "n = Suc nn"
       have "coeff P (p \<otimes>\<^bsub>P\<^esub> \<one>\<^bsub>P\<^esub>) (Suc nn) = coeff P p (Suc nn)"
       proof -
-	have "coeff P (p \<otimes>\<^bsub>P\<^esub> \<one>\<^bsub>P\<^esub>) (Suc nn) = (\<Oplus>i\<in>{..Suc nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" using R by simp
-	also have "\<dots> = coeff P p (Suc nn) \<otimes> (if Suc nn \<le> Suc nn then \<one> else \<zero>) \<oplus> (\<Oplus>i\<in>{..nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))"
-	  using finsum_Suc [of "(\<lambda>i::nat. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" "nn"] unfolding Pi_def using R by simp
-	also have "\<dots> = coeff P p (Suc nn) \<otimes> (if Suc nn \<le> Suc nn then \<one> else \<zero>)"
-	proof -
-	  have "(\<Oplus>i\<in>{..nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>)) = (\<Oplus>i\<in>{..nn}. \<zero>)"
-	    using finsum_cong [of "{..nn}" "{..nn}" "(\<lambda>i::nat. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" "(\<lambda>i::nat. \<zero>)"] using R 
-	    unfolding Pi_def by simp
-	  also have "\<dots> = \<zero>" by simp
-	  finally show ?thesis using r_zero R by simp
-	qed
-	also have "\<dots> = coeff P p (Suc nn)" using R by simp
-	finally show ?thesis by simp
+        have "coeff P (p \<otimes>\<^bsub>P\<^esub> \<one>\<^bsub>P\<^esub>) (Suc nn) = (\<Oplus>i\<in>{..Suc nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" using R by simp
+        also have "\<dots> = coeff P p (Suc nn) \<otimes> (if Suc nn \<le> Suc nn then \<one> else \<zero>) \<oplus> (\<Oplus>i\<in>{..nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))"
+          using finsum_Suc [of "(\<lambda>i::nat. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" "nn"] unfolding Pi_def using R by simp
+        also have "\<dots> = coeff P p (Suc nn) \<otimes> (if Suc nn \<le> Suc nn then \<one> else \<zero>)"
+        proof -
+          have "(\<Oplus>i\<in>{..nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>)) = (\<Oplus>i\<in>{..nn}. \<zero>)"
+            using finsum_cong [of "{..nn}" "{..nn}" "(\<lambda>i::nat. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" "(\<lambda>i::nat. \<zero>)"] using R 
+            unfolding Pi_def by simp
+          also have "\<dots> = \<zero>" by simp
+          finally show ?thesis using r_zero R by simp
+        qed
+        also have "\<dots> = coeff P p (Suc nn)" using R by simp
+        finally show ?thesis by simp
       qed
       then show ?thesis using Succ by simp
     }
@@ -627,11 +627,11 @@
     then show "monom P \<one> (Suc (Suc k)) = monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> (Suc k)"
     proof -
       have lhs: "monom P \<one> (Suc (Suc k)) = monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> k \<otimes>\<^bsub>P\<^esub> monom P \<one> 1"
-	unfolding monom_one_Suc [of "Suc k"] unfolding hypo ..
+        unfolding monom_one_Suc [of "Suc k"] unfolding hypo ..
       note cl = monom_closed [OF R.one_closed, of 1]
       note clk = monom_closed [OF R.one_closed, of k]
       have rhs: "monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> (Suc k) = monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> k \<otimes>\<^bsub>P\<^esub> monom P \<one> 1"
-	unfolding monom_one_Suc [of k] unfolding sym [OF m_assoc  [OF cl clk cl]] ..
+        unfolding monom_one_Suc [of k] unfolding sym [OF m_assoc  [OF cl clk cl]] ..
       from lhs rhs show ?thesis by simp
     qed
   }
@@ -670,25 +670,25 @@
     case True 
     {
       show ?thesis
-	unfolding True [symmetric]
-	  coeff_mult [OF monom_closed [OF a_in_R, of n] monom_closed [OF b_in_R, of m], of "n + m"] 
-	  coeff_monom [OF a_in_R, of n] coeff_monom [OF b_in_R, of m]
-	using R.finsum_cong [of "{.. n + m}" "{.. n + m}" "(\<lambda>i. (if n = i then a else \<zero>) \<otimes> (if m = n + m - i then b else \<zero>))" 
-	  "(\<lambda>i. if n = i then a \<otimes> b else \<zero>)"]
-	  a_in_R b_in_R
-	unfolding simp_implies_def
-	using R.finsum_singleton [of n "{.. n + m}" "(\<lambda>i. a \<otimes> b)"]
-	unfolding Pi_def by auto
+        unfolding True [symmetric]
+          coeff_mult [OF monom_closed [OF a_in_R, of n] monom_closed [OF b_in_R, of m], of "n + m"] 
+          coeff_monom [OF a_in_R, of n] coeff_monom [OF b_in_R, of m]
+        using R.finsum_cong [of "{.. n + m}" "{.. n + m}" "(\<lambda>i. (if n = i then a else \<zero>) \<otimes> (if m = n + m - i then b else \<zero>))" 
+          "(\<lambda>i. if n = i then a \<otimes> b else \<zero>)"]
+          a_in_R b_in_R
+        unfolding simp_implies_def
+        using R.finsum_singleton [of n "{.. n + m}" "(\<lambda>i. a \<otimes> b)"]
+        unfolding Pi_def by auto
     }
   next
     case False
     {
       show ?thesis
-	unfolding coeff_monom [OF R.m_closed [OF a_in_R b_in_R], of "n + m" k] apply (simp add: False)
-	unfolding coeff_mult [OF monom_closed [OF a_in_R, of n] monom_closed [OF b_in_R, of m], of k]
-	unfolding coeff_monom [OF a_in_R, of n] unfolding coeff_monom [OF b_in_R, of m] using False
-	using R.finsum_cong [of "{..k}" "{..k}" "(\<lambda>i. (if n = i then a else \<zero>) \<otimes> (if m = k - i then b else \<zero>))" "(\<lambda>i. \<zero>)"]
-	unfolding Pi_def simp_implies_def using a_in_R b_in_R by force
+        unfolding coeff_monom [OF R.m_closed [OF a_in_R b_in_R], of "n + m" k] apply (simp add: False)
+        unfolding coeff_mult [OF monom_closed [OF a_in_R, of n] monom_closed [OF b_in_R, of m], of k]
+        unfolding coeff_monom [OF a_in_R, of n] unfolding coeff_monom [OF b_in_R, of m] using False
+        using R.finsum_cong [of "{..k}" "{..k}" "(\<lambda>i. (if n = i then a else \<zero>) \<otimes> (if m = k - i then b else \<zero>))" "(\<lambda>i. \<zero>)"]
+        unfolding Pi_def simp_implies_def using a_in_R b_in_R by force
     }
   qed
 qed (simp_all add: a_in_R b_in_R)
@@ -1517,7 +1517,7 @@
       then have max_sl: "max (deg R p) (deg R q) < m" by simp
       then have "deg R (p \<oplus>\<^bsub>P\<^esub> q) < m" using deg_add [OF p_in_P q_in_P] by arith
       with deg_R_p deg_R_q show ?thesis using coeff_add [OF p_in_P q_in_P, of m]
-	using deg_aboveD [of "p \<oplus>\<^bsub>P\<^esub> q" m] using p_in_P q_in_P by simp 
+        using deg_aboveD [of "p \<oplus>\<^bsub>P\<^esub> q" m] using p_in_P q_in_P by simp 
     qed
   qed (simp add: p_in_P q_in_P)
   moreover have deg_ne: "deg R (p \<oplus>\<^bsub>P\<^esub> q) \<noteq> deg R r"
@@ -1582,114 +1582,114 @@
       (*JE: we now apply the induction hypothesis with some additional facts required*)
       from f_in_P deg_g_le_deg_f show ?thesis
       proof (induct n \<equiv> "deg R f" arbitrary: "f" rule: nat_less_induct)
-	fix n f
-	assume hypo: "\<forall>m<n. \<forall>x. x \<in> carrier P \<longrightarrow>
+        fix n f
+        assume hypo: "\<forall>m<n. \<forall>x. x \<in> carrier P \<longrightarrow>
           deg R g \<le> deg R x \<longrightarrow> 
-	  m = deg R x \<longrightarrow>
-	  (\<exists>q r (k::nat). q \<in> carrier P \<and> r \<in> carrier P \<and> lcoeff g (^) k \<odot>\<^bsub>P\<^esub> x = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r & (r = \<zero>\<^bsub>P\<^esub> | deg R r < deg R g))"
-	  and prem: "n = deg R f" and f_in_P [simp]: "f \<in> carrier P"
-	  and deg_g_le_deg_f: "deg R g \<le> deg R f"
-	let ?k = "1::nat" and ?r = "(g \<otimes>\<^bsub>P\<^esub> (monom P (lcoeff f) (deg R f - deg R g))) \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)"
-	  and ?q = "monom P (lcoeff f) (deg R f - deg R g)"
-	show "\<exists> q r (k::nat). q \<in> carrier P \<and> r \<in> carrier P \<and> lcoeff g (^) k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r & (r = \<zero>\<^bsub>P\<^esub> | deg R r < deg R g)"
-	proof -
-	  (*JE: we first extablish the existence of a triple satisfying the previous equation. 
-	    Then we will have to prove the second part of the predicate.*)
-	  have exist: "lcoeff g (^) ?k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r"
-	    using minus_add
-	    using sym [OF a_assoc [of "g \<otimes>\<^bsub>P\<^esub> ?q" "\<ominus>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q)" "lcoeff g \<odot>\<^bsub>P\<^esub> f"]]
-	    using r_neg by auto
-	  show ?thesis
-	  proof (cases "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < deg R g")
-	    (*JE: if the degree of the remainder satisfies the statement property we are done*)
-	    case True
-	    {
-	      show ?thesis
-	      proof (rule exI3 [of _ ?q "\<ominus>\<^bsub>P\<^esub> ?r" ?k], intro conjI)
-		show "lcoeff g (^) ?k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r" using exist by simp
-		show "\<ominus>\<^bsub>P\<^esub> ?r = \<zero>\<^bsub>P\<^esub> \<or> deg R (\<ominus>\<^bsub>P\<^esub> ?r) < deg R g" using True by simp
-	      qed (simp_all)
-	    }
-	  next
-	    case False note n_deg_r_l_deg_g = False
-	    {
-	      (*JE: otherwise, we verify the conditions of the induction hypothesis.*)
-	      show ?thesis
-	      proof (cases "deg R f = 0")
-		(*JE: the solutions are different if the degree of f is zero or not*)
-		case True
-		{
-		  have deg_g: "deg R g = 0" using True using deg_g_le_deg_f by simp
-		  have "lcoeff g (^) (1::nat) \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> f \<oplus>\<^bsub>P\<^esub> \<zero>\<^bsub>P\<^esub>"
-		    unfolding deg_g apply simp
-		    unfolding sym [OF monom_mult_is_smult [OF coeff_closed [OF g_in_P, of 0] f_in_P]]
-		    using deg_zero_impl_monom [OF g_in_P deg_g] by simp
-		  then show ?thesis using f_in_P by blast
-		}
-	      next
-		case False note deg_f_nzero = False
-		{
-		  (*JE: now it only remains the case where the induction hypothesis can be used.*)
-		  (*JE: we first prove that the degree of the remainder is smaller than the one of f*)
-		  have deg_remainder_l_f: "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < n"
-		  proof -
-		    have "deg R (\<ominus>\<^bsub>P\<^esub> ?r) = deg R ?r" using deg_uminus [of ?r] by simp
-		    also have "\<dots> < deg R f"
-		    proof (rule deg_lcoeff_cancel)
-		      show "deg R (\<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)) \<le> deg R f"
-			using deg_smult_ring [of "lcoeff g" f] using prem
-			using lcoeff_nonzero2 [OF g_in_P g_not_zero] by simp
-		      show "deg R (g \<otimes>\<^bsub>P\<^esub> ?q) \<le> deg R f"
-			using monom_deg_mult [OF _ g_in_P, of f "lcoeff f"] and deg_g_le_deg_f
-			by simp
-		      show "coeff P (g \<otimes>\<^bsub>P\<^esub> ?q) (deg R f) = \<ominus> coeff P (\<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)) (deg R f)"
-			unfolding coeff_mult [OF g_in_P monom_closed [OF lcoeff_closed [OF f_in_P], of "deg R f - deg R g"], of "deg R f"]
-			unfolding coeff_monom [OF lcoeff_closed [OF f_in_P], of "(deg R f - deg R g)"]
-			using R.finsum_cong' [of "{..deg R f}" "{..deg R f}" 
-			  "(\<lambda>i. coeff P g i \<otimes> (if deg R f - deg R g = deg R f - i then lcoeff f else \<zero>))" 
-			  "(\<lambda>i. if deg R g = i then coeff P g i \<otimes> lcoeff f else \<zero>)"]
-			using R.finsum_singleton [of "deg R g" "{.. deg R f}" "(\<lambda>i. coeff P g i \<otimes> lcoeff f)"]
-			unfolding Pi_def using deg_g_le_deg_f by force
-		    qed (simp_all add: deg_f_nzero)
-		    finally show "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < n" unfolding prem .
-		  qed
-		  moreover have "\<ominus>\<^bsub>P\<^esub> ?r \<in> carrier P" by simp
-		  moreover obtain m where deg_rem_eq_m: "deg R (\<ominus>\<^bsub>P\<^esub> ?r) = m" by auto
-		  moreover have "deg R g \<le> deg R (\<ominus>\<^bsub>P\<^esub> ?r)" using n_deg_r_l_deg_g by simp
-		    (*JE: now, by applying the induction hypothesis, we obtain new quotient, remainder and exponent.*)
-		  ultimately obtain q' r' k'
-		    where rem_desc: "lcoeff g (^) (k'::nat) \<odot>\<^bsub>P\<^esub> (\<ominus>\<^bsub>P\<^esub> ?r) = g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r'"and rem_deg: "(r' = \<zero>\<^bsub>P\<^esub> \<or> deg R r' < deg R g)"
-		    and q'_in_carrier: "q' \<in> carrier P" and r'_in_carrier: "r' \<in> carrier P"
-		    using hypo by blast
-		      (*JE: we now prove that the new quotient, remainder and exponent can be used to get 
-		      the quotient, remainder and exponent of the long division theorem*)
-		  show ?thesis
-		  proof (rule exI3 [of _ "((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q')" r' "Suc k'"], intro conjI)
-		    show "(lcoeff g (^) (Suc k')) \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q') \<oplus>\<^bsub>P\<^esub> r'"
-		    proof -
-		      have "(lcoeff g (^) (Suc k')) \<odot>\<^bsub>P\<^esub> f = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r)" 
-			using smult_assoc1 exist by simp
-		      also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ( \<ominus>\<^bsub>P\<^esub> ?r))"
-			using UP_smult_r_distr by simp
-		      also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r')"
-			using rem_desc by simp
-		      also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r'"
-			using sym [OF a_assoc [of "lcoeff g (^) k' \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q)" "g \<otimes>\<^bsub>P\<^esub> q'" "r'"]]
-			using q'_in_carrier r'_in_carrier by simp
-		      also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (?q \<otimes>\<^bsub>P\<^esub> g) \<oplus>\<^bsub>P\<^esub> q' \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"
-			using q'_in_carrier by (auto simp add: m_comm)
-		      also have "\<dots> = (((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q) \<otimes>\<^bsub>P\<^esub> g) \<oplus>\<^bsub>P\<^esub> q' \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'" 
-			using smult_assoc2 q'_in_carrier by auto
-		      also have "\<dots> = ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q') \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"
-			using sym [OF l_distr] and q'_in_carrier by auto
-		      finally show ?thesis using m_comm q'_in_carrier by auto
-		    qed
-		  qed (simp_all add: rem_deg q'_in_carrier r'_in_carrier)
-		}
-	      qed
-	    }
-	  qed
-	qed
+          m = deg R x \<longrightarrow>
+          (\<exists>q r (k::nat). q \<in> carrier P \<and> r \<in> carrier P \<and> lcoeff g (^) k \<odot>\<^bsub>P\<^esub> x = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r & (r = \<zero>\<^bsub>P\<^esub> | deg R r < deg R g))"
+          and prem: "n = deg R f" and f_in_P [simp]: "f \<in> carrier P"
+          and deg_g_le_deg_f: "deg R g \<le> deg R f"
+        let ?k = "1::nat" and ?r = "(g \<otimes>\<^bsub>P\<^esub> (monom P (lcoeff f) (deg R f - deg R g))) \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)"
+          and ?q = "monom P (lcoeff f) (deg R f - deg R g)"
+        show "\<exists> q r (k::nat). q \<in> carrier P \<and> r \<in> carrier P \<and> lcoeff g (^) k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r & (r = \<zero>\<^bsub>P\<^esub> | deg R r < deg R g)"
+        proof -
+          (*JE: we first extablish the existence of a triple satisfying the previous equation. 
+            Then we will have to prove the second part of the predicate.*)
+          have exist: "lcoeff g (^) ?k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r"
+            using minus_add
+            using sym [OF a_assoc [of "g \<otimes>\<^bsub>P\<^esub> ?q" "\<ominus>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q)" "lcoeff g \<odot>\<^bsub>P\<^esub> f"]]
+            using r_neg by auto
+          show ?thesis
+          proof (cases "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < deg R g")
+            (*JE: if the degree of the remainder satisfies the statement property we are done*)
+            case True
+            {
+              show ?thesis
+              proof (rule exI3 [of _ ?q "\<ominus>\<^bsub>P\<^esub> ?r" ?k], intro conjI)
+                show "lcoeff g (^) ?k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r" using exist by simp
+                show "\<ominus>\<^bsub>P\<^esub> ?r = \<zero>\<^bsub>P\<^esub> \<or> deg R (\<ominus>\<^bsub>P\<^esub> ?r) < deg R g" using True by simp
+              qed (simp_all)
+            }
+          next
+            case False note n_deg_r_l_deg_g = False
+            {
+              (*JE: otherwise, we verify the conditions of the induction hypothesis.*)
+              show ?thesis
+              proof (cases "deg R f = 0")
+                (*JE: the solutions are different if the degree of f is zero or not*)
+                case True
+                {
+                  have deg_g: "deg R g = 0" using True using deg_g_le_deg_f by simp
+                  have "lcoeff g (^) (1::nat) \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> f \<oplus>\<^bsub>P\<^esub> \<zero>\<^bsub>P\<^esub>"
+                    unfolding deg_g apply simp
+                    unfolding sym [OF monom_mult_is_smult [OF coeff_closed [OF g_in_P, of 0] f_in_P]]
+                    using deg_zero_impl_monom [OF g_in_P deg_g] by simp
+                  then show ?thesis using f_in_P by blast
+                }
+              next
+                case False note deg_f_nzero = False
+                {
+                  (*JE: now it only remains the case where the induction hypothesis can be used.*)
+                  (*JE: we first prove that the degree of the remainder is smaller than the one of f*)
+                  have deg_remainder_l_f: "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < n"
+                  proof -
+                    have "deg R (\<ominus>\<^bsub>P\<^esub> ?r) = deg R ?r" using deg_uminus [of ?r] by simp
+                    also have "\<dots> < deg R f"
+                    proof (rule deg_lcoeff_cancel)
+                      show "deg R (\<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)) \<le> deg R f"
+                        using deg_smult_ring [of "lcoeff g" f] using prem
+                        using lcoeff_nonzero2 [OF g_in_P g_not_zero] by simp
+                      show "deg R (g \<otimes>\<^bsub>P\<^esub> ?q) \<le> deg R f"
+                        using monom_deg_mult [OF _ g_in_P, of f "lcoeff f"] and deg_g_le_deg_f
+                        by simp
+                      show "coeff P (g \<otimes>\<^bsub>P\<^esub> ?q) (deg R f) = \<ominus> coeff P (\<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)) (deg R f)"
+                        unfolding coeff_mult [OF g_in_P monom_closed [OF lcoeff_closed [OF f_in_P], of "deg R f - deg R g"], of "deg R f"]
+                        unfolding coeff_monom [OF lcoeff_closed [OF f_in_P], of "(deg R f - deg R g)"]
+                        using R.finsum_cong' [of "{..deg R f}" "{..deg R f}" 
+                          "(\<lambda>i. coeff P g i \<otimes> (if deg R f - deg R g = deg R f - i then lcoeff f else \<zero>))" 
+                          "(\<lambda>i. if deg R g = i then coeff P g i \<otimes> lcoeff f else \<zero>)"]
+                        using R.finsum_singleton [of "deg R g" "{.. deg R f}" "(\<lambda>i. coeff P g i \<otimes> lcoeff f)"]
+                        unfolding Pi_def using deg_g_le_deg_f by force
+                    qed (simp_all add: deg_f_nzero)
+                    finally show "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < n" unfolding prem .
+                  qed
+                  moreover have "\<ominus>\<^bsub>P\<^esub> ?r \<in> carrier P" by simp
+                  moreover obtain m where deg_rem_eq_m: "deg R (\<ominus>\<^bsub>P\<^esub> ?r) = m" by auto
+                  moreover have "deg R g \<le> deg R (\<ominus>\<^bsub>P\<^esub> ?r)" using n_deg_r_l_deg_g by simp
+                    (*JE: now, by applying the induction hypothesis, we obtain new quotient, remainder and exponent.*)
+                  ultimately obtain q' r' k'
+                    where rem_desc: "lcoeff g (^) (k'::nat) \<odot>\<^bsub>P\<^esub> (\<ominus>\<^bsub>P\<^esub> ?r) = g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r'"and rem_deg: "(r' = \<zero>\<^bsub>P\<^esub> \<or> deg R r' < deg R g)"
+                    and q'_in_carrier: "q' \<in> carrier P" and r'_in_carrier: "r' \<in> carrier P"
+                    using hypo by blast
+                      (*JE: we now prove that the new quotient, remainder and exponent can be used to get 
+                      the quotient, remainder and exponent of the long division theorem*)
+                  show ?thesis
+                  proof (rule exI3 [of _ "((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q')" r' "Suc k'"], intro conjI)
+                    show "(lcoeff g (^) (Suc k')) \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q') \<oplus>\<^bsub>P\<^esub> r'"
+                    proof -
+                      have "(lcoeff g (^) (Suc k')) \<odot>\<^bsub>P\<^esub> f = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r)" 
+                        using smult_assoc1 exist by simp
+                      also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ( \<ominus>\<^bsub>P\<^esub> ?r))"
+                        using UP_smult_r_distr by simp
+                      also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r')"
+                        using rem_desc by simp
+                      also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r'"
+                        using sym [OF a_assoc [of "lcoeff g (^) k' \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q)" "g \<otimes>\<^bsub>P\<^esub> q'" "r'"]]
+                        using q'_in_carrier r'_in_carrier by simp
+                      also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (?q \<otimes>\<^bsub>P\<^esub> g) \<oplus>\<^bsub>P\<^esub> q' \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"
+                        using q'_in_carrier by (auto simp add: m_comm)
+                      also have "\<dots> = (((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q) \<otimes>\<^bsub>P\<^esub> g) \<oplus>\<^bsub>P\<^esub> q' \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'" 
+                        using smult_assoc2 q'_in_carrier by auto
+                      also have "\<dots> = ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q') \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"
+                        using sym [OF l_distr] and q'_in_carrier by auto
+                      finally show ?thesis using m_comm q'_in_carrier by auto
+                    qed
+                  qed (simp_all add: rem_deg q'_in_carrier r'_in_carrier)
+                }
+              qed
+            }
+          qed
+        qed
       qed
     }
   qed
--- a/src/HOL/Algebra/poly/UnivPoly2.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Algebra/poly/UnivPoly2.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -100,7 +100,7 @@
 begin
 
 definition
-  up_add_def:	"p + q = Abs_UP (%n. Rep_UP p n + Rep_UP q n)"
+  up_add_def: "p + q = Abs_UP (%n. Rep_UP p n + Rep_UP q n)"
 
 instance ..
 
@@ -133,7 +133,7 @@
 
 definition
   up_mult_def:  "p * q = Abs_UP (%n::nat. setsum
-		     (%i. Rep_UP p i * Rep_UP q (n-i)) {..n})"
+                     (%i. Rep_UP p i * Rep_UP q (n-i)) {..n})"
 
 instance ..
 
@@ -201,18 +201,18 @@
     have "(%i. f i + g i) : UP"
     proof -
       from fup obtain n where boundn: "bound n f"
-	by (unfold UP_def) fast
+        by (unfold UP_def) fast
       from gup obtain m where boundm: "bound m g"
-	by (unfold UP_def) fast
+        by (unfold UP_def) fast
       have "bound (max n m) (%i. (f i + g i))"
       proof
-	fix i
-	assume "max n m < i"
-	with boundn and boundm show "f i + g i = 0"
+        fix i
+        assume "max n m < i"
+        with boundn and boundm show "f i + g i = 0"
           by (fastsimp simp add: algebra_simps)
       qed
       then show "(%i. (f i + g i)) : UP"
-	by (unfold UP_def) fast
+        by (unfold UP_def) fast
     qed
   }
   then show ?thesis
@@ -228,30 +228,30 @@
     have "(%n. setsum (%i. f i * g (n-i)) {..n}) : UP"
     proof -
       from fup obtain n where "bound n f"
-	by (unfold UP_def) fast
+        by (unfold UP_def) fast
       from gup obtain m where "bound m g"
-	by (unfold UP_def) fast
+        by (unfold UP_def) fast
       have "bound (n + m) (%n. setsum (%i. f i * g (n-i)) {..n})"
       proof
-	fix k
-	assume bound: "n + m < k"
-	{
-	  fix i
-	  have "f i * g (k-i) = 0"
-	  proof cases
-	    assume "n < i"
-	    with `bound n f` show ?thesis by (auto simp add: algebra_simps)
-	  next
-	    assume "~ (n < i)"
-	    with bound have "m < k-i" by arith
-	    with `bound m g` show ?thesis by (auto simp add: algebra_simps)
-	  qed
-	}
-	then show "setsum (%i. f i * g (k-i)) {..k} = 0"
-	  by (simp add: algebra_simps)
+        fix k
+        assume bound: "n + m < k"
+        {
+          fix i
+          have "f i * g (k-i) = 0"
+          proof cases
+            assume "n < i"
+            with `bound n f` show ?thesis by (auto simp add: algebra_simps)
+          next
+            assume "~ (n < i)"
+            with bound have "m < k-i" by arith
+            with `bound m g` show ?thesis by (auto simp add: algebra_simps)
+          qed
+        }
+        then show "setsum (%i. f i * g (k-i)) {..k} = 0"
+          by (simp add: algebra_simps)
       qed
       then show "(%n. setsum (%i. f i * g (n-i)) {..n}) : UP"
-	by (unfold UP_def) fast
+        by (unfold UP_def) fast
     qed
   }
   then show ?thesis
@@ -290,17 +290,17 @@
     {
       fix k and a b c :: "nat=>'a::ring"
       have "k <= n ==> 
-	setsum (%j. setsum (%i. a i * b (j-i)) {..j} * c (n-j)) {..k} = 
-	setsum (%j. a j * setsum  (%i. b i * c (n-j-i)) {..k-j}) {..k}"
-	(is "_ ==> ?eq k")
+        setsum (%j. setsum (%i. a i * b (j-i)) {..j} * c (n-j)) {..k} = 
+        setsum (%j. a j * setsum  (%i. b i * c (n-j-i)) {..k-j}) {..k}"
+        (is "_ ==> ?eq k")
       proof (induct k)
-	case 0 show ?case by simp
+        case 0 show ?case by simp
       next
-	case (Suc k)
-	then have "k <= n" by arith
-	then have "?eq k" by (rule Suc)
-	then show ?case
-	  by (simp add: Suc_diff_le natsum_ldistr)
+        case (Suc k)
+        then have "k <= n" by arith
+        then have "?eq k" by (rule Suc)
+        then show ?case
+          by (simp add: Suc_diff_le natsum_ldistr)
       qed
     }
     then show "coeff ((p * q) * r) n = coeff (p * (q * r)) n"
@@ -326,13 +326,13 @@
       fix k
       fix a b :: "nat=>'a::ring"
       have "k <= n ==> 
-	setsum (%i. a i * b (n-i)) {..k} =
-	setsum (%i. a (k-i) * b (i+n-k)) {..k}"
-	(is "_ ==> ?eq k")
+        setsum (%i. a i * b (n-i)) {..k} =
+        setsum (%i. a (k-i) * b (i+n-k)) {..k}"
+        (is "_ ==> ?eq k")
       proof (induct k)
-	case 0 show ?case by simp
+        case 0 show ?case by simp
       next
-	case (Suc k) then show ?case by (subst natsum_Suc2) simp
+        case (Suc k) then show ?case by (subst natsum_Suc2) simp
       qed
     }
     then show "coeff (p * q) n = coeff (q * p) n"
--- a/src/HOL/Auth/CertifiedEmail.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Auth/CertifiedEmail.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -43,7 +43,7 @@
 
 | FakeSSL: --{*The Spy may open SSL sessions with TTP, who is the only agent
     equipped with the necessary credentials to serve as an SSL server.*}
-	 "[| evsfssl \<in> certified_mail; X \<in> synth(analz(spies evsfssl))|]
+         "[| evsfssl \<in> certified_mail; X \<in> synth(analz(spies evsfssl))|]
           ==> Notes TTP {|Agent Spy, Agent TTP, X|} # evsfssl \<in> certified_mail"
 
 | CM1: --{*The sender approaches the recipient.  The message is a number.*}
@@ -54,14 +54,14 @@
     hs = Hash{|Number cleartext, Nonce q, response S R q, Crypt K (Number m)|};
     S2TTP = Crypt(pubEK TTP) {|Agent S, Number BothAuth, Key K, Agent R, hs|}|]
   ==> Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number BothAuth, 
-		 Number cleartext, Nonce q, S2TTP|} # evs1 
-	\<in> certified_mail"
+                 Number cleartext, Nonce q, S2TTP|} # evs1 
+        \<in> certified_mail"
 
 | CM2: --{*The recipient records @{term S2TTP} while transmitting it and her
      password to @{term TTP} over an SSL channel.*}
  "[|evs2 \<in> certified_mail;
     Gets R {|Agent S, Agent TTP, em, Number BothAuth, Number cleartext, 
-	     Nonce q, S2TTP|} \<in> set evs2;
+             Nonce q, S2TTP|} \<in> set evs2;
     TTP \<noteq> R;  
     hr = Hash {|Number cleartext, Nonce q, response S R q, em|} |]
   ==> 
@@ -289,14 +289,14 @@
  "evs \<in> certified_mail ==>
     Key K \<notin> analz (spies evs) -->
     (\<forall>AO. Crypt (pubEK TTP)
-	   {|Agent S, Number AO, Key K, Agent R, hs|} \<in> used evs -->
+           {|Agent S, Number AO, Key K, Agent R, hs|} \<in> used evs -->
     (\<exists>m ctxt q. 
         hs = Hash{|Number ctxt, Nonce q, response S R q, Crypt K (Number m)|} &
-	Says S R
-	   {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
-	     Number ctxt, Nonce q,
-	     Crypt (pubEK TTP)
-	      {|Agent S, Number AO, Key K, Agent R, hs |}|} \<in> set evs))" 
+        Says S R
+           {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
+             Number ctxt, Nonce q,
+             Crypt (pubEK TTP)
+              {|Agent S, Number AO, Key K, Agent R, hs |}|} \<in> set evs))" 
 apply (erule certified_mail.induct, analz_mono_contra)
 apply (drule_tac [5] CM2_S2TTP_parts_knows_Spy, simp)
 apply (simp add: used_Nil Crypt_notin_initState, simp_all)
@@ -322,11 +322,11 @@
     evs \<in> certified_mail|]
   ==> \<exists>m ctxt q. 
         hs = Hash{|Number ctxt, Nonce q, response S R q, Crypt K (Number m)|} &
-	Says S R
-	   {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
-	     Number ctxt, Nonce q,
-	     Crypt (pubEK TTP)
-	      {|Agent S, Number AO, Key K, Agent R, hs |}|} \<in> set evs" 
+        Says S R
+           {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
+             Number ctxt, Nonce q,
+             Crypt (pubEK TTP)
+              {|Agent S, Number AO, Key K, Agent R, hs |}|} \<in> set evs" 
 by (blast intro: S2TTP_sender_lemma) 
 
 
@@ -401,7 +401,7 @@
                      Number cleartext, Nonce q, S2TTP|} \<in> set evs;
          S2TTP = Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|};
          Key K \<in> analz (spies evs);
-	 evs \<in> certified_mail;
+         evs \<in> certified_mail;
          S\<noteq>Spy|]
       ==> R \<in> bad & Gets S (Crypt (priSK TTP) S2TTP) \<in> set evs"
 apply (erule rev_mp)
@@ -428,7 +428,7 @@
       "[|Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number AO, 
                      Number cleartext, Nonce q, S2TTP|} \<in> set evs;
          S2TTP = Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|};
-	 evs \<in> certified_mail;
+         evs \<in> certified_mail;
          S\<noteq>Spy; R \<notin> bad|]
       ==> Key K \<notin> analz(spies evs)"
 by (blast dest: S_fairness_bad_R) 
@@ -438,10 +438,10 @@
  until @{term S} has access to the return receipt.*} 
 theorem S_guarantee:
      "[|Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number AO, 
-		    Number cleartext, Nonce q, S2TTP|} \<in> set evs;
-	S2TTP = Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|};
-	Notes R {|Agent TTP, Agent R, Key K, hs|} \<in> set evs;
-	S\<noteq>Spy;  evs \<in> certified_mail|]
+                    Number cleartext, Nonce q, S2TTP|} \<in> set evs;
+        S2TTP = Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|};
+        Notes R {|Agent TTP, Agent R, Key K, hs|} \<in> set evs;
+        S\<noteq>Spy;  evs \<in> certified_mail|]
      ==> Gets S (Crypt (priSK TTP) S2TTP) \<in> set evs"
 apply (erule rev_mp)
 apply (erule ssubst)
--- a/src/HOL/Auth/Event.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Auth/Event.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -22,7 +22,7 @@
         | Notes agent       msg
        
 consts 
-  bad    :: "agent set"				(*compromised agents*)
+  bad    :: "agent set"                         -- {* compromised agents *}
   knows  :: "agent => event list => msg set"
 
 
@@ -43,19 +43,19 @@
   knows_Cons:
     "knows A (ev # evs) =
        (if A = Spy then 
-	(case ev of
-	   Says A' B X => insert X (knows Spy evs)
-	 | Gets A' X => knows Spy evs
-	 | Notes A' X  => 
-	     if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs)
-	else
-	(case ev of
-	   Says A' B X => 
-	     if A'=A then insert X (knows A evs) else knows A evs
-	 | Gets A' X    => 
-	     if A'=A then insert X (knows A evs) else knows A evs
-	 | Notes A' X    => 
-	     if A'=A then insert X (knows A evs) else knows A evs))"
+        (case ev of
+           Says A' B X => insert X (knows Spy evs)
+         | Gets A' X => knows Spy evs
+         | Notes A' X  => 
+             if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs)
+        else
+        (case ev of
+           Says A' B X => 
+             if A'=A then insert X (knows A evs) else knows A evs
+         | Gets A' X    => 
+             if A'=A then insert X (knows A evs) else knows A evs
+         | Notes A' X    => 
+             if A'=A then insert X (knows A evs) else knows A evs))"
 
 (*
   Case A=Spy on the Gets event
@@ -71,10 +71,10 @@
 primrec
   used_Nil:   "used []         = (UN B. parts (initState B))"
   used_Cons:  "used (ev # evs) =
-		     (case ev of
-			Says A B X => parts {X} \<union> used evs
-		      | Gets A X   => used evs
-		      | Notes A X  => parts {X} \<union> used evs)"
+                     (case ev of
+                        Says A B X => parts {X} \<union> used evs
+                      | Gets A X   => used evs
+                      | Notes A X  => parts {X} \<union> used evs)"
     --{*The case for @{term Gets} seems anomalous, but @{term Gets} always
         follows @{term Says} in real protocols.  Seems difficult to change.
         See @{text Gets_correct} in theory @{text "Guard/Extensions.thy"}. *}
--- a/src/HOL/Auth/KerberosIV.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Auth/KerberosIV.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -170,18 +170,18 @@
             B \<noteq> Tgs;  authK \<in> symKeys;
             Says A' Tgs \<lbrace>
              (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK,
-				 Number Ta\<rbrace>),
+                                 Number Ta\<rbrace>),
              (Crypt authK \<lbrace>Agent A, Number T2\<rbrace>), Agent B\<rbrace>
-	        \<in> set evs4;
+                \<in> set evs4;
             \<not> expiredAK Ta evs4;
             \<not> expiredA T2 evs4;
             servKlife + (CT evs4) <= authKlife + Ta
          \<rbrakk>
           \<Longrightarrow> Says Tgs A
                 (Crypt authK \<lbrace>Key servK, Agent B, Number (CT evs4),
-			       Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK,
-		 			        Number (CT evs4)\<rbrace> \<rbrace>)
-	        # evs4 \<in> kerbIV"
+                               Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK,
+                                                Number (CT evs4)\<rbrace> \<rbrace>)
+                # evs4 \<in> kerbIV"
 (* Tgs creates a new session key per each request for a service, without
    checking if there is still a fresh one for that service.
    The cipher under Tgs' key is the authTicket, the cipher under B's key
@@ -196,14 +196,14 @@
  | K5:  "\<lbrakk> evs5 \<in> kerbIV; authK \<in> symKeys; servK \<in> symKeys;
             Says A Tgs
                 \<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>,
-		  Agent B\<rbrace>
+                  Agent B\<rbrace>
               \<in> set evs5;
             Says Tgs' A
              (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
                 \<in> set evs5;
             valid Ts wrt T2 \<rbrakk>
           \<Longrightarrow> Says A B \<lbrace>servTicket,
-			 Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
+                         Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
                # evs5 \<in> kerbIV"
 (* Checks similar to those in K3. *)
 
@@ -609,7 +609,7 @@
      evs \<in> kerbIV \<rbrakk>
   \<Longrightarrow> servK \<notin> range shrK &
       (\<exists>A. servTicket =
-	      Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>)
+              Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>)
        | servTicket \<in> analz (spies evs)"
 apply (frule Says_imp_spies [THEN analz.Inj], auto)
  apply (force dest!: servTicket_form)
@@ -1336,15 +1336,15 @@
 
 lemma Confidentiality_lemma [rule_format]:
      "\<lbrakk> Says Tgs A
-	    (Crypt authK
-	       \<lbrace>Key servK, Agent B, Number Ts,
-		 Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
-	   \<in> set evs;
-	Key authK \<notin> analz (spies evs);
+            (Crypt authK
+               \<lbrace>Key servK, Agent B, Number Ts,
+                 Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
+           \<in> set evs;
+        Key authK \<notin> analz (spies evs);
         servK \<in> symKeys;
-	A \<notin> bad;  B \<notin> bad; evs \<in> kerbIV \<rbrakk>
+        A \<notin> bad;  B \<notin> bad; evs \<in> kerbIV \<rbrakk>
       \<Longrightarrow> Key servK \<in> analz (spies evs) \<longrightarrow>
-	  expiredSK Ts evs"
+          expiredSK Ts evs"
 apply (erule rev_mp)
 apply (erule rev_mp)
 apply (erule kerbIV.induct)
--- a/src/HOL/Auth/KerberosIV_Gets.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Auth/KerberosIV_Gets.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -161,18 +161,18 @@
             B \<noteq> Tgs;  authK \<in> symKeys;
             Gets Tgs \<lbrace>
              (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK,
-				 Number Ta\<rbrace>),
+                                 Number Ta\<rbrace>),
              (Crypt authK \<lbrace>Agent A, Number T2\<rbrace>), Agent B\<rbrace>
-	        \<in> set evs4;
+                \<in> set evs4;
             \<not> expiredAK Ta evs4;
             \<not> expiredA T2 evs4;
             servKlife + (CT evs4) <= authKlife + Ta
          \<rbrakk>
           \<Longrightarrow> Says Tgs A
                 (Crypt authK \<lbrace>Key servK, Agent B, Number (CT evs4),
-			       Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK,
-		 			        Number (CT evs4)\<rbrace> \<rbrace>)
-	        # evs4 \<in> kerbIV_gets"
+                               Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK,
+                                                Number (CT evs4)\<rbrace> \<rbrace>)
+                # evs4 \<in> kerbIV_gets"
 (* Tgs creates a new session key per each request for a service, without
    checking if there is still a fresh one for that service.
    The cipher under Tgs' key is the authTicket, the cipher under B's key
@@ -187,14 +187,14 @@
  | K5:  "\<lbrakk> evs5 \<in> kerbIV_gets; authK \<in> symKeys; servK \<in> symKeys;
             Says A Tgs
                 \<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>,
-		  Agent B\<rbrace>
+                  Agent B\<rbrace>
               \<in> set evs5;
             Gets A
              (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
                 \<in> set evs5;
             valid Ts wrt T2 \<rbrakk>
           \<Longrightarrow> Says A B \<lbrace>servTicket,
-			 Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
+                         Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
                # evs5 \<in> kerbIV_gets"
 (* Checks similar to those in K3. *)
 
@@ -495,7 +495,7 @@
      evs \<in> kerbIV_gets \<rbrakk>
   \<Longrightarrow> servK \<notin> range shrK &
       (\<exists>A. servTicket =
-	      Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>)
+              Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>)
        | servTicket \<in> analz (spies evs)"
 apply (frule Gets_imp_knows_Spy [THEN analz.Inj], auto)
  apply (force dest!: servTicket_form)
@@ -1231,15 +1231,15 @@
 
 lemma Confidentiality_lemma [rule_format]:
      "\<lbrakk> Says Tgs A
-	    (Crypt authK
-	       \<lbrace>Key servK, Agent B, Number Ts,
-		 Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
-	   \<in> set evs;
-	Key authK \<notin> analz (spies evs);
+            (Crypt authK
+               \<lbrace>Key servK, Agent B, Number Ts,
+                 Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
+           \<in> set evs;
+        Key authK \<notin> analz (spies evs);
         servK \<in> symKeys;
-	A \<notin> bad;  B \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
+        A \<notin> bad;  B \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
       \<Longrightarrow> Key servK \<in> analz (spies evs) \<longrightarrow>
-	  expiredSK Ts evs"
+          expiredSK Ts evs"
 apply (erule rev_mp)
 apply (erule rev_mp)
 apply (erule kerbIV_gets.induct)
--- a/src/HOL/Auth/KerberosV.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Auth/KerberosV.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -137,9 +137,9 @@
             B \<noteq> Tgs;  authK \<in> symKeys;
             Says A' Tgs \<lbrace>
              (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK,
-				 Number Ta\<rbrace>),
+                                 Number Ta\<rbrace>),
              (Crypt authK \<lbrace>Agent A, Number T2\<rbrace>), Agent B\<rbrace>
-	        \<in> set evs4;
+                \<in> set evs4;
             \<not> expiredAK Ta evs4;
             \<not> expiredA T2 evs4;
             servKlife + (CT evs4) <= authKlife + Ta
@@ -155,14 +155,14 @@
             A \<noteq> Kas; A \<noteq> Tgs;
             Says A Tgs
                 \<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>,
-		  Agent B\<rbrace>
+                  Agent B\<rbrace>
               \<in> set evs5;
             Says Tgs' A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
                           servTicket\<rbrace>
                 \<in> set evs5;
             valid Ts wrt T2 \<rbrakk>
           \<Longrightarrow> Says A B \<lbrace>servTicket,
-			 Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
+                         Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
                # evs5 \<in> kerbV"
 
   | KV6:  "\<lbrakk> evs6 \<in> kerbV; B \<noteq> Kas; B \<noteq> Tgs;
@@ -1081,14 +1081,14 @@
 
 lemma Confidentiality_lemma [rule_format]:
      "\<lbrakk> Says Tgs A
-	    \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
-	      Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>
-	   \<in> set evs;
-	Key authK \<notin> analz (spies evs);
+            \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
+              Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>
+           \<in> set evs;
+        Key authK \<notin> analz (spies evs);
         servK \<in> symKeys;
-	A \<notin> bad;  B \<notin> bad; evs \<in> kerbV \<rbrakk>
+        A \<notin> bad;  B \<notin> bad; evs \<in> kerbV \<rbrakk>
       \<Longrightarrow> Key servK \<in> analz (spies evs) \<longrightarrow>
-	  expiredSK Ts evs"
+          expiredSK Ts evs"
 apply (erule rev_mp)
 apply (erule rev_mp)
 apply (erule kerbV.induct)
--- a/src/HOL/Auth/Kerberos_BAN.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Auth/Kerberos_BAN.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -77,45 +77,45 @@
    Nil:  "[] \<in> bankerberos"
 
  | Fake: "\<lbrakk> evsf \<in> bankerberos;  X \<in> synth (analz (spies evsf)) \<rbrakk>
-	  \<Longrightarrow> Says Spy B X # evsf \<in> bankerberos"
+          \<Longrightarrow> Says Spy B X # evsf \<in> bankerberos"
 
 
  | BK1:  "\<lbrakk> evs1 \<in> bankerberos \<rbrakk>
-	  \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1
-		\<in>  bankerberos"
+          \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1
+                \<in>  bankerberos"
 
 
  | BK2:  "\<lbrakk> evs2 \<in> bankerberos;  Key K \<notin> used evs2; K \<in> symKeys;
-	     Says A' Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
-	  \<Longrightarrow> Says Server A
-		(Crypt (shrK A)
-		   \<lbrace>Number (CT evs2), Agent B, Key K,
-		    (Crypt (shrK B) \<lbrace>Number (CT evs2), Agent A, Key K\<rbrace>)\<rbrace>)
-		# evs2 \<in> bankerberos"
+             Says A' Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
+          \<Longrightarrow> Says Server A
+                (Crypt (shrK A)
+                   \<lbrace>Number (CT evs2), Agent B, Key K,
+                    (Crypt (shrK B) \<lbrace>Number (CT evs2), Agent A, Key K\<rbrace>)\<rbrace>)
+                # evs2 \<in> bankerberos"
 
 
  | BK3:  "\<lbrakk> evs3 \<in> bankerberos;
-	     Says S A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
-	       \<in> set evs3;
-	     Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
-	     \<not> expiredK Tk evs3 \<rbrakk>
-	  \<Longrightarrow> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number (CT evs3)\<rbrace> \<rbrace>
-	       # evs3 \<in> bankerberos"
+             Says S A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
+               \<in> set evs3;
+             Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
+             \<not> expiredK Tk evs3 \<rbrakk>
+          \<Longrightarrow> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number (CT evs3)\<rbrace> \<rbrace>
+               # evs3 \<in> bankerberos"
 
 
  | BK4:  "\<lbrakk> evs4 \<in> bankerberos;
-	     Says A' B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>),
-			 (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4;
-	     \<not> expiredK Tk evs4;  \<not> expiredA Ta evs4 \<rbrakk>
-	  \<Longrightarrow> Says B A (Crypt K (Number Ta)) # evs4
-		\<in> bankerberos"
+             Says A' B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>),
+                         (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4;
+             \<not> expiredK Tk evs4;  \<not> expiredA Ta evs4 \<rbrakk>
+          \<Longrightarrow> Says B A (Crypt K (Number Ta)) # evs4
+                \<in> bankerberos"
 
-	(*Old session keys may become compromised*)
+        (*Old session keys may become compromised*)
  | Oops: "\<lbrakk> evso \<in> bankerberos;
          Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
-	       \<in> set evso;
-	     expiredK Tk evso \<rbrakk>
-	  \<Longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> # evso \<in> bankerberos"
+               \<in> set evso;
+             expiredK Tk evso \<rbrakk>
+          \<Longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> # evso \<in> bankerberos"
 
 
 declare Says_imp_knows_Spy [THEN parts.Inj, dest]
--- a/src/HOL/Auth/Kerberos_BAN_Gets.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Auth/Kerberos_BAN_Gets.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -69,47 +69,47 @@
    Nil:  "[] \<in> bankerb_gets"
 
  | Fake: "\<lbrakk> evsf \<in> bankerb_gets;  X \<in> synth (analz (knows Spy evsf)) \<rbrakk>
-	  \<Longrightarrow> Says Spy B X # evsf \<in> bankerb_gets"
+          \<Longrightarrow> Says Spy B X # evsf \<in> bankerb_gets"
 
  | Reception: "\<lbrakk> evsr\<in> bankerb_gets; Says A B X \<in> set evsr \<rbrakk>
                 \<Longrightarrow> Gets B X # evsr \<in> bankerb_gets"
 
  | BK1:  "\<lbrakk> evs1 \<in> bankerb_gets \<rbrakk>
-	  \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1
-		\<in>  bankerb_gets"
+          \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1
+                \<in>  bankerb_gets"
 
 
  | BK2:  "\<lbrakk> evs2 \<in> bankerb_gets;  Key K \<notin> used evs2; K \<in> symKeys;
-	     Gets Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
-	  \<Longrightarrow> Says Server A
-		(Crypt (shrK A)
-		   \<lbrace>Number (CT evs2), Agent B, Key K,
-		    (Crypt (shrK B) \<lbrace>Number (CT evs2), Agent A, Key K\<rbrace>)\<rbrace>)
-		# evs2 \<in> bankerb_gets"
+             Gets Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
+          \<Longrightarrow> Says Server A
+                (Crypt (shrK A)
+                   \<lbrace>Number (CT evs2), Agent B, Key K,
+                    (Crypt (shrK B) \<lbrace>Number (CT evs2), Agent A, Key K\<rbrace>)\<rbrace>)
+                # evs2 \<in> bankerb_gets"
 
 
  | BK3:  "\<lbrakk> evs3 \<in> bankerb_gets;
-	     Gets A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
-	       \<in> set evs3;
-	     Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
-	     \<not> expiredK Tk evs3 \<rbrakk>
-	  \<Longrightarrow> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number (CT evs3)\<rbrace> \<rbrace>
-	       # evs3 \<in> bankerb_gets"
+             Gets A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
+               \<in> set evs3;
+             Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
+             \<not> expiredK Tk evs3 \<rbrakk>
+          \<Longrightarrow> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number (CT evs3)\<rbrace> \<rbrace>
+               # evs3 \<in> bankerb_gets"
 
 
  | BK4:  "\<lbrakk> evs4 \<in> bankerb_gets;
-	     Gets B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>),
-			 (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4;
-	     \<not> expiredK Tk evs4;  \<not> expiredA Ta evs4 \<rbrakk>
-	  \<Longrightarrow> Says B A (Crypt K (Number Ta)) # evs4
-		\<in> bankerb_gets"
+             Gets B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>),
+                         (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4;
+             \<not> expiredK Tk evs4;  \<not> expiredA Ta evs4 \<rbrakk>
+          \<Longrightarrow> Says B A (Crypt K (Number Ta)) # evs4
+                \<in> bankerb_gets"
 
-	(*Old session keys may become compromised*)
+        (*Old session keys may become compromised*)
  | Oops: "\<lbrakk> evso \<in> bankerb_gets;
          Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
-	       \<in> set evso;
-	     expiredK Tk evso \<rbrakk>
-	  \<Longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> # evso \<in> bankerb_gets"
+               \<in> set evso;
+             expiredK Tk evso \<rbrakk>
+          \<Longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> # evso \<in> bankerb_gets"
 
 
 declare Says_imp_knows_Spy [THEN parts.Inj, dest]
@@ -359,7 +359,7 @@
    "\<lbrakk> Says B A (Crypt K (Number Ta)) \<in> set evs;
       B \<notin> bad; evs \<in> bankerb_gets \<rbrakk>
   \<Longrightarrow> \<exists> Tk. Gets B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,
-	            Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs"
+                    Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs"
 apply (erule rev_mp)
 apply (erule bankerb_gets.induct)
 apply auto
--- a/src/HOL/Auth/Message.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Auth/Message.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -40,13 +40,13 @@
   agent = Server | Friend nat | Spy
 
 datatype
-     msg = Agent  agent	    --{*Agent names*}
+     msg = Agent  agent     --{*Agent names*}
          | Number nat       --{*Ordinary integers, timestamps, ...*}
          | Nonce  nat       --{*Unguessable nonces*}
          | Key    key       --{*Crypto keys*}
-	 | Hash   msg       --{*Hashing*}
-	 | MPair  msg msg   --{*Compound messages*}
-	 | Crypt  key msg   --{*Encryption, public- or shared-key*}
+         | Hash   msg       --{*Hashing*}
+         | MPair  msg msg   --{*Compound messages*}
+         | Crypt  key msg   --{*Encryption, public- or shared-key*}
 
 
 text{*Concrete syntax: messages appear as {|A,B,NA|}, etc...*}
@@ -873,8 +873,8 @@
     (Fake_insert_simp_tac ss 1
      THEN
      IF_UNSOLVED (Blast.depth_tac
-		  (cs addIs [@{thm analz_insertI},
-				   impOfSubs @{thm analz_subset_parts}]) 4 1))
+                  (cs addIs [@{thm analz_insertI},
+                                   impOfSubs @{thm analz_subset_parts}]) 4 1))
 
 fun spy_analz_tac (cs,ss) i =
   DETERM
--- a/src/HOL/Auth/NS_Public.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Auth/NS_Public.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -110,8 +110,8 @@
 lemma A_trusts_NS2_lemma [rule_format]: 
    "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>                     
     \<Longrightarrow> Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace> \<in> parts (spies evs) \<longrightarrow>
-	Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
-	Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
+        Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
+        Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
 apply (erule ns_public.induct, simp_all)
 (*Fake, NS1*)
 apply (blast dest: Spy_not_see_NA)+
@@ -129,8 +129,8 @@
 lemma B_trusts_NS1 [rule_format]:
      "evs \<in> ns_public                                         
       \<Longrightarrow> Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs) \<longrightarrow>
-	  Nonce NA \<notin> analz (spies evs) \<longrightarrow>
-	  Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
+          Nonce NA \<notin> analz (spies evs) \<longrightarrow>
+          Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
 apply (erule ns_public.induct, simp_all)
 (*Fake*)
 apply (blast intro!: analz_insertI)
--- a/src/HOL/Auth/NS_Public_Bad.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Auth/NS_Public_Bad.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -116,8 +116,8 @@
 lemma A_trusts_NS2_lemma [rule_format]: 
    "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>                     
     \<Longrightarrow> Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace> \<in> parts (spies evs) \<longrightarrow>
-	Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
-	Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs"
+        Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
+        Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs"
 apply (erule ns_public.induct)
 apply (auto dest: Spy_not_see_NA unique_NA)
 done
@@ -134,8 +134,8 @@
 lemma B_trusts_NS1 [rule_format]:
      "evs \<in> ns_public                                         
       \<Longrightarrow> Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs) \<longrightarrow>
-	  Nonce NA \<notin> analz (spies evs) \<longrightarrow>
-	  Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
+          Nonce NA \<notin> analz (spies evs) \<longrightarrow>
+          Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
 apply (erule ns_public.induct, simp_all)
 (*Fake*)
 apply (blast intro!: analz_insertI)
--- a/src/HOL/Auth/NS_Shared.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Auth/NS_Shared.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -27,61 +27,61 @@
 
 inductive_set ns_shared :: "event list set"
  where
-	(*Initial trace is empty*)
+        (*Initial trace is empty*)
   Nil:  "[] \<in> ns_shared"
-	(*The spy MAY say anything he CAN say.  We do not expect him to
-	  invent new nonces here, but he can also use NS1.  Common to
-	  all similar protocols.*)
+        (*The spy MAY say anything he CAN say.  We do not expect him to
+          invent new nonces here, but he can also use NS1.  Common to
+          all similar protocols.*)
 | Fake: "\<lbrakk>evsf \<in> ns_shared;  X \<in> synth (analz (spies evsf))\<rbrakk>
-	 \<Longrightarrow> Says Spy B X # evsf \<in> ns_shared"
+         \<Longrightarrow> Says Spy B X # evsf \<in> ns_shared"
 
-	(*Alice initiates a protocol run, requesting to talk to any B*)
+        (*Alice initiates a protocol run, requesting to talk to any B*)
 | NS1:  "\<lbrakk>evs1 \<in> ns_shared;  Nonce NA \<notin> used evs1\<rbrakk>
-	 \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> # evs1  \<in>  ns_shared"
+         \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> # evs1  \<in>  ns_shared"
 
-	(*Server's response to Alice's message.
-	  !! It may respond more than once to A's request !!
-	  Server doesn't know who the true sender is, hence the A' in
-	      the sender field.*)
+        (*Server's response to Alice's message.
+          !! It may respond more than once to A's request !!
+          Server doesn't know who the true sender is, hence the A' in
+              the sender field.*)
 | NS2:  "\<lbrakk>evs2 \<in> ns_shared;  Key KAB \<notin> used evs2;  KAB \<in> symKeys;
-	  Says A' Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs2\<rbrakk>
-	 \<Longrightarrow> Says Server A
-	       (Crypt (shrK A)
-		  \<lbrace>Nonce NA, Agent B, Key KAB,
-		    (Crypt (shrK B) \<lbrace>Key KAB, Agent A\<rbrace>)\<rbrace>)
-	       # evs2 \<in> ns_shared"
+          Says A' Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs2\<rbrakk>
+         \<Longrightarrow> Says Server A
+               (Crypt (shrK A)
+                  \<lbrace>Nonce NA, Agent B, Key KAB,
+                    (Crypt (shrK B) \<lbrace>Key KAB, Agent A\<rbrace>)\<rbrace>)
+               # evs2 \<in> ns_shared"
 
-	 (*We can't assume S=Server.  Agent A "remembers" her nonce.
-	   Need A \<noteq> Server because we allow messages to self.*)
+         (*We can't assume S=Server.  Agent A "remembers" her nonce.
+           Need A \<noteq> Server because we allow messages to self.*)
 | NS3:  "\<lbrakk>evs3 \<in> ns_shared;  A \<noteq> Server;
-	  Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs3;
-	  Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs3\<rbrakk>
-	 \<Longrightarrow> Says A B X # evs3 \<in> ns_shared"
+          Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs3;
+          Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs3\<rbrakk>
+         \<Longrightarrow> Says A B X # evs3 \<in> ns_shared"
 
-	(*Bob's nonce exchange.  He does not know who the message came
-	  from, but responds to A because she is mentioned inside.*)
+        (*Bob's nonce exchange.  He does not know who the message came
+          from, but responds to A because she is mentioned inside.*)
 | NS4:  "\<lbrakk>evs4 \<in> ns_shared;  Nonce NB \<notin> used evs4;  K \<in> symKeys;
-	  Says A' B (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>) \<in> set evs4\<rbrakk>
-	 \<Longrightarrow> Says B A (Crypt K (Nonce NB)) # evs4 \<in> ns_shared"
+          Says A' B (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>) \<in> set evs4\<rbrakk>
+         \<Longrightarrow> Says B A (Crypt K (Nonce NB)) # evs4 \<in> ns_shared"
 
-	(*Alice responds with Nonce NB if she has seen the key before.
-	  Maybe should somehow check Nonce NA again.
-	  We do NOT send NB-1 or similar as the Spy cannot spoof such things.
-	  Letting the Spy add or subtract 1 lets him send all nonces.
-	  Instead we distinguish the messages by sending the nonce twice.*)
+        (*Alice responds with Nonce NB if she has seen the key before.
+          Maybe should somehow check Nonce NA again.
+          We do NOT send NB-1 or similar as the Spy cannot spoof such things.
+          Letting the Spy add or subtract 1 lets him send all nonces.
+          Instead we distinguish the messages by sending the nonce twice.*)
 | NS5:  "\<lbrakk>evs5 \<in> ns_shared;  K \<in> symKeys;
-	  Says B' A (Crypt K (Nonce NB)) \<in> set evs5;
-	  Says S  A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>)
-	    \<in> set evs5\<rbrakk>
-	 \<Longrightarrow> Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) # evs5 \<in> ns_shared"
+          Says B' A (Crypt K (Nonce NB)) \<in> set evs5;
+          Says S  A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>)
+            \<in> set evs5\<rbrakk>
+         \<Longrightarrow> Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) # evs5 \<in> ns_shared"
 
-	(*This message models possible leaks of session keys.
-	  The two Nonces identify the protocol run: the rule insists upon
-	  the true senders in order to make them accurate.*)
+        (*This message models possible leaks of session keys.
+          The two Nonces identify the protocol run: the rule insists upon
+          the true senders in order to make them accurate.*)
 | Oops: "\<lbrakk>evso \<in> ns_shared;  Says B A (Crypt K (Nonce NB)) \<in> set evso;
-	  Says Server A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>)
-	      \<in> set evso\<rbrakk>
-	 \<Longrightarrow> Notes Spy \<lbrace>Nonce NA, Nonce NB, Key K\<rbrace> # evso \<in> ns_shared"
+          Says Server A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>)
+              \<in> set evso\<rbrakk>
+         \<Longrightarrow> Notes Spy \<lbrace>Nonce NA, Nonce NB, Key K\<rbrace> # evso \<in> ns_shared"
 
 
 declare Says_imp_knows_Spy [THEN parts.Inj, dest]
@@ -98,7 +98,7 @@
 apply (intro exI bexI)
 apply (rule_tac [2] ns_shared.Nil
        [THEN ns_shared.NS1, THEN ns_shared.NS2, THEN ns_shared.NS3,
-	THEN ns_shared.NS4, THEN ns_shared.NS5])
+        THEN ns_shared.NS4, THEN ns_shared.NS5])
 apply (possibility, simp add: used_Cons)
 done
 
@@ -275,7 +275,7 @@
 apply blast
 txt{*NS3*}
 apply (blast dest!: Crypt_Spy_analz_bad A_trusts_NS2
-	     dest:  Says_imp_knows_Spy analz.Inj unique_session_keys)
+             dest:  Says_imp_knows_Spy analz.Inj unique_session_keys)
 txt{*Oops*}
 apply (blast dest: unique_session_keys)
 done
@@ -355,8 +355,8 @@
   "\<lbrakk>B \<notin> bad;  evs \<in> ns_shared\<rbrakk> \<Longrightarrow>
      Key K \<notin> analz (spies evs) \<longrightarrow>
      Says Server A
-	  (Crypt (shrK A) \<lbrace>NA, Agent B, Key K,
-			    Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>\<rbrace>) \<in> set evs \<longrightarrow>
+          (Crypt (shrK A) \<lbrace>NA, Agent B, Key K,
+                            Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>\<rbrace>) \<in> set evs \<longrightarrow>
      Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace> \<in> parts (spies evs) \<longrightarrow>
      Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) \<in> set evs"
 apply (erule ns_shared.induct, force)
@@ -366,7 +366,7 @@
 apply (blast dest!: new_keys_not_used Crypt_imp_keysFor)
 txt{*NS5*}
 apply (blast dest!: A_trusts_NS2
-	     dest: Says_imp_knows_Spy [THEN analz.Inj]
+             dest: Says_imp_knows_Spy [THEN analz.Inj]
                    unique_session_keys Crypt_Spy_analz_bad)
 done
 
--- a/src/HOL/Auth/OtwayRees.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Auth/OtwayRees.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -61,7 +61,7 @@
                  # evs3 : otway"
 
          (*Bob receives the Server's (?) message and compares the Nonces with
-	   those in the message he previously sent the Server.
+           those in the message he previously sent the Server.
            Need B \<noteq> Server because we allow messages to self.*)
  | OR4:  "[| evs4 \<in> otway;  B \<noteq> Server;
              Says B Server {|Nonce NA, Agent A, Agent B, X',
--- a/src/HOL/Auth/OtwayReesBella.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Auth/OtwayReesBella.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -22,50 +22,50 @@
   Nil:  "[]\<in> orb"
 
 | Fake: "\<lbrakk>evsa\<in> orb;  X\<in> synth (analz (knows Spy evsa))\<rbrakk>
- 	 \<Longrightarrow> Says Spy B X  # evsa \<in> orb"
+         \<Longrightarrow> Says Spy B X  # evsa \<in> orb"
 
 | Reception: "\<lbrakk>evsr\<in> orb;  Says A B X \<in> set evsr\<rbrakk>
-	      \<Longrightarrow> Gets B X # evsr \<in> orb"
+              \<Longrightarrow> Gets B X # evsr \<in> orb"
 
 | OR1:  "\<lbrakk>evs1\<in> orb;  Nonce NA \<notin> used evs1\<rbrakk>
-	 \<Longrightarrow> Says A B \<lbrace>Nonce M, Agent A, Agent B, 
-		   Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> 
-	       # evs1 \<in> orb"
+         \<Longrightarrow> Says A B \<lbrace>Nonce M, Agent A, Agent B, 
+                   Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> 
+               # evs1 \<in> orb"
 
 | OR2:  "\<lbrakk>evs2\<in> orb;  Nonce NB \<notin> used evs2;
-	   Gets B \<lbrace>Nonce M, Agent A, Agent B, X\<rbrace> \<in> set evs2\<rbrakk>
-	\<Longrightarrow> Says B Server 
-		\<lbrace>Nonce M, Agent A, Agent B, X, 
-	   Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
-	       # evs2 \<in> orb"
+           Gets B \<lbrace>Nonce M, Agent A, Agent B, X\<rbrace> \<in> set evs2\<rbrakk>
+        \<Longrightarrow> Says B Server 
+                \<lbrace>Nonce M, Agent A, Agent B, X, 
+           Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
+               # evs2 \<in> orb"
 
 | OR3:  "\<lbrakk>evs3\<in> orb;  Key KAB \<notin> used evs3;
-	  Gets Server 
-	     \<lbrace>Nonce M, Agent A, Agent B, 
-	       Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>, 
-	       Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
-	  \<in> set evs3\<rbrakk>
-	\<Longrightarrow> Says Server B \<lbrace>Nonce M,
-		    Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,
-				      Nonce NB, Key KAB\<rbrace>\<rbrace>
-	       # evs3 \<in> orb"
+          Gets Server 
+             \<lbrace>Nonce M, Agent A, Agent B, 
+               Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>, 
+               Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
+          \<in> set evs3\<rbrakk>
+        \<Longrightarrow> Says Server B \<lbrace>Nonce M,
+                    Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,
+                                      Nonce NB, Key KAB\<rbrace>\<rbrace>
+               # evs3 \<in> orb"
 
   (*B can only check that the message he is bouncing is a ciphertext*)
   (*Sending M back is omitted*)   
 | OR4:  "\<lbrakk>evs4\<in> orb; B \<noteq> Server; \<forall> p q. X \<noteq> \<lbrace>p, q\<rbrace>; 
-	  Says B Server \<lbrace>Nonce M, Agent A, Agent B, X', 
-		Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
-	    \<in> set evs4;
-	  Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce NB, Key KAB\<rbrace>\<rbrace>
-	    \<in> set evs4\<rbrakk>
-	\<Longrightarrow> Says B A \<lbrace>Nonce M, X\<rbrace> # evs4 \<in> orb"
+          Says B Server \<lbrace>Nonce M, Agent A, Agent B, X', 
+                Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
+            \<in> set evs4;
+          Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce NB, Key KAB\<rbrace>\<rbrace>
+            \<in> set evs4\<rbrakk>
+        \<Longrightarrow> Says B A \<lbrace>Nonce M, X\<rbrace> # evs4 \<in> orb"
 
 
 | Oops: "\<lbrakk>evso\<in> orb;  
-	   Says Server B \<lbrace>Nonce M,
-		    Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,
-				      Nonce NB, Key KAB\<rbrace>\<rbrace> 
-	     \<in> set evso\<rbrakk>
+           Says Server B \<lbrace>Nonce M,
+                    Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,
+                                      Nonce NB, Key KAB\<rbrace>\<rbrace> 
+             \<in> set evso\<rbrakk>
  \<Longrightarrow> Notes Spy \<lbrace>Agent A, Agent B, Nonce NA, Nonce NB, Key KAB\<rbrace> # evso 
      \<in> orb"
 
--- a/src/HOL/Auth/OtwayRees_AN.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Auth/OtwayRees_AN.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -35,7 +35,7 @@
         
  | Reception: --{*A message that has been sent can be received by the
                   intended recipient.*}
-	      "[| evsr \<in> otway;  Says A B X \<in>set evsr |]
+              "[| evsr \<in> otway;  Says A B X \<in>set evsr |]
                ==> Gets B X # evsr \<in> otway"
 
  | OR1:  --{*Alice initiates a protocol run*}
@@ -43,14 +43,14 @@
           ==> Says A B {|Agent A, Agent B, Nonce NA|} # evs1 \<in> otway"
 
  | OR2:  --{*Bob's response to Alice's message.*}
-	 "[| evs2 \<in> otway;
+         "[| evs2 \<in> otway;
              Gets B {|Agent A, Agent B, Nonce NA|} \<in>set evs2 |]
           ==> Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|}
                  # evs2 \<in> otway"
 
  | OR3:  --{*The Server receives Bob's message.  Then he sends a new
            session key to Bob with a packet for forwarding to Alice.*}
-	 "[| evs3 \<in> otway;  Key KAB \<notin> used evs3;
+         "[| evs3 \<in> otway;  Key KAB \<notin> used evs3;
              Gets Server {|Agent A, Agent B, Nonce NA, Nonce NB|}
                \<in>set evs3 |]
           ==> Says Server B
@@ -59,9 +59,9 @@
               # evs3 \<in> otway"
 
  | OR4:  --{*Bob receives the Server's (?) message and compares the Nonces with
-	     those in the message he previously sent the Server.
+             those in the message he previously sent the Server.
              Need @{term "B \<noteq> Server"} because we allow messages to self.*}
-	 "[| evs4 \<in> otway;  B \<noteq> Server;
+         "[| evs4 \<in> otway;  B \<noteq> Server;
              Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|} \<in>set evs4;
              Gets B {|X, Crypt(shrK B){|Nonce NB,Agent A,Agent B,Key K|}|}
                \<in>set evs4 |]
@@ -69,7 +69,7 @@
 
  | Oops: --{*This message models possible leaks of session keys.  The nonces
              identify the protocol run.*}
-	 "[| evso \<in> otway;
+         "[| evso \<in> otway;
              Says Server B
                       {|Crypt (shrK A) {|Nonce NA, Agent A, Agent B, Key K|},
                         Crypt (shrK B) {|Nonce NB, Agent A, Agent B, Key K|}|}
--- a/src/HOL/Auth/OtwayRees_Bad.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Auth/OtwayRees_Bad.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -32,18 +32,18 @@
         
  | Reception: --{*A message that has been sent can be received by the
                   intended recipient.*}
-	      "[| evsr \<in> otway;  Says A B X \<in>set evsr |]
+              "[| evsr \<in> otway;  Says A B X \<in>set evsr |]
                ==> Gets B X # evsr \<in> otway"
 
  | OR1:  --{*Alice initiates a protocol run*}
-	 "[| evs1 \<in> otway;  Nonce NA \<notin> used evs1 |]
+         "[| evs1 \<in> otway;  Nonce NA \<notin> used evs1 |]
           ==> Says A B {|Nonce NA, Agent A, Agent B,
                          Crypt (shrK A) {|Nonce NA, Agent A, Agent B|} |}
                  # evs1 \<in> otway"
 
  | OR2:  --{*Bob's response to Alice's message.
              This variant of the protocol does NOT encrypt NB.*}
-	 "[| evs2 \<in> otway;  Nonce NB \<notin> used evs2;
+         "[| evs2 \<in> otway;  Nonce NB \<notin> used evs2;
              Gets B {|Nonce NA, Agent A, Agent B, X|} \<in> set evs2 |]
           ==> Says B Server
                   {|Nonce NA, Agent A, Agent B, X, Nonce NB,
@@ -53,7 +53,7 @@
  | OR3:  --{*The Server receives Bob's message and checks that the three NAs
            match.  Then he sends a new session key to Bob with a packet for
            forwarding to Alice.*}
-	 "[| evs3 \<in> otway;  Key KAB \<notin> used evs3;
+         "[| evs3 \<in> otway;  Key KAB \<notin> used evs3;
              Gets Server
                   {|Nonce NA, Agent A, Agent B,
                     Crypt (shrK A) {|Nonce NA, Agent A, Agent B|},
@@ -67,9 +67,9 @@
                  # evs3 \<in> otway"
 
  | OR4:  --{*Bob receives the Server's (?) message and compares the Nonces with
-	     those in the message he previously sent the Server.
+             those in the message he previously sent the Server.
              Need @{term "B \<noteq> Server"} because we allow messages to self.*}
-	 "[| evs4 \<in> otway;  B \<noteq> Server;
+         "[| evs4 \<in> otway;  B \<noteq> Server;
              Says B Server {|Nonce NA, Agent A, Agent B, X', Nonce NB,
                              Crypt (shrK B) {|Nonce NA, Agent A, Agent B|}|}
                \<in> set evs4;
@@ -79,7 +79,7 @@
 
  | Oops: --{*This message models possible leaks of session keys.  The nonces
              identify the protocol run.*}
-	 "[| evso \<in> otway;
+         "[| evso \<in> otway;
              Says Server B {|Nonce NA, X, Crypt (shrK B) {|Nonce NB, Key K|}|}
                \<in> set evso |]
           ==> Notes Spy {|Nonce NA, Nonce NB, Key K|} # evso \<in> otway"
--- a/src/HOL/Auth/Recur.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Auth/Recur.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -101,7 +101,7 @@
      etc.
 
    Oops:  "[| evso \<in> recur;  Says Server B RB \<in> set evso;
-	      RB \<in> responses evs';  Key K \<in> parts {RB} |]
+              RB \<in> responses evs';  Key K \<in> parts {RB} |]
            ==> Notes Spy {|Key K, RB|} # evso \<in> recur"
   *)
 
@@ -140,10 +140,10 @@
 apply (rule_tac [2] 
           recur.Nil
            [THEN recur.RA1 [of _ NA], 
-	    THEN recur.RA2 [of _ NB],
-	    THEN recur.RA3 [OF _ _ respond.One 
+            THEN recur.RA2 [of _ NB],
+            THEN recur.RA3 [OF _ _ respond.One 
                                      [THEN respond.Cons [of _ _ K _ K']]],
-	    THEN recur.RA4], possibility)
+            THEN recur.RA4], possibility)
 apply (auto simp add: used_Cons)
 done
 
@@ -241,7 +241,7 @@
                    (K \<in> KK | Key K \<in> analz (insert RB H))"
 apply (erule responses.induct)
 apply (simp_all del: image_insert
-	        add: analz_image_freshK_simps, auto)
+                add: analz_image_freshK_simps, auto)
 done 
 
 
--- a/src/HOL/Auth/Smartcard/EventSC.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Auth/Smartcard/EventSC.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -38,13 +38,13 @@
 specification (stolen)
   (*The server's card is secure by assumption\<dots>*)
   Card_Server_not_stolen [iff]: "Card Server \<notin> stolen"
-  Card_Spy_not_stolen  	 [iff]: "Card Spy \<notin> stolen"
+  Card_Spy_not_stolen    [iff]: "Card Spy \<notin> stolen"
   apply blast done
 
 specification (cloned)
   (*\<dots> the spy's card is secure because she already can use it freely*)
   Card_Server_not_cloned [iff]: "Card Server \<notin> cloned"
-  Card_Spy_not_cloned  	 [iff]: "Card Spy \<notin> cloned"
+  Card_Spy_not_cloned    [iff]: "Card Spy \<notin> cloned"
   apply blast done
 
 
@@ -52,28 +52,28 @@
           assumption of secure means*)
   knows_Nil:   "knows A [] = initState A"
   knows_Cons:  "knows A (ev # evs) =
-	 (case ev of
-	    Says A' B X => 
+         (case ev of
+            Says A' B X => 
                 if (A=A' | A=Spy) then insert X (knows A evs) else knows A evs
-	  | Notes A' X  => 
-	        if (A=A' | (A=Spy & A'\<in>bad)) then insert X (knows A evs) 
+          | Notes A' X  => 
+                if (A=A' | (A=Spy & A'\<in>bad)) then insert X (knows A evs) 
                                              else knows A evs 
           | Gets A' X   =>
-		if (A=A' & A \<noteq> Spy) then insert X (knows A evs) 
+                if (A=A' & A \<noteq> Spy) then insert X (knows A evs) 
                                      else knows A evs
           | Inputs A' C X =>
-	      if secureM then
+              if secureM then
                 if A=A' then insert X (knows A evs) else knows A evs
-	      else
-	        if (A=A' | A=Spy) then insert X (knows A evs) else knows A evs
+              else
+                if (A=A' | A=Spy) then insert X (knows A evs) else knows A evs
           | C_Gets C X   => knows A evs
           | Outpts C A' X =>
-	      if secureM then
+              if secureM then
                 if A=A' then insert X (knows A evs) else knows A evs
               else
-	        if A=Spy then insert X (knows A evs) else knows A evs
+                if A=Spy then insert X (knows A evs) else knows A evs
           | A_Gets A' X   =>
-		if (A=A' & A \<noteq> Spy) then insert X (knows A evs) 
+                if (A=A' & A \<noteq> Spy) then insert X (knows A evs) 
                                      else knows A evs)"
 
 
@@ -86,14 +86,14 @@
 primrec
   used_Nil:   "used []         = (UN B. parts (initState B))"
   used_Cons:  "used (ev # evs) =
-	         (case ev of
-		    Says A B X => parts {X} \<union> (used evs)
-		  | Notes A X  => parts {X} \<union> (used evs)
-		  | Gets A X   => used evs
+                 (case ev of
+                    Says A B X => parts {X} \<union> (used evs)
+                  | Notes A X  => parts {X} \<union> (used evs)
+                  | Gets A X   => used evs
                   | Inputs A C X  => parts{X} \<union> (used evs) 
-		  | C_Gets C X   => used evs
+                  | C_Gets C X   => used evs
                   | Outpts C A X  => parts{X} \<union> (used evs)
-		  | A_Gets A X   => used evs)"
+                  | A_Gets A X   => used evs)"
     --{*@{term Gets} always follows @{term Says} in real protocols. 
        Likewise, @{term C_Gets} will always have to follow @{term Inputs}
        and @{term A_Gets} will always have to follow @{term Outpts}*}
--- a/src/HOL/Auth/Smartcard/Smartcard.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Auth/Smartcard/Smartcard.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -386,8 +386,8 @@
 
 val analz_image_freshK_ss = 
      @{simpset} delsimps [image_insert, image_Un]
-	       delsimps [@{thm imp_disjL}]    (*reduces blow-up*)
-	       addsimps @{thms analz_image_freshK_simps}
+               delsimps [@{thm imp_disjL}]    (*reduces blow-up*)
+               addsimps @{thms analz_image_freshK_simps}
 end
 *}
 
--- a/src/HOL/Auth/TLS.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Auth/TLS.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -36,7 +36,7 @@
 Proofs would be simpler if ClientKeyExch included A's name within
 Crypt KB (Nonce PMS).  As things stand, there is much overlap between proofs
 about that message (which B receives) and the stronger event
-	Notes A {|Agent B, Nonce PMS|}.
+Notes A {|Agent B, Nonce PMS|}.
 *)
 
 header{*The TLS Protocol: Transport Layer Security*}
@@ -112,30 +112,30 @@
  | SpyKeys: --{*The spy may apply @{term PRF} and @{term sessionK}
                 to available nonces*}
          "[| evsSK \<in> tls;
-	     {Nonce NA, Nonce NB, Nonce M} <= analz (spies evsSK) |]
+             {Nonce NA, Nonce NB, Nonce M} <= analz (spies evsSK) |]
           ==> Notes Spy {| Nonce (PRF(M,NA,NB)),
-			   Key (sessionK((NA,NB,M),role)) |} # evsSK \<in> tls"
+                           Key (sessionK((NA,NB,M),role)) |} # evsSK \<in> tls"
 
  | ClientHello:
-	 --{*(7.4.1.2)
-	   PA represents @{text CLIENT_VERSION}, @{text CIPHER_SUITES} and @{text COMPRESSION_METHODS}.
-	   It is uninterpreted but will be confirmed in the FINISHED messages.
-	   NA is CLIENT RANDOM, while SID is @{text SESSION_ID}.
+         --{*(7.4.1.2)
+           PA represents @{text CLIENT_VERSION}, @{text CIPHER_SUITES} and @{text COMPRESSION_METHODS}.
+           It is uninterpreted but will be confirmed in the FINISHED messages.
+           NA is CLIENT RANDOM, while SID is @{text SESSION_ID}.
            UNIX TIME is omitted because the protocol doesn't use it.
            May assume @{term "NA \<notin> range PRF"} because CLIENT RANDOM is 
            28 bytes while MASTER SECRET is 48 bytes*}
          "[| evsCH \<in> tls;  Nonce NA \<notin> used evsCH;  NA \<notin> range PRF |]
           ==> Says A B {|Agent A, Nonce NA, Number SID, Number PA|}
-	        # evsCH  \<in>  tls"
+                # evsCH  \<in>  tls"
 
  | ServerHello:
          --{*7.4.1.3 of the TLS Internet-Draft
-	   PB represents @{text CLIENT_VERSION}, @{text CIPHER_SUITE} and @{text COMPRESSION_METHOD}.
+           PB represents @{text CLIENT_VERSION}, @{text CIPHER_SUITE} and @{text COMPRESSION_METHOD}.
            SERVER CERTIFICATE (7.4.2) is always present.
            @{text CERTIFICATE_REQUEST} (7.4.4) is implied.*}
          "[| evsSH \<in> tls;  Nonce NB \<notin> used evsSH;  NB \<notin> range PRF;
              Says A' B {|Agent A, Nonce NA, Number SID, Number PA|}
-	       \<in> set evsSH |]
+               \<in> set evsSH |]
           ==> Says B A {|Nonce NB, Number SID, Number PB|} # evsSH  \<in>  tls"
 
  | Certificate:
@@ -148,28 +148,28 @@
            She encrypts PMS using the supplied KB, which ought to be pubK B.
            We assume @{term "PMS \<notin> range PRF"} because a clash betweem the PMS
            and another MASTER SECRET is highly unlikely (even though
-	   both items have the same length, 48 bytes).
+           both items have the same length, 48 bytes).
            The Note event records in the trace that she knows PMS
                (see REMARK at top). *}
          "[| evsCX \<in> tls;  Nonce PMS \<notin> used evsCX;  PMS \<notin> range PRF;
              Says B' A (certificate B KB) \<in> set evsCX |]
           ==> Says A B (Crypt KB (Nonce PMS))
-	      # Notes A {|Agent B, Nonce PMS|}
-	      # evsCX  \<in>  tls"
+              # Notes A {|Agent B, Nonce PMS|}
+              # evsCX  \<in>  tls"
 
  | CertVerify:
-	--{*The optional Certificate Verify (7.4.8) message contains the
+        --{*The optional Certificate Verify (7.4.8) message contains the
           specific components listed in the security analysis, F.1.1.2.
           It adds the pre-master-secret, which is also essential!
           Checking the signature, which is the only use of A's certificate,
           assures B of A's presence*}
          "[| evsCV \<in> tls;
              Says B' A {|Nonce NB, Number SID, Number PB|} \<in> set evsCV;
-	     Notes A {|Agent B, Nonce PMS|} \<in> set evsCV |]
+             Notes A {|Agent B, Nonce PMS|} \<in> set evsCV |]
           ==> Says A B (Crypt (priK A) (Hash{|Nonce NB, Agent B, Nonce PMS|}))
               # evsCV  \<in>  tls"
 
-	--{*Finally come the FINISHED messages (7.4.8), confirming PA and PB
+        --{*Finally come the FINISHED messages (7.4.8), confirming PA and PB
           among other things.  The master-secret is PRF(PMS,NA,NB).
           Either party may send its message first.*}
 
@@ -181,60 +181,60 @@
           could simply put @{term "A\<noteq>Spy"} into the rule, but one should not
           expect the spy to be well-behaved.*}
          "[| evsCF \<in> tls;
-	     Says A  B {|Agent A, Nonce NA, Number SID, Number PA|}
-	       \<in> set evsCF;
+             Says A  B {|Agent A, Nonce NA, Number SID, Number PA|}
+               \<in> set evsCF;
              Says B' A {|Nonce NB, Number SID, Number PB|} \<in> set evsCF;
              Notes A {|Agent B, Nonce PMS|} \<in> set evsCF;
-	     M = PRF(PMS,NA,NB) |]
+             M = PRF(PMS,NA,NB) |]
           ==> Says A B (Crypt (clientK(NA,NB,M))
-			(Hash{|Number SID, Nonce M,
-			       Nonce NA, Number PA, Agent A,
-			       Nonce NB, Number PB, Agent B|}))
+                        (Hash{|Number SID, Nonce M,
+                               Nonce NA, Number PA, Agent A,
+                               Nonce NB, Number PB, Agent B|}))
               # evsCF  \<in>  tls"
 
  | ServerFinished:
-	--{*Keeping A' and A'' distinct means B cannot even check that the
+        --{*Keeping A' and A'' distinct means B cannot even check that the
           two messages originate from the same source. *}
          "[| evsSF \<in> tls;
-	     Says A' B  {|Agent A, Nonce NA, Number SID, Number PA|}
-	       \<in> set evsSF;
-	     Says B  A  {|Nonce NB, Number SID, Number PB|} \<in> set evsSF;
-	     Says A'' B (Crypt (pubK B) (Nonce PMS)) \<in> set evsSF;
-	     M = PRF(PMS,NA,NB) |]
+             Says A' B  {|Agent A, Nonce NA, Number SID, Number PA|}
+               \<in> set evsSF;
+             Says B  A  {|Nonce NB, Number SID, Number PB|} \<in> set evsSF;
+             Says A'' B (Crypt (pubK B) (Nonce PMS)) \<in> set evsSF;
+             M = PRF(PMS,NA,NB) |]
           ==> Says B A (Crypt (serverK(NA,NB,M))
-			(Hash{|Number SID, Nonce M,
-			       Nonce NA, Number PA, Agent A,
-			       Nonce NB, Number PB, Agent B|}))
+                        (Hash{|Number SID, Nonce M,
+                               Nonce NA, Number PA, Agent A,
+                               Nonce NB, Number PB, Agent B|}))
               # evsSF  \<in>  tls"
 
  | ClientAccepts:
-	--{*Having transmitted ClientFinished and received an identical
+        --{*Having transmitted ClientFinished and received an identical
           message encrypted with serverK, the client stores the parameters
           needed to resume this session.  The "Notes A ..." premise is
           used to prove @{text Notes_master_imp_Crypt_PMS}.*}
          "[| evsCA \<in> tls;
-	     Notes A {|Agent B, Nonce PMS|} \<in> set evsCA;
-	     M = PRF(PMS,NA,NB);
-	     X = Hash{|Number SID, Nonce M,
-	               Nonce NA, Number PA, Agent A,
-		       Nonce NB, Number PB, Agent B|};
+             Notes A {|Agent B, Nonce PMS|} \<in> set evsCA;
+             M = PRF(PMS,NA,NB);
+             X = Hash{|Number SID, Nonce M,
+                       Nonce NA, Number PA, Agent A,
+                       Nonce NB, Number PB, Agent B|};
              Says A  B (Crypt (clientK(NA,NB,M)) X) \<in> set evsCA;
              Says B' A (Crypt (serverK(NA,NB,M)) X) \<in> set evsCA |]
           ==>
              Notes A {|Number SID, Agent A, Agent B, Nonce M|} # evsCA  \<in>  tls"
 
  | ServerAccepts:
-	--{*Having transmitted ServerFinished and received an identical
+        --{*Having transmitted ServerFinished and received an identical
           message encrypted with clientK, the server stores the parameters
           needed to resume this session.  The "Says A'' B ..." premise is
           used to prove @{text Notes_master_imp_Crypt_PMS}.*}
          "[| evsSA \<in> tls;
-	     A \<noteq> B;
+             A \<noteq> B;
              Says A'' B (Crypt (pubK B) (Nonce PMS)) \<in> set evsSA;
-	     M = PRF(PMS,NA,NB);
-	     X = Hash{|Number SID, Nonce M,
-	               Nonce NA, Number PA, Agent A,
-		       Nonce NB, Number PB, Agent B|};
+             M = PRF(PMS,NA,NB);
+             X = Hash{|Number SID, Nonce M,
+                       Nonce NA, Number PA, Agent A,
+                       Nonce NB, Number PB, Agent B|};
              Says B  A (Crypt (serverK(NA,NB,M)) X) \<in> set evsSA;
              Says A' B (Crypt (clientK(NA,NB,M)) X) \<in> set evsSA |]
           ==>
@@ -244,27 +244,27 @@
          --{*If A recalls the @{text SESSION_ID}, then she sends a FINISHED
              message using the new nonces and stored MASTER SECRET.*}
          "[| evsCR \<in> tls;
-	     Says A  B {|Agent A, Nonce NA, Number SID, Number PA|}: set evsCR;
+             Says A  B {|Agent A, Nonce NA, Number SID, Number PA|}: set evsCR;
              Says B' A {|Nonce NB, Number SID, Number PB|} \<in> set evsCR;
              Notes A {|Number SID, Agent A, Agent B, Nonce M|} \<in> set evsCR |]
           ==> Says A B (Crypt (clientK(NA,NB,M))
-			(Hash{|Number SID, Nonce M,
-			       Nonce NA, Number PA, Agent A,
-			       Nonce NB, Number PB, Agent B|}))
+                        (Hash{|Number SID, Nonce M,
+                               Nonce NA, Number PA, Agent A,
+                               Nonce NB, Number PB, Agent B|}))
               # evsCR  \<in>  tls"
 
  | ServerResume:
          --{*Resumption (7.3):  If B finds the @{text SESSION_ID} then he can 
              send a FINISHED message using the recovered MASTER SECRET*}
          "[| evsSR \<in> tls;
-	     Says A' B {|Agent A, Nonce NA, Number SID, Number PA|}: set evsSR;
-	     Says B  A {|Nonce NB, Number SID, Number PB|} \<in> set evsSR;
+             Says A' B {|Agent A, Nonce NA, Number SID, Number PA|}: set evsSR;
+             Says B  A {|Nonce NB, Number SID, Number PB|} \<in> set evsSR;
              Notes B {|Number SID, Agent A, Agent B, Nonce M|} \<in> set evsSR |]
           ==> Says B A (Crypt (serverK(NA,NB,M))
-			(Hash{|Number SID, Nonce M,
-			       Nonce NA, Number PA, Agent A,
-			       Nonce NB, Number PB, Agent B|})) # evsSR
-	        \<in>  tls"
+                        (Hash{|Number SID, Nonce M,
+                               Nonce NA, Number PA, Agent A,
+                               Nonce NB, Number PB, Agent B|})) # evsSR
+                \<in>  tls"
 
  | Oops:
          --{*The most plausible compromise is of an old session key.  Losing
@@ -273,7 +273,7 @@
            otherwise the Spy could learn session keys merely by 
            replaying messages!*}
          "[| evso \<in> tls;  A \<noteq> Spy;
-	     Says A B (Crypt (sessionK((NA,NB,M),role)) X) \<in> set evso |]
+             Says A B (Crypt (sessionK((NA,NB,M),role)) X) \<in> set evso |]
           ==> Says A Spy (Key (sessionK((NA,NB,M),role))) # evso  \<in>  tls"
 
 (*
@@ -328,7 +328,7 @@
 
 
 (** These proofs assume that the Nonce_supply nonces
-	(which have the form  @ N. Nonce N \<notin> used evs)
+        (which have the form  @ N. Nonce N \<notin> used evs)
     lie outside the range of PRF.  It seems reasonable, but as it is needed
     only for the possibility theorems, it is not taken as an axiom.
 **)
@@ -381,11 +381,11 @@
           \<forall>evs. (@ N. Nonce N \<notin> used evs) \<notin> range PRF;
           A \<noteq> B |]
       ==> \<exists>NA PA NB PB X. \<exists>evs \<in> tls.
-		X = Hash{|Number SID, Nonce M,
-			  Nonce NA, Number PA, Agent A,
-			  Nonce NB, Number PB, Agent B|}  &
-		Says A B (Crypt (clientK(NA,NB,M)) X) \<in> set evs  &
-		Says B A (Crypt (serverK(NA,NB,M)) X) \<in> set evs"
+                X = Hash{|Number SID, Nonce M,
+                          Nonce NA, Number PA, Agent A,
+                          Nonce NB, Number PB, Agent B|}  &
+                Says A B (Crypt (clientK(NA,NB,M)) X) \<in> set evs  &
+                Says B A (Crypt (serverK(NA,NB,M)) X) \<in> set evs"
 apply (intro exI bexI)
 apply (rule_tac [2] tls.ClientHello
                     [THEN tls.ServerHello,
@@ -570,7 +570,7 @@
           (priK B \<in> KK | B \<in> bad)"
 apply (erule tls.induct)
 apply (simp_all (no_asm_simp)
-		del: image_insert
+                del: image_insert
                 add: image_Un [THEN sym]
                      insert_Key_image Un_assoc [THEN sym])
 txt{*Fake*}
@@ -598,16 +598,16 @@
 lemma analz_image_keys [rule_format]:
      "evs \<in> tls ==>
       \<forall>KK. KK <= range sessionK -->
-	      (Nonce N \<in> analz (Key`KK Un (spies evs))) =
-	      (Nonce N \<in> analz (spies evs))"
+              (Nonce N \<in> analz (Key`KK Un (spies evs))) =
+              (Nonce N \<in> analz (spies evs))"
 apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
 apply (safe del: iffI)
 apply (safe del: impI iffI intro!: analz_image_keys_lemma)
 apply (simp_all (no_asm_simp)               (*faster*)
                 del: image_insert imp_disjL (*reduces blow-up*)
-		add: image_Un [THEN sym]  Un_assoc [THEN sym]
-		     insert_Key_singleton
-		     range_sessionkeys_not_priK analz_image_priK)
+                add: image_Un [THEN sym]  Un_assoc [THEN sym]
+                     insert_Key_singleton
+                     range_sessionkeys_not_priK analz_image_priK)
 apply (simp_all add: insert_absorb)
 txt{*Fake*}
 apply spy_analz
@@ -901,11 +901,11 @@
            down to 433s on albatross*)
 
 (*5/5/01: conversion to Isar script
-	  loads in 137s (perch)
+          loads in 137s (perch)
           the last ML version loaded in 122s on perch, a 600MHz machine:
-		twice as fast as pike.  No idea why it's so much slower!
-	  The Isar script is slower still, perhaps because simp_all simplifies
-	  the assumptions be default.
+                twice as fast as pike.  No idea why it's so much slower!
+          The Isar script is slower still, perhaps because simp_all simplifies
+          the assumptions be default.
 *)
 
 end
--- a/src/HOL/Auth/Yahalom.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Auth/Yahalom.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -58,7 +58,7 @@
            uses the new session key to send Bob his Nonce.  The premise
            @{term "A \<noteq> Server"} is needed for @{text Says_Server_not_range}.
            Alice can check that K is symmetric by its length.*}
-	 "[| evs4 \<in> yahalom;  A \<noteq> Server;  K \<in> symKeys;
+         "[| evs4 \<in> yahalom;  A \<noteq> Server;  K \<in> symKeys;
              Gets A {|Crypt(shrK A) {|Agent B, Key K, Nonce NA, Nonce NB|}, X|}
                 \<in> set evs4;
              Says A B {|Agent A, Nonce NA|} \<in> set evs4 |]
@@ -481,9 +481,9 @@
 text{*A vital theorem for B, that nonce NB remains secure from the Spy.*}
 lemma Spy_not_see_NB :
      "[| Says B Server
-	        {|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|}
-	   \<in> set evs;
-	 (\<forall>k. Notes Spy {|Nonce NA, Nonce NB, k|} \<notin> set evs);
+                {|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|}
+           \<in> set evs;
+         (\<forall>k. Notes Spy {|Nonce NA, Nonce NB, k|} \<notin> set evs);
          A \<notin> bad;  B \<notin> bad;  evs \<in> yahalom |]
       ==> Nonce NB \<notin> analz (knows Spy evs)"
 apply (erule rev_mp, erule rev_mp)
--- a/src/HOL/Auth/Yahalom2.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Auth/Yahalom2.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -51,7 +51,7 @@
            Both agents are quoted in the 2nd certificate to prevent attacks!*)
  | YM3:  "[| evs3 \<in> yahalom;  Key KAB \<notin> used evs3;
              Gets Server {|Agent B, Nonce NB,
-			   Crypt (shrK B) {|Agent A, Nonce NA|}|}
+                           Crypt (shrK B) {|Agent A, Nonce NA|}|}
                \<in> set evs3 |]
           ==> Says Server A
                {|Nonce NB,
--- a/src/HOL/Auth/ZhouGollmann.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Auth/ZhouGollmann.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -35,7 +35,7 @@
   Nil:  "[] \<in> zg"
 
 | Fake: "[| evsf \<in> zg;  X \<in> synth (analz (spies evsf)) |]
-	 ==> Says Spy B X  # evsf \<in> zg"
+         ==> Says Spy B X  # evsf \<in> zg"
 
 | Reception:  "[| evsr \<in> zg; Says A B X \<in> set evsr |] ==> Gets B X # evsr \<in> zg"
 
@@ -44,26 +44,26 @@
     We just assume that the protocol's objective is to deliver K fairly,
     rather than to keep M secret.*)
 | ZG1: "[| evs1 \<in> zg;  Nonce L \<notin> used evs1; C = Crypt K (Number m);
-	   K \<in> symKeys;
-	   NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|}|]
+           K \<in> symKeys;
+           NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|}|]
        ==> Says A B {|Number f_nro, Agent B, Nonce L, C, NRO|} # evs1 \<in> zg"
 
   (*B must check that NRO is A's signature to learn the sender's name*)
 | ZG2: "[| evs2 \<in> zg;
-	   Gets B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs2;
-	   NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|};
-	   NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|}|]
+           Gets B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs2;
+           NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|};
+           NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|}|]
        ==> Says B A {|Number f_nrr, Agent A, Nonce L, NRR|} # evs2  \<in>  zg"
 
   (*A must check that NRR is B's signature to learn the sender's name;
     without spy, the matching label would be enough*)
 | ZG3: "[| evs3 \<in> zg; C = Crypt K M; K \<in> symKeys;
-	   Says A B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs3;
-	   Gets A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs3;
-	   NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|};
-	   sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|}|]
+           Says A B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs3;
+           Gets A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs3;
+           NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|};
+           sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|}|]
        ==> Says A TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|}
-	     # evs3 \<in> zg"
+             # evs3 \<in> zg"
 
  (*TTP checks that sub_K is A's signature to learn who issued K, then
    gives credentials to A and B.  The Notes event models the availability of
@@ -72,15 +72,15 @@
    also allowing lemma @{text Crypt_used_imp_spies} to omit the condition
    @{term "K \<noteq> priK TTP"}. *)
 | ZG4: "[| evs4 \<in> zg; K \<in> symKeys;
-	   Gets TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|}
-	     \<in> set evs4;
-	   sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|};
-	   con_K = Crypt (priK TTP) {|Number f_con, Agent A, Agent B,
-				      Nonce L, Key K|}|]
+           Gets TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|}
+             \<in> set evs4;
+           sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|};
+           con_K = Crypt (priK TTP) {|Number f_con, Agent A, Agent B,
+                                      Nonce L, Key K|}|]
        ==> Says TTP Spy con_K
            #
-	   Notes TTP {|Number f_con, Agent A, Agent B, Nonce L, Key K, con_K|}
-	   # evs4 \<in> zg"
+           Notes TTP {|Number f_con, Agent A, Agent B, Nonce L, Key K, con_K|}
+           # evs4 \<in> zg"
 
 
 declare Says_imp_knows_Spy [THEN analz.Inj, dest]
--- a/src/HOL/Bali/AxCompl.thy	Wed Oct 21 16:54:04 2009 +0200
+++ b/src/HOL/Bali/AxCompl.thy	Wed Oct 21 16:57:57 2009 +0200
@@ -319,13 +319,13 @@
       case 0
       with is_cls
       show ?thesis
-	by - (rule ax_impossible [THEN conseq1],fastsimp dest: nyinitcls_emptyD)
+        by - (rule ax_impossible [THEN conseq1],fastsimp dest: nyinitcls_emptyD)
     next
       case (Suc m)
       with mgf_hyp have mgf_hyp': "\<And> t. G,A\<turnstile>{=:m} t\<succ> {G\<rightarrow>}"
-	by simp
+        by simp
       from is_cls obtain c where c: "the (class G C) = c"
-	by auto
+        by auto
       let ?Q= "(\<lambda>Y s' (x,s) . 
           G\<turnstile> (x,init_class_obj G C s) 
              \<midarrow> (if C=Object then Skip else Init (super (the (class G C))))\<rightarrow> s'
@@ -333,45 +333,45 @@
       from c
       show ?thesis
       proof (rule ax_derivs.Init [where ?Q="?Q"])
-	let ?P' = "Normal ((\<lambda>Y s' s. s' = supd (init_class_obj G C) s 
+        let ?P' = "Normal ((\<lambda>Y s' s. s' = supd (init_class_obj G C) s 
                            \<and> normal s \<and> \<not> initd C s) \<and>. G\<turnstile>init\<le>m)" 
-	show "G,A\<turnstile>{Normal (?P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C))}
+        show "G,A\<turnstile>{Normal (?P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C))}
                   .(if C = Object then Skip else Init (super c)). 
                   {?Q}"
-	proof (rule conseq1 [where ?P'="?P'"])
-	  show "G,A\<turnstile>{?P'} .(if C = Object then Skip else Init (super c)). {?Q}"
-	  proof (cases "C=Object")
-	    case True
-	    have "G,A\<turnstile>{?P'} .Skip. {?Q}"
-	      by (rule ax_derivs.Skip [THEN conseq1])
-	         (auto simp add: True intro: eval.Skip)
+        proof (rule conseq1 [where ?P'="?P'"])
+          show "G,A\<turnstile>{?P'} .(if C = Object then Skip else Init (super c)). {?Q}"
+          proof (cases "C=Object")
+            case True
+            have "G,A\<turnstile>{?P'} .Skip. {?Q}"
+              by (rule ax_derivs.Skip [THEN conseq1])
+                 (auto simp add: True intro: eval.Skip)
             with True show ?thesis 
-	      by simp
-	  next
-	    case False
-	    from mgf_hyp'
-	    have "G,A\<turnstile>{?P'} .Init (super c). {?Q}"
-	      by (rule MGFnD' [THEN conseq12]) (fastsimp simp add: False c)
-	    with False show ?thesis
-	      by simp
-	  qed
-	next
-	  from Suc is_cls
-	  show "Normal (?P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C))
+              by simp
+          next
+            case False
+            from mgf_hyp'
+            have "G,A\<turnstile>{?P'} .Init (super c). {?Q}"
+              by (rule MGFnD' [THEN conseq12]) (fastsimp simp add: False c)
+            with False show ?thesis
+              by simp
+          qed
+        next
+          from Suc is_cls
+          show "Normal (?P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C))
                 \<Rightarrow> ?P'"
-	    by (fastsimp elim: nyinitcls_le_SucD)
-	qed
+            by (fastsimp elim: nyinitcls_le_SucD)
+        qed
       next
-	from mgf_hyp'
-	show "\<forall>l. G,A\<turnstile>{?Q \<and>. (\<lambda>s. l = locals (snd s)) ;. set_lvars empty} 
+        from mgf_hyp'
+        show "\<forall>l. G,A\<turnstile>{?Q \<and>. (\<lambda>s. l = locals (snd s)) ;. set_lvars empty} 
                       .init c.
                       {set_lvars l .; ?R}"
-	  apply (rule MGFnD' [THEN conseq12, THEN allI])
-	  apply (clarsimp simp add: split_paired_all)
-	  apply (rule eval.Init [OF c])
-	  apply (insert c)
-	  apply auto
-	  done
+          apply (rule MGFnD' [THEN conseq12, THEN allI])
+          apply (clarsimp simp add: split_paired_all)
+          apply (rule eval.Init [OF c])
+          apply (insert c)
+          apply auto
+          done
       qed
     qed
     thus "G,A\<turnstile>{Normal ?P  \<and>. Not \<circ> initd C} .Init C. {?R}"
@@ -399,7 +399,7 @@
                  mode: "mode = invmode statM e" and
                     T: "T =Inl (resTy statM)" and
         eq_accC_accC': "accC=accC'"
-	by cases fastsimp+
+        by cases fastsimp+
   let ?Q="(\<lambda>Y s1 (x,s) . x = None \<and> 
               (\<exists>a. G\<turnstile>Norm s \<midarrow>e-\<succ>a\<rightarrow> s1 \<and> 
                    (normal s1 \<longrightarrow> G, store s1\<turnstile>a\<Colon>\<preceq>RefT statT)
@@ -435,29 +435,29 @@
               (\<exists>P. \<lparr>prg=G,cls=accC',lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P))
             \<and> s1\<Colon>\<preceq>(G, L)"
       proof -
-	from da obtain C where
-	  da_e:  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
+        from da obtain C where
+          da_e:  "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
                     dom (locals (store ((Norm s0)::state)))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> C" and
-	  da_ps: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm C \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> E" 
-	  by cases (simp add: eq_accC_accC')
-	from eval_e conf_s0 wt_e da_e wf
-	obtain "(abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT)"
-	  and  "s1\<Colon>\<preceq>(G, L)"
-	  by (rule eval_type_soundE) simp
-	moreover
-	{
-	  assume normal_s1: "normal s1"
-	  have "\<exists>P. \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P"
-	  proof -
-	    from eval_e wt_e da_e wf normal_s1
-	    have "nrm C \<subseteq>  dom (locals (store s1))"
-	      by (cases rule: da_good_approxE') iprover
-	    with da_ps show ?thesis
-	      by (rule da_weakenE) iprover
-	  qed
-	}
-	ultimately show ?thesis
-	  using eq_accC_accC' by simp
+          da_ps: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm C \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> E" 
+          by cases (simp add: eq_accC_accC')
+        from eval_e conf_s0 wt_e da_e wf
+        obtain "(abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT)"
+          and  "s1\<Colon>\<preceq>(G, L)"
+          by (rule eval_type_soundE) simp
+        moreover
+        {
+          assume normal_s1: "normal s1"
+          have "\<exists>P. \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P"
+          proof -
+            from eval_e wt_e da_e wf normal_s1
+            have "nrm C \<subseteq>  dom (locals (store s1))"
+              by (cases rule: da_good_approxE') iprover
+            with da_ps show ?thesis
+              by (rule da_weakenE) iprover
+          qed
+        }
+        ultimately show ?thesis
+          using eq_accC_accC' by simp
       qed
     qed
   next
@@ -467,36 +467,36 @@
       show "?PS a"
       proof (rule MGFnD' [OF mgf_ps, THEN conseq12],
              clarsimp simp add: eq_accC_accC' [symmetric])
-	fix s0 s1 s2 vs
-	assume conf_s1: "s1\<Colon>\<preceq>(G, L)"
-	assume eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
-	assume conf_a: "abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
-	assume eval_ps: "G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2"
-	assume da_ps: "abrupt s1 = None \<longrightarrow> 
+        fix s0 s1 s2 vs
+        assume conf_s1: "s1\<Colon>\<preceq>(G, L)"
+        assume eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
+        assume conf_a: "abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
+        assume eval_ps: "G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2"
+        assume da_ps: "abrupt s1 = None \<longrightarrow> 
                        (\<exists>P. \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> 
                                dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P)"
-	show "(\<exists>s1. G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1 \<and>
+        show "(\<exists>s1. G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1 \<and>
                 (abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT) \<and>
                 G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2) \<and>
               s2\<Colon>\<preceq>(G, L)"
-	proof (cases "normal s1")
-	  case True
-	  with da_ps obtain P where
-	   "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P"
-	    by auto
-	  from eval_ps conf_s1 wt_args this wf
-	  have "s2\<Colon>\<preceq>(G, L)"
-	    by (rule eval_type_soundE)
-	  with eval_e conf_a eval_ps 
-	  show ?thesis 
-	    by auto
-	next
-	  case False
-	  with eval_ps have "s2=s1" by auto
-	  with eval_e conf_a eval_ps conf_s1 
-	  show ?thesis 
-	    by auto
-	qed
+        proof (cases "normal s1")
+          case True
+          with da_ps obtain P where
+           "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P"
+            by auto
+          from eval_ps conf_s1 wt_args this wf
+          have "s2\<Colon>\<preceq>(G, L)"
+            by (rule eval_type_soundE)
+          with eval_e conf_a eval_ps 
+          show ?thesis 
+            by auto
+        next
+          case False
+          with eval_ps have "s2=s1" by auto
+          with eval_e conf_a eval_ps conf_s1 
+          show ?thesis 
+            by auto
+        qed
       qed
     qed
   next
@@ -517,52 +517,52 @@
       from mgf_methds [rule_format]
       show "?METHD a vs invC declC l"
       proof (rule MGFnD' [THEN conseq12],clarsimp)
-	fix s4 s2 s1::state
-	fix s0 v
-	let ?D= "invocation_declclass G mode (store s2) a statT 
+        fix s4 s2 s1::state
+        fix s0 v
+        let ?D= "invocation_declclass G mode (store s2) a statT 
                     \<lparr>name=mn,parTs=pTs'\<rparr>"
-	let ?s3= "init_lvars G ?D \<lparr>name=mn, parTs=pTs'\<rparr> mode a vs s2"
-	assume inv_prop: "abrupt ?s3=None 
+        let ?s3= "init_lvars G ?D \<lparr>name=mn, parTs=pTs'\<rparr> mode a vs s2"
+        assume inv_prop: "abrupt ?s3=None 
              \<longrightarrow> G\<turnstile>mode\<rightarrow>invocation_class mode (store s2) a statT\<preceq>statT"
-	assume conf_s2: "s2\<Colon>\<preceq>(G, L)"
-	assume conf_a: "abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
-	assume eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
-	assume eval_ps: "G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2"
-	assume eval_mthd: "G\<turnstile>?s3 \<midarrow>Methd ?D \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ>v\<rightarrow> s4"
-	show "G\<turnstile>Norm s0 \<midarrow>{accC,statT,mode}e\<cdot>mn( {pTs'}ps)-\<succ>v
+        assume conf_s2: "s2\<Colon>\<preceq>(G, L)"
+        assume conf_a: "abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
+        assume eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
+        assume eval_ps: "G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2"
+        assume eval_mthd: "G\<turnstile>?s3 \<midarrow>Methd ?D \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ>v\<rightarrow> s4"
+        show "G\<turnstile>Norm s0 \<midarrow>{accC,statT,mode}e\<cdot>mn( {pTs'}ps)-\<succ>v
                         \<rightarrow> (set_lvars (locals (store s2))) s4"
-	proof -
-	  obtain D where D: "D=?D" by simp
-	  obtain s3 where s3: "s3=?s3" by simp
-	  obtain s3' where 
-	    s3': "s3' = check_method_access G accC statT mode 
+        proof -
+          obtain D where D: "D=?D" by simp
+          obtain s3 where s3: "s3=?s3" by simp
+          obtain s3' where 
+            s3': "s3' = check_method_access G accC statT mode 
                            \<lparr>name=mn,parTs=pTs'\<rparr> a s3"
-	    by simp
-	  have eq_s3'_s3: "s3'=s3"
-	  proof -
-	    from inv_prop s3 mode
-	    have "normal s3 \<Longrightarrow> 
+            by simp
+          have eq_s3'_s3: "s3'=s3"
+          proof -
+            from inv_prop s3 mode
+            have "normal s3 \<Longrightarrow> 
              G\<turnstile>invmode statM e\<rightarrow>invocation_class mode (store s2) a statT\<preceq>statT"
-	      by auto
-	    with eval_ps wt_e statM conf_s2 conf_a [rule_format] 
-	    have "check_method_access G accC statT (invmode statM e)
+              by auto
+            with eval_ps wt_e statM conf_s2 conf_a [rule_format] 
+            have "check_method_access G accC statT (invmode statM e)
                       \<lparr>name=mn,parTs=pTs'\<rparr> a s3 = s3"
-	      by (rule error_free_call_access) (auto simp add: s3 mode wf)
-	    thus ?thesis 
-	      by (simp add: s3' mode)
-	  qed
-	  with eval_mthd D s3
-	  have "G\<turnstile>s3' \<midarrow>Methd D \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ>v\<rightarrow> s4"
-	    by simp
-	  with eval_e eval_ps D _ s3' 
-	  show ?thesis
-	    by (rule eval_Call) (auto simp add: s3 mode D)
-	qed
+              by (rule error_free_call_access) (auto simp add: s3 mode wf)
+            thus ?thesis 
+              by (simp add: s3' mode)
+          qed
+          with eval_mthd D s3
+          have "G\<turnstile>s3' \<midarrow>Methd D \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ>v\<rightarrow> s4"
+            by simp
+          with eval_e eval_ps D _ s3' 
+          show ?thesis
+            by (rule eval_Call) (auto simp add: s3 mode D)
+        qed
       qed
     qed
   qed
 qed
-	  	  
+                  
 lemma eval_expression_no_jump':
   assumes eval: "G\<turnstile>s0 \<midarrow>e-\<succ>v\<rightarrow> s1"
   and   no_jmp: "abrupt s0 \<noteq> Some (Jump j)"
@@ -610,36 +610,36 @@
       case True
       with normal_t eval_e normal_termination
       show ?thesis
-	by (auto intro: eval.Loop)
+        by (auto intro: eval.Loop)
     next
       case False
       note abrupt_s' = this
       from eval_e _ wt wf
       have no_cont: "abrupt s' \<noteq> Some (Jump (Cont l))"
-	by (rule eval_expression_no_jump') (insert normal_t,simp)
+        by (rule eval_expression_no_jump') (insert normal_t,simp)
       have
-	"if the_Bool v 
+        "if the_Bool v 
              then (G\<turnstile>s' \<midarrow>c\<rightarrow> s' \<and> 
                    G\<turnstile>(abupd (absorb (Cont l)) s') \<midarrow>l\<bullet> While(e) c\<rightarrow> s')
-	     else s' = s'"
+             else s' = s'"
       proof (cases "the_Bool v")
-	case False thus ?thesis by simp
+        case False thus ?thesis by simp
       next
-	case True
-	with abrupt_s' have "G\<turnstile>s' \<midarrow>c\<rightarrow> s'" by auto
-	moreover from abrupt_s' no_cont 
-	have no_absorb: "(abupd (absorb (Cont l)) s')=s'"
-	  by (cases s') (simp add: absorb_def split: split_if)
-	moreover
-	from no_absorb abrupt_s'
-	have "G\<turnstile>(abupd (absorb (Cont l)) s') \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
-	  by auto
-	ultimately show ?thesis
-	  using True by simp
+        case True
+        with abrupt_s' have "G\<turnstile>s' \<midarrow>c\<rightarrow> s'" by auto
+        moreover from abrupt_s' no_cont 
+        have no_absorb: "(abupd (absorb (Cont l)) s')=s'"
+          by (cases s') (simp add: absorb_def split: split_if)
+        moreover
+        from no_absorb abrupt_s'
+        have "G\<turnstile>(abupd (absorb (Cont l)) s') \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
+          by auto
+        ultimately show ?thesis
+          using True by simp
       qed
       with eval_e 
       show ?thesis
-	using normal_t by (auto intro: eval.Loop)
+        using normal_t by (auto intro: eval.Loop)
     qed
   qed
 next
@@ -690,49 +690,49 @@
                                   Y = In1 b \<and> G\<turnstile>t \<midarrow>e-\<succ>b\<rightarrow> s')) 
                               \<and>. G\<turnstile>init\<le>n)\<leftarrow>=False\<down>=\<diamondsuit>}"
       proof (rule ax_derivs.Loop)
-	from mfg_e
-	show "G,A\<turnstile>{(\<lambda>Y s' s. (s, s') \<in> (unroll G l e c)\<^sup>*) \<and>. G\<turnstile>init\<le>n} 
+        from mfg_e
+        show "G,A\<turnstile>{(\<lambda>Y s' s. (s, s') \<in> (unroll G l e c)\<^sup>*) \<and>. G\<turnstile>init\<le>n} 
                    e-\<succ>
                   {(\<lambda>Y s' s. (\<exists>t b. (s, t) \<in> (unroll G l e c)\<^sup>* \<and> 
                                      Y = In1 b \<and> G\<turnstile>t \<midarrow>e-\<succ>b\<rightarrow> s')) 
                    \<and>. G\<turnstile>init\<le>n}"
-	proof (rule MGFnD' [THEN conseq12],clarsimp)
-	  fix s Z s' v
-	  assume "(Z, s) \<in> (unroll G l e c)\<^sup>*"
-	  moreover
-	  assume "G\<turnstile>s \<midarrow>e-\<succ>v\<rightarrow> s'"
-	  ultimately
-	  show "\<exists>t. (Z, t) \<in> (unroll G l e c)\<^sup>* \<and> G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> s'"
-	    by blast
-	qed
+        proof (rule MGFnD' [THEN conseq12],clarsimp)
+          fix s Z s' v
+          assume "(Z, s) \<in> (unroll G l e c)\<^sup>*"
+          moreover
+          assume "G\<turnstile>s \<midarrow>e-\<succ>v\<rightarrow> s'"
+          ultimately
+          show "\<exists>t. (Z, t) \<in> (unroll G l e c)\<^sup>* \<and> G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> s'"
+            by blast
+        qed
       next
-	from mfg_c
-	show "G,A\<turnstile>{Normal (((\<lambda>Y s' s. \<exists>t b. (s, t) \<in> (unroll G l e c)\<^sup>* \<and>
+        from mfg_c
+        show "G,A\<turnstile>{Normal (((\<lambda>Y s' s. \<exists>t b. (s, t) \<in> (unroll G l e c)\<^sup>* \<and>
                                        Y = \<lfloor>b\<rfloor>\<^sub>e \<and> G\<turnstile>t \<midarrow>e-\<succ>b\<rightarrow> s') 
                           \<and>. G\<turnstile>init\<le>n)\<leftarrow>=True)}
                   .c.
                   {abupd (absorb (Cont l)) .;
                    ((\<lambda>Y s' s. (s, s') \<in> (unroll G l e c)\<^sup>*) \<and>. G\<turnstile>init\<le>n)}"
-	proof (rule MGFnD' [THEN conseq12],clarsimp)
-	  fix Z s' s v t
-	  assume unroll: "(Z, t) \<in> (unroll G l e c)\<^sup>*"
-	  assume eval_e: "G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> Norm s" 
-	  assume true: "the_Bool v"
-	  assume eval_c: "G\<turnstile>Norm s \<midarrow>c\<rightarrow> s'"
-	  show "(Z, abupd (absorb (Cont l)) s') \<in> (unroll G l e c)\<^sup>*"
-	  proof -
-	    note unroll
-	    also
-	    from eval_e true eval_c
-	    have "(t,abupd (absorb (Cont l)) s') \<in> unroll G l e c" 
-	      by (unfold unroll_def) force
-	    ultimately show ?thesis ..
-	  qed
-	qed
+        proof (rule MGFnD' [THEN conseq12],clarsimp)
+          fix Z s' s v t
+          assume unroll: "(Z, t) \<in> (unroll G l e c)\<^sup>*"
+          assume eval_e: "G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> Norm s" 
+          assume true: "the_Bool v"
+          assume eval_c: "G\<turnstile>Norm s \<midarrow>c\<rightarrow> s'"
+          show "(Z, abupd (absorb (Cont l)) s') \<in> (unroll G l e c)\<^sup>*"
+          proof -
+            note unroll
+            also
+            from eval_e true eval_c
+            have "(t,abupd (absorb (Cont l)) s') \<in> unroll G l e c" 
+              by (unfold unroll_def) force
+            ultimately show ?thesis ..
+          qed
+        qed
       qed
     next
       show 
-	"\<forall>Y s Z.
+        "\<forall>Y s Z.
          (Normal ((\<lambda>Y' s' s. s' = s \<and> normal s) \<and>. G\<turnstile>init\<le>n)) Y s Z 
          \<longrightarrow> (\<forall>Y' s'.
                (\<forall>Y Z'. 
@@ -742,28 +742,28 @@
                      \<and>. G\<turnstile>init\<le>n)\<leftarrow>=False\<down>=\<diamondsuit>) Y' s' Z') 
                \<longrightarrow> G\<turnstile>Z \<midarrow>\<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<succ>\<rightarrow> (Y', s'))"
       proof (clarsimp)
-	fix Y' s' s
-	assume asm:
-	  "\<forall>Z'. (Z', Norm s) \<in> (unroll G l e c)\<^sup>* 
+        fix Y' s' s
+        assume asm:
+          "\<forall>Z'. (Z', Norm s) \<in> (unroll G l e c)\<^sup>* 
                  \<longrightarrow> card (nyinitcls G s') \<le> n \<and>
                      (\<exists>v. (\<exists>t. (Z', t) \<in> (unroll G l e c)\<^sup>* \<and> G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> s') \<and>
                      (fst s' = None \<longrightarrow> \<not> the_Bool v)) \<and> Y' = \<diamondsuit>"
-	show "Y' = \<diamondsuit> \<and> G\<turnstile>Norm s \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
-	proof -
-	  from asm obtain v t where 
-	    -- {* @{term "Z'"} gets instantiated with @{term "Norm s"} *}  
-	    unroll: "(Norm s, t) \<in> (unroll G l e c)\<^sup>*" and
+        show "Y' = \<diamondsuit> \<and> G\<turnstile>Norm s \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
+        proof -
+          from asm obtain v t where 
+            -- {* @{term "Z'"} gets instantiated with @{term "Norm s"} *}  
+            unroll: "(Norm s, t) \<in> (unroll G l e c)\<^sup>*" and
             eval_e: "G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> s'" and
             normal_termination: "normal s' \<longrightarrow> \<not> the_Bool v" and
-	     Y': "Y' = \<diamondsuit>"
-	    by auto
-	  from unroll eval_e normal_termination wt_e wf
-	  have "G\<turnstile>Norm s \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
-	    by (rule unroll_while)
-	  with Y' 
-	  show ?thesis
-	    by simp
-	qed
+             Y': "Y' = \<diamondsuit>"
+            by auto
+          from unroll eval_e normal_termination wt_e wf
+          have "G\<turnstile>Norm s \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
+            by (rule unroll_while)
+          with Y' 
+          show ?thesis
+            by simp
+        qed
       qed
     qed
   qed
@@ -810,39 +810,39 @@
       show "(\<exists>E. \<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile> dom (locals (store s')) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E) \<and>
             s'\<Colon>\<preceq>(G, L)"
       proof -
-	from da 
-	obtain E where
-	  "\<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile> dom (locals s) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
-	  by cases simp
-	moreover
-	from eval_init
-	have "dom (locals s) \<subseteq> dom (locals (store s'))"
-	  by (rule dom_locals_eval_mono [elim_format]) simp
-	ultimately obtain E' where
-	  "\<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile> dom (locals (store s')) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
-	  by (rule da_weakenE)
-	moreover
-	have "s'\<Colon>\<preceq>(G, L)"
-	proof -
-	  have wt_init: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(Init statDeclC)\<Colon>\<surd>"
-	  proof -
-	    from wf wt_e 
-	    have iscls_statC: "is_class G statC"
-	      by (auto dest: ty_expr_is_type type_is_class)
-	    with wf accfield 
-	    have iscls_statDeclC: "is_class G statDeclC"
-	      by (auto dest!: accfield_fields dest: fields_declC)
-	    thus ?thesis by simp
-	  qed
-	  obtain I where 
-	    da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+        from da 
+        obtain E where
+          "\<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile> dom (locals s) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+          by cases simp
+        moreover
+        from eval_init
+        have "dom (locals s) \<subseteq> dom (locals (store s'))"
+          by (rule dom_locals_eval_mono [elim_format]) simp
+        ultimately obtain E' where
+          "\<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile> dom (locals (store s')) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
+          by (rule da_weakenE)
+        moreover
+        have "s'\<Colon>\<preceq>(G, L)"
+        proof -
+          have wt_init: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(Init statDeclC)\<Colon>\<surd>"
+          proof -
+            from wf wt_e 
+            have iscls_statC: "is_class G statC"
+              by (auto dest: ty_expr_is_type type_is_class)
+            with wf accfield 
+            have iscls_statDeclC: "is_class G statDeclC"
+              by (auto dest!: accfield_fields dest: fields_declC)
+            thus ?thesis by simp
+          qed
+          obtain I where 
+            da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
                \<turnstile> dom (locals (store ((Norm s)::state))) \<guillemotright>\<langle>Init statDeclC\<rangle>\<^sub>s\<guillemotright> I"
-	    by (auto intro: da_Init [simplified] assigned.select_convs)
-	  from eval_init conf_s wt_init da_init  wf
-	  show ?thesis
-	    by (rule eval_type_soundE)
-	qed
-	ultimately show ?thesis by iprover
+            by (auto intro: da_Init [simplified] assigned.select_convs)
+          from eval_init conf_s wt_init da_init  wf
+          show ?thesis
+            by (rule eval_type_soundE)
+        qed
+        ultimately show ?thesis by iprover
       qed
     qed
   next
@@ -857,30 +857,30 @@
       assume da_e: "\<lparr>prg=G,cls=accC',lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
       show "G\<turnstile>Norm s0 \<midarrow>{accC,statDeclC,stat}e..fn=\<succ>fst ?fvar\<rightarrow> snd ?fvar"
       proof -
-	obtain v s2' where
-	  v: "v=fst ?fvar" and s2': "s2'=snd ?fvar"
-	  by simp
-	obtain s3 where
-	  s3: "s3= check_field_access G accC' statDeclC fn stat a s2'"
-	  by simp
-	have eq_s3_s2': "s3=s2'"
-	proof -
-	  from eval_e conf_s1 wt_e da_e wf obtain
-	    conf_s2: "s2\<Colon>\<preceq>(G, L)"  and
-	    conf_a: "normal s2 \<Longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
-	    by (rule eval_type_soundE) simp
-	  from accfield wt_e eval_init eval_e conf_s2 conf_a _ wf
-	  show ?thesis
-	    by (rule  error_free_field_access 
+        obtain v s2' where
+          v: "v=fst ?fvar" and s2': "s2'=snd ?fvar"
+          by simp
+        obtain s3 where
+          s3: "s3= check_field_access G accC' statDeclC fn stat a s2'"
+          by simp
+        have eq_s3_s2': "s3=s2'"
+        proof -
+          from eval_e conf_s1 wt_e da_e wf obtain
+            conf_s2: "s2\<Colon>\<preceq>(G, L)"  and
+            conf_a: "normal s2 \<Longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
+            by (rule eval_type_soundE) simp
+          from accfield wt_e eval_init eval_e conf_s2 conf_a _ wf
+          show ?thesis
+            by (rule  error_free_field_access 
                       [where ?v=v and ?s2'=s2',elim_format])
-	       (simp add: s3 v s2' stat)+
+               (simp add: s3 v s2' stat)+
         qed
-	from eval_init eval_e 
-	show ?thesis
-	  apply (rule eval.FVar [where ?s2'=s2'])
-	  apply  (simp add: s2')
-	  apply  (simp add: s3 [symmetric]   eq_s3_s2' eq_accC s2' [symmetric])
-	  done
+        from eval_init eval_e 
+        show ?thesis
+          apply (rule eval.FVar [where ?s2'=s2'])
+          apply  (simp add: s2')
+          apply  (simp add: s3 [symmetric]   eq_s3_s2' eq_accC s2' [symmetric])
+          done
       qed
     qed
   qed
@@ -918,7 +918,7 @@
       fix s0
       assume "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals s0) \<guillemotright>\<langle>c1 Finally c2\<rangle>\<^sub>s\<guillemotright> C"
       thus "\<exists>C1. \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals s0) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1"
-	by cases (auto simp add: inj_term_simps)
+        by cases (auto simp add: inj_term_simps)
     qed
   next
     from mgf_c2
@@ -933,25 +933,25 @@
       show "G\<turnstile>Norm s0 \<midarrow>c1 Finally c2
                \<rightarrow> abupd (abrupt_if (\<exists>y. abrupt s1 = Some y) (abrupt s1)) s2"
       proof -
-	obtain abr1 str1 where s1: "s1=(abr1,str1)"
-	  by (cases s1)
-	with eval_c1 eval_c2 obtain
-	  eval_c1': "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (abr1,str1)" and
-	  eval_c2': "G\<turnstile>Norm str1 \<midarrow>c2\<rightarrow> s2"
-	  by simp
-	obtain s3 where 
-	  s3: "s3 = (if \<exists>err. abr1 = Some (Error err) 
-	                then (abr1, str1)
+        obtain abr1 str1 where s1: "s1=(abr1,str1)"
+          by (cases s1)
+        with eval_c1 eval_c2 obtain
+          eval_c1': "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (abr1,str1)" and
+          eval_c2': "G\<turnstile>Norm str1 \<midarrow>c2\<rightarrow> s2"
+          by simp
+        obtain s3 where 
+          s3: "s3 = (if \<exists>err. abr1 = Some (Error err) 
+                        then (abr1, str1)
                         else abupd (abrupt_if (abr1 \<noteq> None) abr1) s2)"
-	  by simp
-	from eval_c1' conf_s0 wt_c1 _ wf 
-	have "error_free (abr1,str1)"
-	  by (rule eval_type_soundE) (insert da_c1,auto)
-	with s3 have eq_s3: "s3=abupd (abrupt_if (abr1 \<noteq> None) abr1) s2"
-	  by (simp add: error_free_def)
-	from eval_c1' eval_c2' s3
-	show ?thesis
-	  by (rule eval.Fin [elim_format]) (simp add: s1 eq_s3)
+          by simp
+        from eval_c1' conf_s0 wt_c1 _ wf 
+        have "error_free (abr1,str1)"
+          by (rule eval_type_soundE) (insert da_c1,auto)
+        with s3 have eq_s3: "s3=abupd (abrupt_if (abr1 \<noteq> None) abr1) s2"
+          by (simp add: error_free_def)
+        from eval_c1' eval_c2' s3
+        show ?thesis
+          by (rule eval.Fin [elim_format]) (simp add: s1 eq_s3)
       qed
     qed 
   qed
@@ -1009,7 +1009,7 @@
       fix s0
       assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals s0) \<guillemotright>\<langle>Body D c\<rangle>\<^sub>e\<guillemotright> E"
       thus "jumpNestingOkS {Ret} c"
-	by cases simp
+        by cases simp
     qed
   next
     from mgf_c
@@ -1022,22 +1022,22 @@
       show "G\<turnstile>Norm s0 \<midarrow>Body D c-\<succ>the (locals (store s2) Result)
               \<rightarrow> abupd (absorb Ret) s2"
       proof -
-	from wt obtain d where 
+        from wt obtain d where 
           d: "class G D=Some d" and
           wt_c: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c\<Colon>\<surd>"
-	  by cases auto
-	obtain s3 where 
-	  s3: "s3= (if \<exists>l. fst s2 = Some (Jump (Break l)) \<or>
+          by cases auto
+        obtain s3 where 
+          s3: "s3= (if \<exists>l. fst s2 = Some (Jump (Break l)) \<or>
                            fst s2 = Some (Jump (Cont l))
                        then abupd (\<lambda>x. Some (Error CrossMethodJump)) s2 
                        else s2)"
-	  by simp
-	from eval_init eval_c nestingOk wt_c d wf
-	have eq_s3_s2: "s3=s2"
-	  by (rule Body_no_break [elim_format]) (simp add: s3)
-	from eval_init eval_c s3
-	show ?thesis
-	  by (rule eval.Body [elim_format]) (simp add: eq_s3_s2)
+          by simp
+        from eval_init eval_c nestingOk wt_c d wf
+        have eq_s3_s2: "s3=s2"
+          by (rule Body_no_break [elim_format]) (simp add: s3)
+        from eval_init eval_c s3
+        show ?thesis
+          by (rule eval.Body [elim_format]) (simp add: eq_s3_s2)
       qed
     qed
   qed
@@ -1062,312 +1062,312 @@
     proof (induct rule: var_expr_stmt.inducts)
       case (LVar v)
       show "G,A\<turnstile>{=:n} \<langle>LVar v\<rangle>\<^sub>v\<succ> {G\<rightarrow>}"
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.LVar [THEN conseq1])
-	apply (clarsimp)
-	apply (rule eval.LVar)
-	done
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.LVar [THEN conseq1])
+        apply (clarsimp)
+        apply (rule eval.LVar)
+        done
     next
       case (FVar accC statDeclC stat e fn)
       from MGFn_Init [OF hyp] and `G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}` and wf
       show ?case
-	by (rule MGFn_FVar)
+        by (rule MGFn_FVar)
     next
       case (AVar e1 e2)
       note mgf_e1 = `G,A\<turnstile>{=:n} \<langle>e1\<rangle>\<^sub>e\<succ> {G\<rightarrow>}`
       note mgf_e2 = `G,A\<turnstile>{=:n} \<langle>e2\<rangle>\<^sub>e\<succ> {G\<rightarrow>}`
       show "G,A\<turnstile>{=:n} \<langle>e1.[e2]\<rangle>\<^sub>v\<succ> {G\<rightarrow>}"
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.AVar)
-	apply  (rule MGFnD [OF mgf_e1, THEN ax_NormalD])
-	apply (rule allI)
-	apply (rule MGFnD' [OF mgf_e2, THEN conseq12])
-	apply (fastsimp intro: eval.AVar)
-	done
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.AVar)
+        apply  (rule MGFnD [OF mgf_e1, THEN ax_NormalD])
+        apply (rule allI)
+        apply (rule MGFnD' [OF mgf_e2, THEN conseq12])
+        apply (fastsimp intro: eval.AVar)
+        done
     next
       case (InsInitV c v)
       show ?case
-	by (rule MGFn_NormalI) (rule ax_derivs.InsInitV)
+        by (rule MGFn_NormalI) (rule ax_derivs.InsInitV)
     next
       case (NewC C)
       show ?case
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.NewC)
-	apply (rule MGFn_InitD [OF hyp, THEN conseq2])
-	apply (fastsimp intro: eval.NewC)
-	done
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.NewC)
+        apply (rule MGFn_InitD [OF hyp, THEN conseq2])
+        apply (fastsimp intro: eval.NewC)
+        done
     next
       case (NewA T e)
       thus ?case
-	apply -
-	apply (rule MGFn_NormalI) 
-	apply (rule ax_derivs.NewA 
+        apply -
+        apply (rule MGFn_NormalI) 
+        apply (rule ax_derivs.NewA 
                [where ?Q = "(\<lambda>Y' s' s. normal s \<and> G\<turnstile>s \<midarrow>In1r (init_comp_ty T) 
                               \<succ>\<rightarrow> (Y',s')) \<and>. G\<turnstile>init\<le>n"])
-	apply  (simp add: init_comp_ty_def split add: split_if)
-	apply  (rule conjI, clarsimp)
-	apply   (rule MGFn_InitD [OF hyp, THEN conseq2])
-	apply   (clarsimp intro: eval.Init)
-	apply  clarsimp
-	apply  (rule ax_derivs.Skip [THEN conseq1])
-	apply  (clarsimp intro: eval.Skip)
-	apply (erule MGFnD' [THEN conseq12])
-	apply (fastsimp intro: eval.NewA)
-	done
+        apply  (simp add: init_comp_ty_def split add: split_if)
+        apply  (rule conjI, clarsimp)
+        apply   (rule MGFn_InitD [OF hyp, THEN conseq2])
+        apply   (clarsimp intro: eval.Init)
+        apply  clarsimp
+        apply  (rule ax_derivs.Skip [THEN conseq1])
+        apply  (clarsimp intro: eval.Skip)
+        apply (erule MGFnD' [THEN conseq12])
+        apply (fastsimp intro: eval.NewA)
+        done
     next
       case (Cast C e)
       thus ?case
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Cast])
-	apply (fastsimp intro: eval.Cast)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Cast])
+        apply (fastsimp intro: eval.Cast)
+        done
     next
       case (Inst e C)
       thus ?case
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Inst])
-	apply (fastsimp intro: eval.Inst)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Inst])
+        apply (fastsimp intro: eval.Inst)
+        done
     next
       case (Lit v)
       show ?case
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.Lit [THEN conseq1])
-	apply (fastsimp intro: eval.Lit)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.Lit [THEN conseq1])
+        apply (fastsimp intro: eval.Lit)
+        done
     next
       case (UnOp unop e)
       thus ?case
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.UnOp)
-	apply (erule MGFnD' [THEN conseq12])
-	apply (fastsimp intro: eval.UnOp)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.UnOp)
+        apply (erule MGFnD' [THEN conseq12])
+        apply (fastsimp intro: eval.UnOp)
+        done
     next
       case (BinOp binop e1 e2)
       thus ?case
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.BinOp)
-	apply  (erule MGFnD [THEN ax_NormalD])
-	apply (rule allI)
-	apply (case_tac "need_second_arg binop v1")
-	apply  simp
-	apply  (erule MGFnD' [THEN conseq12])
-	apply  (fastsimp intro: eval.BinOp)
-	apply simp
-	apply (rule ax_Normal_cases)
-	apply  (rule ax_derivs.Skip [THEN conseq1])
-	apply  clarsimp
-	apply  (rule eval_BinOp_arg2_indepI)
-	apply   simp
-	apply  simp
-	apply (rule ax_derivs.Abrupt [THEN conseq1], clarsimp simp add: Let_def)
-	apply (fastsimp intro: eval.BinOp)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.BinOp)
+        apply  (erule MGFnD [THEN ax_NormalD])
+        apply (rule allI)
+        apply (case_tac "need_second_arg binop v1")
+        apply  simp
+        apply  (erule MGFnD' [THEN conseq12])
+        apply  (fastsimp intro: eval.BinOp)
+        apply simp
+        apply (rule ax_Normal_cases)
+        apply  (rule ax_derivs.Skip [THEN conseq1])
+        apply  clarsimp
+        apply  (rule eval_BinOp_arg2_indepI)
+        apply   simp
+        apply  simp
+        apply (rule ax_derivs.Abrupt [THEN conseq1], clarsimp simp add: Let_def)
+        apply (fastsimp intro: eval.BinOp)
+        done
     next
       case Super
       show ?case
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.Super [THEN conseq1])
-	apply (fastsimp intro: eval.Super)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.Super [THEN conseq1])
+        apply (fastsimp intro: eval.Super)
+        done
     next
       case (Acc v)
       thus ?case
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Acc])
-	apply (fastsimp intro: eval.Acc simp add: split_paired_all)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Acc])
+        apply (fastsimp intro: eval.Acc simp add: split_paired_all)
+        done
     next
       case (Ass v e)
       thus "G,A\<turnstile>{=:n} \<langle>v:=e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.Ass)
-	apply  (erule MGFnD [THEN ax_NormalD])
-	apply (rule allI)
-	apply (erule MGFnD'[THEN conseq12])
-	apply (fastsimp intro: eval.Ass simp add: split_paired_all)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.Ass)
+        apply  (erule MGFnD [THEN ax_NormalD])
+        apply (rule allI)
+        apply (erule MGFnD'[THEN conseq12])
+        apply (fastsimp intro: eval.Ass simp add: split_paired_all)
+        done
     next
       case (Cond e1 e2 e3)
       thus "G,A\<turnstile>{=:n} \<langle>e1 ? e2 : e3\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.Cond)
-	apply  (erule MGFnD [THEN ax_NormalD])
-	apply (rule allI)
-	apply (rule ax_Normal_cases)
-	prefer 2
-	apply  (rule ax_derivs.Abrupt [THEN conseq1],clarsimp simp add: Let_def)
-	apply  (fastsimp intro: eval.Cond)
-	apply (case_tac "b")
-	apply  simp
-	apply  (erule MGFnD'[THEN conseq12])
-	apply  (fastsimp intro: eval.Cond)
-	apply simp
-	apply (erule MGFnD'[THEN conseq12])
-	apply (fastsimp intro: eval.Cond)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.Cond)
+        apply  (erule MGFnD [THEN ax_NormalD])
+        apply (rule allI)
+        apply (rule ax_Normal_cases)
+        prefer 2
+        apply  (rule ax_derivs.Abrupt [THEN conseq1],clarsimp simp add: Let_def)
+        apply  (fastsimp intro: eval.Cond)
+        apply (case_tac "b")
+        apply  simp
+        apply  (erule MGFnD'[THEN conseq12])
+        apply  (fastsimp intro: eval.Cond)
+        apply simp
+        apply (erule MGFnD'[THEN conseq12])
+        apply (fastsimp intro: eval.Cond)
+        done
     next
       case (Call accC statT mode e mn pTs' ps)
       note mgf_e = `G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}`
       note mgf_ps = `G,A\<turnstile>{=:n} \<langle>ps\<rangle>\<^sub>l\<succ> {G\<rightarrow>}`
       from mgf_methds mgf_e mgf_ps wf
       show "G,A\<turnstile>{=:n} \<langle>{accC,statT,mode}e\<cdot>mn({pTs'}ps)\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
-	by (rule MGFn_Call)
+        by (rule MGFn_Call)
     next
       case (Methd D mn)
       from mgf_methds
       show "G,A\<turnstile>{=:n} \<langle>Methd D mn\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
-	by simp
+        by simp
     next
       case (Body D c)
       note mgf_c = `G,A\<turnstile>{=:n} \<langle>c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}`
       from wf MGFn_Init [OF hyp] mgf_c
       show "G,A\<turnstile>{=:n} \<langle>Body D c\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
-	by (rule MGFn_Body)
+        by (rule MGFn_Body)
     next
       case (InsInitE c e)
       show ?case
-	by (rule MGFn_NormalI) (rule ax_derivs.InsInitE)
+        by (rule MGFn_NormalI) (rule ax_derivs.InsInitE)
     next
       case (Callee l e)
       show ?case
-	by (rule MGFn_NormalI) (rule ax_derivs.Callee)
+        by (rule MGFn_NormalI) (rule ax_derivs.Callee)
     next
       case Skip
       show ?case
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.Skip [THEN conseq1])
-	apply (fastsimp intro: eval.Skip)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.Skip [THEN conseq1])
+        apply (fastsimp intro: eval.Skip)
+        done
     next
       case (Expr e)
       thus ?case
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Expr])
-	apply (fastsimp intro: eval.Expr)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Expr])
+        apply (fastsimp intro: eval.Expr)
+        done
     next
       case (Lab l c)
       thus "G,A\<turnstile>{=:n} \<langle>l\<bullet> c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Lab])
-	apply (fastsimp intro: eval.Lab)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Lab])
+        apply (fastsimp intro: eval.Lab)
+        done
     next
       case (Comp c1 c2)
       thus "G,A\<turnstile>{=:n} \<langle>c1;; c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.Comp)
-	apply  (erule MGFnD [THEN ax_NormalD])
-	apply (erule MGFnD' [THEN conseq12])
-	apply (fastsimp intro: eval.Comp) 
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.Comp)
+        apply  (erule MGFnD [THEN ax_NormalD])
+        apply (erule MGFnD' [THEN conseq12])
+        apply (fastsimp intro: eval.Comp) 
+        done
     next
       case (If' e c1 c2)
       thus "G,A\<turnstile>{=:n} \<langle>If(e) c1 Else c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.If)
-	apply  (erule MGFnD [THEN ax_NormalD])
-	apply (rule allI)
-	apply (rule ax_Normal_cases)
-	prefer 2
-	apply  (rule ax_derivs.Abrupt [THEN conseq1],clarsimp simp add: Let_def)
-	apply  (fastsimp intro: eval.If)
-	apply (case_tac "b")
-	apply  simp
-	apply  (erule MGFnD' [THEN conseq12])
-	apply  (fastsimp intro: eval.If)
-	apply simp
-	apply (erule MGFnD' [THEN conseq12])
-	apply (fastsimp intro: eval.If)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.If)
+        apply  (erule MGFnD [THEN ax_NormalD])
+        apply (rule allI)
+        apply (rule ax_Normal_cases)
+        prefer 2
+        apply  (rule ax_derivs.Abrupt [THEN conseq1],clarsimp simp add: Let_def)
+        apply  (fastsimp intro: eval.If)
+        apply (case_tac "b")
+        apply  simp
+        apply  (erule MGFnD' [THEN conseq12])
+        apply  (fastsimp intro: eval.If)
+        apply simp
+        apply (erule MGFnD' [THEN conseq12])
+        apply (fastsimp intro: eval.If)
+        done
     next
       case (Loop l e c)
       note mgf_e = `G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}`
       note mgf_c = `G,A\<turnstile>{=:n} \<langle>c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}`
       from mgf_e mgf_c wf
       show "G,A\<turnstile>{=:n} \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
-	by (rule MGFn_Loop)
+        by (rule MGFn_Loop)
     next
       case (Jmp j)
       thus ?case
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.Jmp [THEN conseq1])
-	apply (auto intro: eval.Jmp simp add: abupd_def2)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.Jmp [THEN conseq1])
+        apply (auto intro: eval.Jmp simp add: abupd_def2)
+        done
     next
       case (Throw e)
       thus ?case
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Throw])
-	apply (fastsimp intro: eval.Throw)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Throw])
+        apply (fastsimp intro: eval.Throw)
+        done
     next
       case (TryC c1 C vn c2)
       thus "G,A\<turnstile>{=:n} \<langle>Try c1 Catch(C vn) c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.Try [where 
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.Try [where 
           ?Q = " (\<lambda>Y' s' s. normal s \<and> (\<exists>s''. G\<turnstile>s \<midarrow>\<langle>c1\<rangle>\<^sub>s\<succ>\<rightarrow> (Y',s'') \<and> 
                             G\<turnstile>s'' \<midarrow>sxalloc\<rightarrow> s')) \<and>. G\<turnstile>init\<le>n"])
-	apply   (erule MGFnD [THEN ax_NormalD, THEN conseq2])
-	apply   (fastsimp elim: sxalloc_gext [THEN card_nyinitcls_gext])
-	apply  (erule MGFnD'[THEN conseq12])
-	apply  (fastsimp intro: eval.Try)
-	apply (fastsimp intro: eval.Try)
-	done
+        apply   (erule MGFnD [THEN ax_NormalD, THEN conseq2])
+        apply   (fastsimp elim: sxalloc_gext [THEN card_nyinitcls_gext])
+        apply  (erule MGFnD'[THEN conseq12])
+        apply  (fastsimp intro: eval.Try)
+        apply (fastsimp intro: eval.Try)
+        done
     next
       case (Fin c1 c2)
       note mgf_c1 = `G,A\<turnstile>{=:n} \<langle>c1\<rangle>\<^sub>s\<succ> {G\<rightarrow>}`
       note mgf_c2 = `G,A\<turnstile>{=:n} \<langle>c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}`
       from wf mgf_c1 mgf_c2
       show "G,A\<turnstile>{=:n} \<langle>c1 Finally c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
-	by (rule MGFn_Fin)
+        by (rule MGFn_Fin)
     next
       case (FinA abr c)
       show ?case
-	by (rule MGFn_NormalI) (rule ax_derivs.FinA)
+        by (rule MGFn_NormalI) (rule ax_derivs.FinA)
     next
       case (Init C)
       from hyp
       show "G,A\<turnstile>{=:n} \<langle>Init C\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
-	by (rule MGFn_Init)
+        by (rule MGFn_Init)
     next
       case Nil_expr
       show "G,A\<turnstile>{=:n} \<langle>[]\<rangle>\<^sub>l\<succ> {G\<rightarrow>}"
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.Nil [THEN conseq1])
-	apply (fastsimp intro: eval.Nil)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.Nil [THEN conseq1])
+        apply (fastsimp intro: eval.Nil)
+        done
     next
       case (Cons_expr e es)
       thus "G,A\<turnstile>{=:n} \<langle>e# es\<rangle>\<^sub>l\<succ> {G\<rightarrow>}"
-	apply -
-	apply (rule MGFn_NormalI)
-	apply (rule ax_derivs.Cons)
-	apply  (erule MGFnD [THEN ax_NormalD])
-	apply (rule allI)
-	apply (erule MGFnD'[THEN conseq12])
-	apply (fastsimp intro: eval.Cons)
-	done
+        apply -
+        apply (rule MGFn_NormalI)
+        apply (rule ax_derivs.Cons)
+        apply  (erule MGFnD [THEN ax_NormalD])
+        apply (rule allI)
+        apply (erule MGFnD'[THEN conseq12])
+        apply (fastsimp intro: eval.Cons)
+        done
     qed
   }
   thus ?thesis
@@ -1519,15 +1519,15 @@
     proof -
       from eval_t type_ok wf 
       obtain n where evaln: "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y', s')"
-	by (rule eval_to_evaln [elim_format]) iprover
+        by (rule eval_to_evaln [elim_format]) iprover
       from valid have 
-	valid_expanded:
-	"\<forall>n Y s Z. P Y s Z \<longrightarrow> type_ok G t s 
+        valid_expanded:
+        "\<forall>n Y s Z. P Y s Z \<longrightarrow> type_ok G t s 
                    \<longrightarrow> (\<forall>Y' s'. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y', s') \<longrightarrow> Q Y' s' Z)"
-	by (simp add: ax_valids_def triple_valid_def)
+        by (simp add: ax_valids_def triple_valid_def)
       from P type_ok evaln
       show "Q Y' s' Z"
-	by (rule valid_expanded [rule_format])