merged
authorhaftmann
Wed, 21 Oct 2009 08:16:25 +0200
changeset 33039 5018f6a76b3f
parent 33036 c61fe520602b (diff)
parent 33038 8f9594c31de4 (current diff)
child 33040 cffdb7b28498
child 33041 6793b02a3409
merged
src/HOL/Import/proof_kernel.ML
src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML
src/HOL/Library/normarith.ML
src/HOL/Library/positivstellensatz.ML
src/HOL/Set.thy
src/HOL/Tools/Groebner_Basis/groebner.ML
src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML
src/HOL/Tools/Qelim/cooper.ML
src/HOL/Tools/TFL/tfl.ML
src/HOL/Tools/metis_tools.ML
src/HOL/Tools/prop_logic.ML
src/HOL/Tools/record.ML
src/HOL/Tools/refute.ML
src/HOL/Tools/res_atp.ML
src/HOL/Tools/res_axioms.ML
src/HOL/Tools/res_hol_clause.ML
src/HOL/Tools/res_reconstruct.ML
src/HOL/Tools/sat_solver.ML
src/HOL/ex/predicate_compile.ML
src/Pure/Tools/find_theorems.ML
src/Pure/meta_simplifier.ML
--- a/Admin/isatest/isatest-stats	Wed Oct 21 08:14:38 2009 +0200
+++ b/Admin/isatest/isatest-stats	Wed Oct 21 08:16:25 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/doc-src/Main/Docs/Main_Doc.thy	Wed Oct 21 08:14:38 2009 +0200
+++ b/doc-src/Main/Docs/Main_Doc.thy	Wed Oct 21 08:16:25 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 08:14:38 2009 +0200
+++ b/doc-src/Main/Docs/document/Main_Doc.tex	Wed Oct 21 08:16:25 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/src/HOL/Decision_Procs/Approximation.thy	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Decision_Procs/Approximation.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -3462,9 +3462,9 @@
 
   fun approx_form prec ctxt t =
           realify t
-       |> prepare_form (Context.theory_of_proof ctxt)
+       |> prepare_form (ProofContext.theory_of ctxt)
        |> (fn arith_term =>
-          reify_form (Context.theory_of_proof ctxt) arith_term
+          reify_form (ProofContext.theory_of ctxt) arith_term
        |> HOLogic.dest_Trueprop |> dest_interpret_form
        |> (fn (data, xs) =>
           mk_approx_form_eval prec data (HOLogic.mk_list @{typ "(float * float) option"}
--- a/src/HOL/HOL.thy	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/HOL.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -839,9 +839,6 @@
 ML_Antiquote.value "claset"
   (Scan.succeed "Classical.claset_of (ML_Context.the_local_context ())");
 
-structure ResAtpset = Named_Thms
-  (val name = "atp" val description = "ATP rules");
-
 structure ResBlacklist = Named_Thms
   (val name = "noatp" val description = "theorems blacklisted for ATP");
 *}
@@ -860,7 +857,6 @@
   Hypsubst.hypsubst_setup
   #> ContextRules.addSWrapper (fn tac => hyp_subst_tac' ORELSE' tac)
   #> Classical.setup
-  #> ResAtpset.setup
   #> ResBlacklist.setup
 end
 *}
--- a/src/HOL/Import/proof_kernel.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Import/proof_kernel.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -584,7 +584,7 @@
 
 fun input_types thyname (Elem("tylist",[("i",i)],xtys)) =
     let
-        val types = Array.array(valOf (Int.fromString i),XMLty (Elem("",[],[])))
+        val types = Array.array(the (Int.fromString i),XMLty (Elem("",[],[])))
         fun IT _ [] = ()
           | IT n (xty::xtys) =
             (Array.update(types,n,XMLty xty);
@@ -651,7 +651,7 @@
 
 fun input_terms thyname types (Elem("tmlist",[("i",i)],xtms)) =
     let
-        val terms = Array.array(valOf(Int.fromString i),XMLtm (Elem("",[],[])))
+        val terms = Array.array(the (Int.fromString i), XMLtm (Elem("",[],[])))
 
         fun IT _ [] = ()
           | IT n (xtm::xtms) =
@@ -1240,7 +1240,7 @@
         val (newstr,u) = pairself Substring.string (Substring.splitr (fn c => c = #"_") f)
     in
         if not (idx = "") andalso u = "_"
-        then SOME (newstr,valOf(Int.fromString idx))
+        then SOME (newstr, the (Int.fromString idx))
         else NONE
     end
     handle _ => NONE  (* FIXME avoid handle _ *)
@@ -1915,7 +1915,7 @@
 fun new_definition thyname constname rhs thy =
     let
         val constname = rename_const thyname thy constname
-        val redeclared = isSome (Sign.const_type thy (Sign.intern_const thy constname));
+        val redeclared = is_some (Sign.const_type thy (Sign.intern_const thy constname));
         val _ = warning ("Introducing constant " ^ constname)
         val (thmname,thy) = get_defname thyname constname thy
         val (info,rhs') = disamb_term rhs
--- a/src/HOL/Import/replay.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Import/replay.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -291,7 +291,7 @@
                 fun get_facts facts =
                     case TextIO.inputLine is of
                         NONE => (case facts of
-                                   i::facts => (valOf (Int.fromString i),map P.protect_factname (rev facts))
+                                   i::facts => (the (Int.fromString i),map P.protect_factname (rev facts))
                                  | _ => raise ERR "replay_thm" "Bad facts.lst file")
                       | SOME fact => get_facts ((String.substring(fact,0,String.size fact -1 ))::facts)
             in
--- a/src/HOL/IsaMakefile	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/IsaMakefile	Wed Oct 21 08:16:25 2009 +0200
@@ -25,11 +25,11 @@
   HOL-IOA \
   HOL-Imperative_HOL \
   HOL-Induct \
-  HOL-Isar_examples \
+  HOL-Isar_Examples \
   HOL-Lambda \
   HOL-Lattice \
   HOL-Matrix \
-  HOL-MetisExamples \
+  HOL-Metis_Examples \
   HOL-MicroJava \
   HOL-Mirabelle \
   HOL-Modelcheck \
@@ -38,7 +38,7 @@
   HOL-Number_Theory \
   HOL-Old_Number_Theory \
   HOL-Prolog \
-  HOL-SET-Protocol \
+  HOL-SET_Protocol \
   HOL-SizeChange \
   HOL-SMT-Examples \
   HOL-Statespace \
@@ -556,16 +556,16 @@
 	@$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL Hoare_Parallel
 
 
-## HOL-MetisExamples
+## HOL-Metis_Examples
 
-HOL-MetisExamples: HOL $(LOG)/HOL-MetisExamples.gz
+HOL-Metis_Examples: HOL $(LOG)/HOL-Metis_Examples.gz
 
-$(LOG)/HOL-MetisExamples.gz: $(OUT)/HOL MetisExamples/ROOT.ML	\
-  MetisExamples/Abstraction.thy MetisExamples/BigO.thy		\
-  MetisExamples/BT.thy MetisExamples/Message.thy		\
-  MetisExamples/Tarski.thy MetisExamples/TransClosure.thy	\
-  MetisExamples/set.thy
-	@$(ISABELLE_TOOL) usedir $(OUT)/HOL MetisExamples
+$(LOG)/HOL-Metis_Examples.gz: $(OUT)/HOL Metis_Examples/ROOT.ML	\
+  Metis_Examples/Abstraction.thy Metis_Examples/BigO.thy	\
+  Metis_Examples/BT.thy Metis_Examples/Message.thy		\
+  Metis_Examples/Tarski.thy Metis_Examples/TransClosure.thy	\
+  Metis_Examples/set.thy
+	@$(ISABELLE_TOOL) usedir $(OUT)/HOL Metis_Examples
 
 
 ## HOL-Algebra
@@ -837,14 +837,15 @@
 
 HOL-Bali: HOL $(LOG)/HOL-Bali.gz
 
-$(LOG)/HOL-Bali.gz: $(OUT)/HOL Bali/Bali.thy Bali/AxCompl.thy Bali/AxExample.thy	\
-  Bali/AxSem.thy Bali/AxSound.thy Bali/Basis.thy Bali/Conform.thy	\
-  Bali/Decl.thy Bali/DeclConcepts.thy Bali/Eval.thy Bali/Evaln.thy	\
-  Bali/Example.thy Bali/Name.thy Bali/ROOT.ML Bali/State.thy		\
-  Bali/Table.thy Bali/Term.thy Bali/Trans.thy Bali/Type.thy		\
-  Bali/TypeRel.thy Bali/TypeSafe.thy Bali/Value.thy Bali/WellForm.thy	\
-  Bali/DefiniteAssignment.thy Bali/DefiniteAssignmentCorrect.thy	\
-  Bali/WellType.thy Bali/document/root.tex
+$(LOG)/HOL-Bali.gz: $(OUT)/HOL Bali/Bali.thy Bali/AxCompl.thy		\
+  Bali/AxExample.thy Bali/AxSem.thy Bali/AxSound.thy Bali/Basis.thy	\
+  Bali/Conform.thy Bali/Decl.thy Bali/DeclConcepts.thy Bali/Eval.thy	\
+  Bali/Evaln.thy Bali/Example.thy Bali/Name.thy Bali/ROOT.ML		\
+  Bali/State.thy Bali/Table.thy Bali/Term.thy Bali/Trans.thy		\
+  Bali/Type.thy Bali/TypeRel.thy Bali/TypeSafe.thy Bali/Value.thy	\
+  Bali/WellForm.thy Bali/DefiniteAssignment.thy				\
+  Bali/DefiniteAssignmentCorrect.thy Bali/WellType.thy			\
+  Bali/document/root.tex
 	@$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL Bali
 
 
@@ -885,63 +886,62 @@
 HOL-ex: HOL $(LOG)/HOL-ex.gz
 
 $(LOG)/HOL-ex.gz: $(OUT)/HOL Library/Commutative_Ring.thy		\
-  Number_Theory/Primes.thy ex/Abstract_NAT.thy ex/Antiquote.thy		\
-  ex/Arith_Examples.thy				\
+  Number_Theory/Primes.thy						\
+  Tools/Predicate_Compile/predicate_compile_core.ML			\
+  ex/Abstract_NAT.thy ex/Antiquote.thy ex/Arith_Examples.thy		\
   ex/Arithmetic_Series_Complex.thy ex/BT.thy ex/BinEx.thy		\
   ex/Binary.thy ex/CTL.thy ex/Chinese.thy ex/Classical.thy		\
-  ex/CodegenSML_Test.thy ex/Codegenerator_Candidates.thy				\
-  ex/Codegenerator_Pretty.thy ex/Codegenerator_Test.thy \
-  ex/Codegenerator_Pretty_Test.thy ex/Coherent.thy				\
-  ex/Commutative_RingEx.thy ex/Commutative_Ring_Complete.thy		\
-  ex/Efficient_Nat_examples.thy		\
-  ex/Eval_Examples.thy	ex/Fundefs.thy			\
-  ex/Groebner_Examples.thy ex/Guess.thy ex/HarmonicSeries.thy		\
-  ex/Hebrew.thy ex/Hex_Bin_Examples.thy ex/Higher_Order_Logic.thy	\
-  ex/Hilbert_Classical.thy			\
-  ex/Induction_Scheme.thy ex/InductiveInvariant.thy			\
-  ex/InductiveInvariant_examples.thy ex/Intuitionistic.thy		\
-  ex/Lagrange.thy ex/LocaleTest2.thy ex/MT.thy ex/MergeSort.thy		\
-  ex/Meson_Test.thy ex/MonoidGroup.thy ex/Multiquote.thy ex/NatSum.thy	\
-  ex/Numeral.thy ex/PER.thy ex/PresburgerEx.thy ex/Primrec.thy		\
-  ex/Quickcheck_Examples.thy ex/ROOT.ML	\
-  ex/Recdefs.thy ex/Records.thy ex/ReflectionEx.thy			\
+  ex/CodegenSML_Test.thy ex/Codegenerator_Candidates.thy		\
+  ex/Codegenerator_Pretty.thy ex/Codegenerator_Pretty_Test.thy		\
+  ex/Codegenerator_Test.thy ex/Coherent.thy ex/Commutative_RingEx.thy	\
+  ex/Commutative_Ring_Complete.thy ex/Efficient_Nat_examples.thy	\
+  ex/Eval_Examples.thy ex/Fundefs.thy ex/Groebner_Examples.thy		\
+  ex/Guess.thy ex/HarmonicSeries.thy ex/Hebrew.thy			\
+  ex/Hex_Bin_Examples.thy ex/Higher_Order_Logic.thy			\
+  ex/Hilbert_Classical.thy ex/Induction_Scheme.thy			\
+  ex/InductiveInvariant.thy ex/InductiveInvariant_examples.thy		\
+  ex/Intuitionistic.thy ex/Lagrange.thy ex/LocaleTest2.thy ex/MT.thy	\
+  ex/MergeSort.thy ex/Meson_Test.thy ex/MonoidGroup.thy			\
+  ex/Multiquote.thy ex/NatSum.thy ex/Numeral.thy ex/PER.thy		\
+  ex/Predicate_Compile.thy ex/Predicate_Compile_ex.thy			\
+  ex/PresburgerEx.thy ex/Primrec.thy ex/Quickcheck_Examples.thy		\
+  ex/ROOT.ML ex/Recdefs.thy ex/Records.thy ex/ReflectionEx.thy		\
   ex/Refute_Examples.thy ex/SAT_Examples.thy ex/SVC_Oracle.thy		\
-  ex/Serbian.thy ex/Sqrt.thy ex/Sqrt_Script.thy \
-  ex/Sudoku.thy ex/Tarski.thy \
-  ex/Termination.thy ex/Transfer_Ex.thy ex/Unification.thy ex/document/root.bib		\
-  ex/document/root.tex ex/set.thy ex/svc_funcs.ML ex/svc_test.thy \
-  ex/Predicate_Compile.thy Tools/Predicate_Compile/predicate_compile_core.ML ex/Predicate_Compile_ex.thy
+  ex/Serbian.thy ex/Sqrt.thy ex/Sqrt_Script.thy ex/Sudoku.thy		\
+  ex/Tarski.thy ex/Termination.thy ex/Transfer_Ex.thy			\
+  ex/Unification.thy ex/document/root.bib ex/document/root.tex		\
+  ex/set.thy ex/svc_funcs.ML ex/svc_test.thy
 	@$(ISABELLE_TOOL) usedir $(OUT)/HOL ex
 
 
-## HOL-Isar_examples
+## HOL-Isar_Examples
 
-HOL-Isar_examples: HOL $(LOG)/HOL-Isar_examples.gz
+HOL-Isar_Examples: HOL $(LOG)/HOL-Isar_Examples.gz
 
-$(LOG)/HOL-Isar_examples.gz: $(OUT)/HOL Isar_examples/Basic_Logic.thy	\
-  Isar_examples/Cantor.thy Isar_examples/Drinker.thy			\
-  Isar_examples/Expr_Compiler.thy Isar_examples/Fibonacci.thy		\
-  Isar_examples/Group.thy Isar_examples/Hoare.thy			\
-  Isar_examples/Hoare_Ex.thy Isar_examples/Knaster_Tarski.thy		\
-  Isar_examples/Mutilated_Checkerboard.thy				\
-  Isar_examples/Nested_Datatype.thy Isar_examples/Peirce.thy		\
-  Isar_examples/Puzzle.thy Isar_examples/Summation.thy			\
-  Isar_examples/ROOT.ML Isar_examples/document/proof.sty		\
-  Isar_examples/document/root.bib Isar_examples/document/root.tex	\
-  Isar_examples/document/style.tex Hoare/hoare_tac.ML
-	@$(ISABELLE_TOOL) usedir $(OUT)/HOL Isar_examples
+$(LOG)/HOL-Isar_Examples.gz: $(OUT)/HOL Isar_Examples/Basic_Logic.thy	\
+  Isar_Examples/Cantor.thy Isar_Examples/Drinker.thy			\
+  Isar_Examples/Expr_Compiler.thy Isar_Examples/Fibonacci.thy		\
+  Isar_Examples/Group.thy Isar_Examples/Hoare.thy			\
+  Isar_Examples/Hoare_Ex.thy Isar_Examples/Knaster_Tarski.thy		\
+  Isar_Examples/Mutilated_Checkerboard.thy				\
+  Isar_Examples/Nested_Datatype.thy Isar_Examples/Peirce.thy		\
+  Isar_Examples/Puzzle.thy Isar_Examples/Summation.thy			\
+  Isar_Examples/ROOT.ML Isar_Examples/document/proof.sty		\
+  Isar_Examples/document/root.bib Isar_Examples/document/root.tex	\
+  Isar_Examples/document/style.tex Hoare/hoare_tac.ML
+	@$(ISABELLE_TOOL) usedir $(OUT)/HOL Isar_Examples
 
 
-## HOL-SET-Protocol
+## HOL-SET_Protocol
 
-HOL-SET-Protocol: HOL $(LOG)/HOL-SET-Protocol.gz
+HOL-SET_Protocol: HOL $(LOG)/HOL-SET_Protocol.gz
 
-$(LOG)/HOL-SET-Protocol.gz: $(OUT)/HOL SET-Protocol/ROOT.ML		\
-  SET-Protocol/MessageSET.thy SET-Protocol/EventSET.thy			\
-  SET-Protocol/PublicSET.thy SET-Protocol/Cardholder_Registration.thy	\
-  SET-Protocol/Merchant_Registration.thy SET-Protocol/Purchase.thy	\
-  SET-Protocol/document/root.tex
-	@$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL SET-Protocol
+$(LOG)/HOL-SET_Protocol.gz: $(OUT)/HOL SET_Protocol/ROOT.ML		\
+  SET_Protocol/Message_SET.thy SET_Protocol/Event_SET.thy		\
+  SET_Protocol/Public_SET.thy SET_Protocol/Cardholder_Registration.thy	\
+  SET_Protocol/Merchant_Registration.thy SET_Protocol/Purchase.thy	\
+  SET_Protocol/document/root.tex
+	@$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL SET_Protocol
 
 
 ## HOL-Matrix
@@ -987,7 +987,8 @@
 
 TLA-Buffer: TLA $(LOG)/TLA-Buffer.gz
 
-$(LOG)/TLA-Buffer.gz: $(OUT)/TLA TLA/Buffer/Buffer.thy TLA/Buffer/DBuffer.thy
+$(LOG)/TLA-Buffer.gz: $(OUT)/TLA TLA/Buffer/Buffer.thy	\
+  TLA/Buffer/DBuffer.thy
 	@cd TLA; $(ISABELLE_TOOL) usedir $(OUT)/TLA Buffer
 
 
@@ -1056,8 +1057,7 @@
 
 HOL-Word: HOL $(OUT)/HOL-Word
 
-$(OUT)/HOL-Word: $(OUT)/HOL Word/ROOT.ML	\
-  Library/Boolean_Algebra.thy			\
+$(OUT)/HOL-Word: $(OUT)/HOL Word/ROOT.ML Library/Boolean_Algebra.thy	\
   Library/Numeral_Type.thy Word/Num_Lemmas.thy Word/TdThs.thy		\
   Word/Size.thy Word/BinGeneral.thy Word/BinOperations.thy		\
   Word/BinBoolList.thy Word/BitSyntax.thy Word/WordDefinition.thy	\
@@ -1092,30 +1092,13 @@
 
 HOL-NSA: HOL $(OUT)/HOL-NSA
 
-$(OUT)/HOL-NSA: $(OUT)/HOL \
-  NSA/CLim.thy \
-  NSA/CStar.thy \
-  NSA/NSCA.thy \
-  NSA/NSComplex.thy \
-  NSA/HDeriv.thy \
-  NSA/HLim.thy \
-  NSA/HLog.thy \
-  NSA/HSEQ.thy \
-  NSA/HSeries.thy \
-  NSA/HTranscendental.thy \
-  NSA/Hypercomplex.thy \
-  NSA/HyperDef.thy \
-  NSA/HyperNat.thy \
-  NSA/Hyperreal.thy \
-  NSA/Filter.thy \
-  NSA/NatStar.thy \
-  NSA/NSA.thy \
-  NSA/StarDef.thy \
-  NSA/Star.thy \
-  NSA/transfer.ML \
-  Library/Infinite_Set.thy \
-  Library/Zorn.thy \
-  NSA/ROOT.ML
+$(OUT)/HOL-NSA: $(OUT)/HOL NSA/CLim.thy NSA/CStar.thy NSA/NSCA.thy	\
+  NSA/NSComplex.thy NSA/HDeriv.thy NSA/HLim.thy NSA/HLog.thy		\
+  NSA/HSEQ.thy NSA/HSeries.thy NSA/HTranscendental.thy			\
+  NSA/Hypercomplex.thy NSA/HyperDef.thy NSA/HyperNat.thy		\
+  NSA/Hyperreal.thy NSA/Filter.thy NSA/NatStar.thy NSA/NSA.thy		\
+  NSA/StarDef.thy NSA/Star.thy NSA/transfer.ML				\
+  Library/Infinite_Set.thy Library/Zorn.thy NSA/ROOT.ML
 	@cd NSA; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL HOL-NSA
 
 
@@ -1132,28 +1115,28 @@
 
 HOL-Mirabelle: HOL $(LOG)/HOL-Mirabelle.gz
 
-$(LOG)/HOL-Mirabelle.gz: $(OUT)/HOL Mirabelle/Mirabelle_Test.thy \
-  Mirabelle/Mirabelle.thy Mirabelle/Tools/mirabelle.ML Mirabelle/ROOT.ML \
-  Mirabelle/Tools/mirabelle_arith.ML \
-  Mirabelle/Tools/mirabelle_metis.ML \
-  Mirabelle/Tools/mirabelle_quickcheck.ML \
-  Mirabelle/Tools/mirabelle_refute.ML \
+$(LOG)/HOL-Mirabelle.gz: $(OUT)/HOL Mirabelle/Mirabelle_Test.thy	\
+  Mirabelle/Mirabelle.thy Mirabelle/Tools/mirabelle.ML			\
+  Mirabelle/ROOT.ML Mirabelle/Tools/mirabelle_arith.ML			\
+  Mirabelle/Tools/mirabelle_metis.ML					\
+  Mirabelle/Tools/mirabelle_quickcheck.ML				\
+  Mirabelle/Tools/mirabelle_refute.ML					\
   Mirabelle/Tools/mirabelle_sledgehammer.ML
 	@$(ISABELLE_TOOL) usedir $(OUT)/HOL Mirabelle
 
 
 ## HOL-SMT
 
-HOL-SMT: HOL-Word $(LOG)/HOL-SMT.gz
+HOL-SMT: HOL-Word $(OUT)/HOL-SMT
 
-$(LOG)/HOL-SMT.gz: $(OUT)/HOL-Word SMT/SMT_Base.thy SMT/Z3.thy SMT/SMT.thy \
-  SMT/Tools/smt_normalize.ML SMT/Tools/smt_monomorph.ML \
-  SMT/Tools/smt_translate.ML SMT/Tools/smt_builtin.ML \
-  SMT/Tools/smtlib_interface.ML SMT/Tools/smt_solver.ML \
-  SMT/Tools/cvc3_solver.ML SMT/Tools/yices_solver.ML \
-  SMT/Tools/z3_proof_terms.ML SMT/Tools/z3_proof_rules.ML \
-  SMT/Tools/z3_proof.ML SMT/Tools/z3_model.ML SMT/Tools/z3_interface.ML \
-  SMT/Tools/z3_solver.ML
+$(OUT)/HOL-SMT: $(OUT)/HOL-Word SMT/SMT_Base.thy SMT/Z3.thy		\
+  SMT/SMT.thy SMT/Tools/smt_normalize.ML SMT/Tools/smt_monomorph.ML	\
+  SMT/Tools/smt_translate.ML SMT/Tools/smt_builtin.ML			\
+  SMT/Tools/smtlib_interface.ML SMT/Tools/smt_solver.ML			\
+  SMT/Tools/cvc3_solver.ML SMT/Tools/yices_solver.ML			\
+  SMT/Tools/z3_proof_terms.ML SMT/Tools/z3_proof_rules.ML		\
+  SMT/Tools/z3_proof.ML SMT/Tools/z3_model.ML				\
+  SMT/Tools/z3_interface.ML SMT/Tools/z3_solver.ML
 	@cd SMT; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL-Word HOL-SMT
 
 
@@ -1161,203 +1144,157 @@
 
 HOL-SMT-Examples: HOL-SMT $(LOG)/HOL-SMT-Examples.gz
 
-$(LOG)/HOL-SMT-Examples.gz: $(OUT)/HOL-SMT SMT/Examples/ROOT.ML	\
-  SMT/Examples/SMT_Examples.thy \
-  SMT/Examples/cert/z3_arith_quant_01 \
-  SMT/Examples/cert/z3_arith_quant_01.proof \
-  SMT/Examples/cert/z3_arith_quant_02 \
-  SMT/Examples/cert/z3_arith_quant_02.proof \
-  SMT/Examples/cert/z3_arith_quant_03 \
-  SMT/Examples/cert/z3_arith_quant_03.proof \
-  SMT/Examples/cert/z3_arith_quant_04 \
-  SMT/Examples/cert/z3_arith_quant_04.proof \
-  SMT/Examples/cert/z3_arith_quant_05 \
-  SMT/Examples/cert/z3_arith_quant_05.proof \
-  SMT/Examples/cert/z3_arith_quant_06 \
-  SMT/Examples/cert/z3_arith_quant_06.proof \
-  SMT/Examples/cert/z3_arith_quant_07 \
-  SMT/Examples/cert/z3_arith_quant_07.proof \
-  SMT/Examples/cert/z3_arith_quant_08 \
-  SMT/Examples/cert/z3_arith_quant_08.proof \
-  SMT/Examples/cert/z3_arith_quant_09 \
-  SMT/Examples/cert/z3_arith_quant_09.proof \
-  SMT/Examples/cert/z3_arith_quant_10 \
-  SMT/Examples/cert/z3_arith_quant_10.proof \
-  SMT/Examples/cert/z3_arith_quant_11 \
-  SMT/Examples/cert/z3_arith_quant_11.proof \
-  SMT/Examples/cert/z3_arith_quant_12 \
-  SMT/Examples/cert/z3_arith_quant_12.proof \
-  SMT/Examples/cert/z3_arith_quant_13 \
-  SMT/Examples/cert/z3_arith_quant_13.proof \
-  SMT/Examples/cert/z3_arith_quant_14 \
-  SMT/Examples/cert/z3_arith_quant_14.proof \
-  SMT/Examples/cert/z3_arith_quant_15 \
-  SMT/Examples/cert/z3_arith_quant_15.proof \
-  SMT/Examples/cert/z3_arith_quant_16 \
-  SMT/Examples/cert/z3_arith_quant_16.proof \
-  SMT/Examples/cert/z3_arith_quant_17 \
-  SMT/Examples/cert/z3_arith_quant_17.proof \
-  SMT/Examples/cert/z3_arith_quant_18 \
-  SMT/Examples/cert/z3_arith_quant_18.proof \
-  SMT/Examples/cert/z3_bv_01 \
-  SMT/Examples/cert/z3_bv_01.proof \
-  SMT/Examples/cert/z3_bv_02 \
-  SMT/Examples/cert/z3_bv_02.proof \
-  SMT/Examples/cert/z3_bv_arith_01 \
-  SMT/Examples/cert/z3_bv_arith_01.proof \
-  SMT/Examples/cert/z3_bv_arith_02 \
-  SMT/Examples/cert/z3_bv_arith_02.proof \
-  SMT/Examples/cert/z3_bv_arith_03 \
-  SMT/Examples/cert/z3_bv_arith_03.proof \
-  SMT/Examples/cert/z3_bv_arith_04 \
-  SMT/Examples/cert/z3_bv_arith_04.proof \
-  SMT/Examples/cert/z3_bv_arith_05 \
-  SMT/Examples/cert/z3_bv_arith_05.proof \
-  SMT/Examples/cert/z3_bv_arith_06 \
-  SMT/Examples/cert/z3_bv_arith_06.proof \
-  SMT/Examples/cert/z3_bv_arith_07 \
-  SMT/Examples/cert/z3_bv_arith_07.proof \
-  SMT/Examples/cert/z3_bv_arith_08 \
-  SMT/Examples/cert/z3_bv_arith_08.proof \
-  SMT/Examples/cert/z3_bv_arith_09 \
-  SMT/Examples/cert/z3_bv_arith_09.proof \
-  SMT/Examples/cert/z3_bv_arith_10 \
-  SMT/Examples/cert/z3_bv_arith_10.proof \
-  SMT/Examples/cert/z3_bv_bit_01 \
-  SMT/Examples/cert/z3_bv_bit_01.proof \
-  SMT/Examples/cert/z3_bv_bit_02 \
-  SMT/Examples/cert/z3_bv_bit_02.proof \
-  SMT/Examples/cert/z3_bv_bit_03 \
-  SMT/Examples/cert/z3_bv_bit_03.proof \
-  SMT/Examples/cert/z3_bv_bit_04 \
-  SMT/Examples/cert/z3_bv_bit_04.proof \
-  SMT/Examples/cert/z3_bv_bit_05 \
-  SMT/Examples/cert/z3_bv_bit_05.proof \
-  SMT/Examples/cert/z3_bv_bit_06 \
-  SMT/Examples/cert/z3_bv_bit_06.proof \
-  SMT/Examples/cert/z3_bv_bit_07 \
-  SMT/Examples/cert/z3_bv_bit_07.proof \
-  SMT/Examples/cert/z3_bv_bit_08 \
-  SMT/Examples/cert/z3_bv_bit_08.proof \
-  SMT/Examples/cert/z3_bv_bit_09 \
-  SMT/Examples/cert/z3_bv_bit_09.proof \
-  SMT/Examples/cert/z3_bv_bit_10 \
-  SMT/Examples/cert/z3_bv_bit_10.proof \
-  SMT/Examples/cert/z3_bv_bit_11 \
-  SMT/Examples/cert/z3_bv_bit_11.proof \
-  SMT/Examples/cert/z3_bv_bit_12 \
-  SMT/Examples/cert/z3_bv_bit_12.proof \
-  SMT/Examples/cert/z3_bv_bit_13 \
-  SMT/Examples/cert/z3_bv_bit_13.proof \
-  SMT/Examples/cert/z3_bv_bit_14 \
-  SMT/Examples/cert/z3_bv_bit_14.proof \
-  SMT/Examples/cert/z3_bv_bit_15 \
-  SMT/Examples/cert/z3_bv_bit_15.proof \
-  SMT/Examples/cert/z3_fol_01 \
-  SMT/Examples/cert/z3_fol_01.proof \
-  SMT/Examples/cert/z3_fol_02 \
-  SMT/Examples/cert/z3_fol_02.proof \
-  SMT/Examples/cert/z3_fol_03 \
-  SMT/Examples/cert/z3_fol_03.proof \
-  SMT/Examples/cert/z3_fol_04 \
-  SMT/Examples/cert/z3_fol_04.proof \
-  SMT/Examples/cert/z3_hol_01 \
-  SMT/Examples/cert/z3_hol_01.proof \
-  SMT/Examples/cert/z3_hol_02 \
-  SMT/Examples/cert/z3_hol_02.proof \
-  SMT/Examples/cert/z3_hol_03 \
-  SMT/Examples/cert/z3_hol_03.proof \
-  SMT/Examples/cert/z3_hol_04 \
-  SMT/Examples/cert/z3_hol_04.proof \
-  SMT/Examples/cert/z3_hol_05 \
-  SMT/Examples/cert/z3_hol_05.proof \
-  SMT/Examples/cert/z3_hol_06 \
-  SMT/Examples/cert/z3_hol_06.proof \
-  SMT/Examples/cert/z3_hol_07 \
-  SMT/Examples/cert/z3_hol_07.proof \
-  SMT/Examples/cert/z3_hol_08 \
-  SMT/Examples/cert/z3_hol_08.proof \
-  SMT/Examples/cert/z3_linarith_01 \
-  SMT/Examples/cert/z3_linarith_01.proof \
-  SMT/Examples/cert/z3_linarith_02 \
-  SMT/Examples/cert/z3_linarith_02.proof \
-  SMT/Examples/cert/z3_linarith_03 \
-  SMT/Examples/cert/z3_linarith_03.proof \
-  SMT/Examples/cert/z3_linarith_04 \
-  SMT/Examples/cert/z3_linarith_04.proof \
-  SMT/Examples/cert/z3_linarith_05 \
-  SMT/Examples/cert/z3_linarith_05.proof \
-  SMT/Examples/cert/z3_linarith_06 \
-  SMT/Examples/cert/z3_linarith_06.proof \
-  SMT/Examples/cert/z3_linarith_07 \
-  SMT/Examples/cert/z3_linarith_07.proof \
-  SMT/Examples/cert/z3_linarith_08 \
-  SMT/Examples/cert/z3_linarith_08.proof \
-  SMT/Examples/cert/z3_linarith_09 \
-  SMT/Examples/cert/z3_linarith_09.proof \
-  SMT/Examples/cert/z3_linarith_10 \
-  SMT/Examples/cert/z3_linarith_10.proof \
-  SMT/Examples/cert/z3_linarith_11 \
-  SMT/Examples/cert/z3_linarith_11.proof \
-  SMT/Examples/cert/z3_linarith_12 \
-  SMT/Examples/cert/z3_linarith_12.proof \
-  SMT/Examples/cert/z3_linarith_13 \
-  SMT/Examples/cert/z3_linarith_13.proof \
-  SMT/Examples/cert/z3_linarith_14 \
-  SMT/Examples/cert/z3_linarith_14.proof \
-  SMT/Examples/cert/z3_linarith_15 \
-  SMT/Examples/cert/z3_linarith_15.proof \
-  SMT/Examples/cert/z3_linarith_16 \
-  SMT/Examples/cert/z3_linarith_16.proof \
-  SMT/Examples/cert/z3_mono_01 \
-  SMT/Examples/cert/z3_mono_01.proof \
-  SMT/Examples/cert/z3_mono_02 \
-  SMT/Examples/cert/z3_mono_02.proof \
-  SMT/Examples/cert/z3_nat_arith_01 \
-  SMT/Examples/cert/z3_nat_arith_01.proof \
-  SMT/Examples/cert/z3_nat_arith_02 \
-  SMT/Examples/cert/z3_nat_arith_02.proof \
-  SMT/Examples/cert/z3_nat_arith_03 \
-  SMT/Examples/cert/z3_nat_arith_03.proof \
-  SMT/Examples/cert/z3_nat_arith_04 \
-  SMT/Examples/cert/z3_nat_arith_04.proof \
-  SMT/Examples/cert/z3_nat_arith_05 \
-  SMT/Examples/cert/z3_nat_arith_05.proof \
-  SMT/Examples/cert/z3_nat_arith_06 \
-  SMT/Examples/cert/z3_nat_arith_06.proof \
-  SMT/Examples/cert/z3_nat_arith_07 \
-  SMT/Examples/cert/z3_nat_arith_07.proof \
-  SMT/Examples/cert/z3_nlarith_01 \
-  SMT/Examples/cert/z3_nlarith_01.proof \
-  SMT/Examples/cert/z3_nlarith_02 \
-  SMT/Examples/cert/z3_nlarith_02.proof \
-  SMT/Examples/cert/z3_nlarith_03 \
-  SMT/Examples/cert/z3_nlarith_03.proof \
-  SMT/Examples/cert/z3_nlarith_04 \
-  SMT/Examples/cert/z3_nlarith_04.proof \
-  SMT/Examples/cert/z3_pair_01 \
-  SMT/Examples/cert/z3_pair_01.proof \
-  SMT/Examples/cert/z3_pair_02 \
-  SMT/Examples/cert/z3_pair_02.proof \
-  SMT/Examples/cert/z3_prop_01 \
-  SMT/Examples/cert/z3_prop_01.proof \
-  SMT/Examples/cert/z3_prop_02 \
-  SMT/Examples/cert/z3_prop_02.proof \
-  SMT/Examples/cert/z3_prop_03 \
-  SMT/Examples/cert/z3_prop_03.proof \
-  SMT/Examples/cert/z3_prop_04 \
-  SMT/Examples/cert/z3_prop_04.proof \
-  SMT/Examples/cert/z3_prop_05 \
-  SMT/Examples/cert/z3_prop_05.proof \
-  SMT/Examples/cert/z3_prop_06 \
-  SMT/Examples/cert/z3_prop_06.proof \
-  SMT/Examples/cert/z3_prop_07 \
-  SMT/Examples/cert/z3_prop_07.proof \
-  SMT/Examples/cert/z3_prop_08 \
-  SMT/Examples/cert/z3_prop_08.proof \
-  SMT/Examples/cert/z3_prop_09 \
-  SMT/Examples/cert/z3_prop_09.proof \
-  SMT/Examples/cert/z3_prop_10 \
+$(LOG)/HOL-SMT-Examples.gz: $(OUT)/HOL-SMT SMT/Examples/ROOT.ML		\
+  SMT/Examples/SMT_Examples.thy SMT/Examples/cert/z3_arith_quant_01	\
+  SMT/Examples/cert/z3_arith_quant_01.proof				\
+  SMT/Examples/cert/z3_arith_quant_02					\
+  SMT/Examples/cert/z3_arith_quant_02.proof				\
+  SMT/Examples/cert/z3_arith_quant_03					\
+  SMT/Examples/cert/z3_arith_quant_03.proof				\
+  SMT/Examples/cert/z3_arith_quant_04					\
+  SMT/Examples/cert/z3_arith_quant_04.proof				\
+  SMT/Examples/cert/z3_arith_quant_05					\
+  SMT/Examples/cert/z3_arith_quant_05.proof				\
+  SMT/Examples/cert/z3_arith_quant_06					\
+  SMT/Examples/cert/z3_arith_quant_06.proof				\
+  SMT/Examples/cert/z3_arith_quant_07					\
+  SMT/Examples/cert/z3_arith_quant_07.proof				\
+  SMT/Examples/cert/z3_arith_quant_08					\
+  SMT/Examples/cert/z3_arith_quant_08.proof				\
+  SMT/Examples/cert/z3_arith_quant_09					\
+  SMT/Examples/cert/z3_arith_quant_09.proof				\
+  SMT/Examples/cert/z3_arith_quant_10					\
+  SMT/Examples/cert/z3_arith_quant_10.proof				\
+  SMT/Examples/cert/z3_arith_quant_11					\
+  SMT/Examples/cert/z3_arith_quant_11.proof				\
+  SMT/Examples/cert/z3_arith_quant_12					\
+  SMT/Examples/cert/z3_arith_quant_12.proof				\
+  SMT/Examples/cert/z3_arith_quant_13					\
+  SMT/Examples/cert/z3_arith_quant_13.proof				\
+  SMT/Examples/cert/z3_arith_quant_14					\
+  SMT/Examples/cert/z3_arith_quant_14.proof				\
+  SMT/Examples/cert/z3_arith_quant_15					\
+  SMT/Examples/cert/z3_arith_quant_15.proof				\
+  SMT/Examples/cert/z3_arith_quant_16					\
+  SMT/Examples/cert/z3_arith_quant_16.proof				\
+  SMT/Examples/cert/z3_arith_quant_17					\
+  SMT/Examples/cert/z3_arith_quant_17.proof				\
+  SMT/Examples/cert/z3_arith_quant_18					\
+  SMT/Examples/cert/z3_arith_quant_18.proof SMT/Examples/cert/z3_bv_01	\
+  SMT/Examples/cert/z3_bv_01.proof SMT/Examples/cert/z3_bv_02		\
+  SMT/Examples/cert/z3_bv_02.proof SMT/Examples/cert/z3_bv_arith_01	\
+  SMT/Examples/cert/z3_bv_arith_01.proof				\
+  SMT/Examples/cert/z3_bv_arith_02					\
+  SMT/Examples/cert/z3_bv_arith_02.proof				\
+  SMT/Examples/cert/z3_bv_arith_03					\
+  SMT/Examples/cert/z3_bv_arith_03.proof				\
+  SMT/Examples/cert/z3_bv_arith_04					\
+  SMT/Examples/cert/z3_bv_arith_04.proof				\
+  SMT/Examples/cert/z3_bv_arith_05					\
+  SMT/Examples/cert/z3_bv_arith_05.proof				\
+  SMT/Examples/cert/z3_bv_arith_06					\
+  SMT/Examples/cert/z3_bv_arith_06.proof				\
+  SMT/Examples/cert/z3_bv_arith_07					\
+  SMT/Examples/cert/z3_bv_arith_07.proof				\
+  SMT/Examples/cert/z3_bv_arith_08					\
+  SMT/Examples/cert/z3_bv_arith_08.proof				\
+  SMT/Examples/cert/z3_bv_arith_09					\
+  SMT/Examples/cert/z3_bv_arith_09.proof				\
+  SMT/Examples/cert/z3_bv_arith_10					\
+  SMT/Examples/cert/z3_bv_arith_10.proof				\
+  SMT/Examples/cert/z3_bv_bit_01 SMT/Examples/cert/z3_bv_bit_01.proof	\
+  SMT/Examples/cert/z3_bv_bit_02 SMT/Examples/cert/z3_bv_bit_02.proof	\
+  SMT/Examples/cert/z3_bv_bit_03 SMT/Examples/cert/z3_bv_bit_03.proof	\
+  SMT/Examples/cert/z3_bv_bit_04 SMT/Examples/cert/z3_bv_bit_04.proof	\
+  SMT/Examples/cert/z3_bv_bit_05 SMT/Examples/cert/z3_bv_bit_05.proof	\
+  SMT/Examples/cert/z3_bv_bit_06 SMT/Examples/cert/z3_bv_bit_06.proof	\
+  SMT/Examples/cert/z3_bv_bit_07 SMT/Examples/cert/z3_bv_bit_07.proof	\
+  SMT/Examples/cert/z3_bv_bit_08 SMT/Examples/cert/z3_bv_bit_08.proof	\
+  SMT/Examples/cert/z3_bv_bit_09 SMT/Examples/cert/z3_bv_bit_09.proof	\
+  SMT/Examples/cert/z3_bv_bit_10 SMT/Examples/cert/z3_bv_bit_10.proof	\
+  SMT/Examples/cert/z3_bv_bit_11 SMT/Examples/cert/z3_bv_bit_11.proof	\
+  SMT/Examples/cert/z3_bv_bit_12 SMT/Examples/cert/z3_bv_bit_12.proof	\
+  SMT/Examples/cert/z3_bv_bit_13 SMT/Examples/cert/z3_bv_bit_13.proof	\
+  SMT/Examples/cert/z3_bv_bit_14 SMT/Examples/cert/z3_bv_bit_14.proof	\
+  SMT/Examples/cert/z3_bv_bit_15 SMT/Examples/cert/z3_bv_bit_15.proof	\
+  SMT/Examples/cert/z3_fol_01 SMT/Examples/cert/z3_fol_01.proof		\
+  SMT/Examples/cert/z3_fol_02 SMT/Examples/cert/z3_fol_02.proof		\
+  SMT/Examples/cert/z3_fol_03 SMT/Examples/cert/z3_fol_03.proof		\
+  SMT/Examples/cert/z3_fol_04 SMT/Examples/cert/z3_fol_04.proof		\
+  SMT/Examples/cert/z3_hol_01 SMT/Examples/cert/z3_hol_01.proof		\
+  SMT/Examples/cert/z3_hol_02 SMT/Examples/cert/z3_hol_02.proof		\
+  SMT/Examples/cert/z3_hol_03 SMT/Examples/cert/z3_hol_03.proof		\
+  SMT/Examples/cert/z3_hol_04 SMT/Examples/cert/z3_hol_04.proof		\
+  SMT/Examples/cert/z3_hol_05 SMT/Examples/cert/z3_hol_05.proof		\
+  SMT/Examples/cert/z3_hol_06 SMT/Examples/cert/z3_hol_06.proof		\
+  SMT/Examples/cert/z3_hol_07 SMT/Examples/cert/z3_hol_07.proof		\
+  SMT/Examples/cert/z3_hol_08 SMT/Examples/cert/z3_hol_08.proof		\
+  SMT/Examples/cert/z3_linarith_01					\
+  SMT/Examples/cert/z3_linarith_01.proof				\
+  SMT/Examples/cert/z3_linarith_02					\
+  SMT/Examples/cert/z3_linarith_02.proof				\
+  SMT/Examples/cert/z3_linarith_03					\
+  SMT/Examples/cert/z3_linarith_03.proof				\
+  SMT/Examples/cert/z3_linarith_04					\
+  SMT/Examples/cert/z3_linarith_04.proof				\
+  SMT/Examples/cert/z3_linarith_05					\
+  SMT/Examples/cert/z3_linarith_05.proof				\
+  SMT/Examples/cert/z3_linarith_06					\
+  SMT/Examples/cert/z3_linarith_06.proof				\
+  SMT/Examples/cert/z3_linarith_07					\
+  SMT/Examples/cert/z3_linarith_07.proof				\
+  SMT/Examples/cert/z3_linarith_08					\
+  SMT/Examples/cert/z3_linarith_08.proof				\
+  SMT/Examples/cert/z3_linarith_09					\
+  SMT/Examples/cert/z3_linarith_09.proof				\
+  SMT/Examples/cert/z3_linarith_10					\
+  SMT/Examples/cert/z3_linarith_10.proof				\
+  SMT/Examples/cert/z3_linarith_11					\
+  SMT/Examples/cert/z3_linarith_11.proof				\
+  SMT/Examples/cert/z3_linarith_12					\
+  SMT/Examples/cert/z3_linarith_12.proof				\
+  SMT/Examples/cert/z3_linarith_13					\
+  SMT/Examples/cert/z3_linarith_13.proof				\
+  SMT/Examples/cert/z3_linarith_14					\
+  SMT/Examples/cert/z3_linarith_14.proof				\
+  SMT/Examples/cert/z3_linarith_15					\
+  SMT/Examples/cert/z3_linarith_15.proof				\
+  SMT/Examples/cert/z3_linarith_16					\
+  SMT/Examples/cert/z3_linarith_16.proof SMT/Examples/cert/z3_mono_01	\
+  SMT/Examples/cert/z3_mono_01.proof SMT/Examples/cert/z3_mono_02	\
+  SMT/Examples/cert/z3_mono_02.proof SMT/Examples/cert/z3_nat_arith_01	\
+  SMT/Examples/cert/z3_nat_arith_01.proof				\
+  SMT/Examples/cert/z3_nat_arith_02					\
+  SMT/Examples/cert/z3_nat_arith_02.proof				\
+  SMT/Examples/cert/z3_nat_arith_03					\
+  SMT/Examples/cert/z3_nat_arith_03.proof				\
+  SMT/Examples/cert/z3_nat_arith_04					\
+  SMT/Examples/cert/z3_nat_arith_04.proof				\
+  SMT/Examples/cert/z3_nat_arith_05					\
+  SMT/Examples/cert/z3_nat_arith_05.proof				\
+  SMT/Examples/cert/z3_nat_arith_06					\
+  SMT/Examples/cert/z3_nat_arith_06.proof				\
+  SMT/Examples/cert/z3_nat_arith_07					\
+  SMT/Examples/cert/z3_nat_arith_07.proof				\
+  SMT/Examples/cert/z3_nlarith_01					\
+  SMT/Examples/cert/z3_nlarith_01.proof					\
+  SMT/Examples/cert/z3_nlarith_02					\
+  SMT/Examples/cert/z3_nlarith_02.proof					\
+  SMT/Examples/cert/z3_nlarith_03					\
+  SMT/Examples/cert/z3_nlarith_03.proof					\
+  SMT/Examples/cert/z3_nlarith_04					\
+  SMT/Examples/cert/z3_nlarith_04.proof SMT/Examples/cert/z3_pair_01	\
+  SMT/Examples/cert/z3_pair_01.proof SMT/Examples/cert/z3_pair_02	\
+  SMT/Examples/cert/z3_pair_02.proof SMT/Examples/cert/z3_prop_01	\
+  SMT/Examples/cert/z3_prop_01.proof SMT/Examples/cert/z3_prop_02	\
+  SMT/Examples/cert/z3_prop_02.proof SMT/Examples/cert/z3_prop_03	\
+  SMT/Examples/cert/z3_prop_03.proof SMT/Examples/cert/z3_prop_04	\
+  SMT/Examples/cert/z3_prop_04.proof SMT/Examples/cert/z3_prop_05	\
+  SMT/Examples/cert/z3_prop_05.proof SMT/Examples/cert/z3_prop_06	\
+  SMT/Examples/cert/z3_prop_06.proof SMT/Examples/cert/z3_prop_07	\
+  SMT/Examples/cert/z3_prop_07.proof SMT/Examples/cert/z3_prop_08	\
+  SMT/Examples/cert/z3_prop_08.proof SMT/Examples/cert/z3_prop_09	\
+  SMT/Examples/cert/z3_prop_09.proof SMT/Examples/cert/z3_prop_10	\
   SMT/Examples/cert/z3_prop_10.proof
 	@cd SMT; $(ISABELLE_TOOL) usedir $(OUT)/HOL-SMT Examples
 
@@ -1367,7 +1304,7 @@
 clean:
 	@rm -f $(OUT)/HOL-Plain $(OUT)/HOL-Main $(OUT)/HOL		\
 		$(OUT)/HOL-Nominal $(OUT)/TLA $(LOG)/HOL.gz		\
-		$(LOG)/TLA.gz $(LOG)/HOL-Isar_examples.gz		\
+		$(LOG)/TLA.gz $(LOG)/HOL-Isar_Examples.gz		\
 		$(LOG)/HOL-Induct.gz $(LOG)/HOL-ex.gz			\
 		$(LOG)/HOL-Subst.gz $(LOG)/HOL-IMP.gz			\
 		$(LOG)/HOL-IMPP.gz $(LOG)/HOL-Hoare.gz			\
@@ -1378,7 +1315,7 @@
 		$(LOG)/HOL-MicroJava.gz $(LOG)/HOL-NanoJava.gz		\
 		$(LOG)/HOL-Nominal-Examples.gz $(LOG)/HOL-IOA.gz	\
 		$(LOG)/HOL-Lattice $(LOG)/HOL-Matrix			\
-		$(LOG)/HOL-Hahn_Banach.gz $(LOG)/HOL-SET-Protocol.gz	\
+		$(LOG)/HOL-Hahn_Banach.gz $(LOG)/HOL-SET_Protocol.gz	\
 		$(LOG)/TLA-Inc.gz $(LOG)/TLA-Buffer.gz			\
 		$(LOG)/TLA-Memory.gz $(LOG)/HOL-Library.gz		\
 		$(LOG)/HOL-Unix.gz $(OUT)/HOL-Word $(LOG)/HOL-Word.gz	\
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_Examples/Basic_Logic.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,448 @@
+(*  Title:      HOL/Isar_Examples/Basic_Logic.thy
+    Author:     Markus Wenzel, TU Muenchen
+
+Basic propositional and quantifier reasoning.
+*)
+
+header {* Basic logical reasoning *}
+
+theory Basic_Logic
+imports Main
+begin
+
+
+subsection {* Pure backward reasoning *}
+
+text {*
+  In order to get a first idea of how Isabelle/Isar proof documents
+  may look like, we consider the propositions @{text I}, @{text K},
+  and @{text S}.  The following (rather explicit) proofs should
+  require little extra explanations.
+*}
+
+lemma I: "A --> A"
+proof
+  assume A
+  show A by fact
+qed
+
+lemma K: "A --> B --> A"
+proof
+  assume A
+  show "B --> A"
+  proof
+    show A by fact
+  qed
+qed
+
+lemma S: "(A --> B --> C) --> (A --> B) --> A --> C"
+proof
+  assume "A --> B --> C"
+  show "(A --> B) --> A --> C"
+  proof
+    assume "A --> B"
+    show "A --> C"
+    proof
+      assume A
+      show C
+      proof (rule mp)
+        show "B --> C" by (rule mp) fact+
+        show B by (rule mp) fact+
+      qed
+    qed
+  qed
+qed
+
+text {*
+  Isar provides several ways to fine-tune the reasoning, avoiding
+  excessive detail.  Several abbreviated language elements are
+  available, enabling the writer to express proofs in a more concise
+  way, even without referring to any automated proof tools yet.
+
+  First of all, proof by assumption may be abbreviated as a single
+  dot.
+*}
+
+lemma "A --> A"
+proof
+  assume A
+  show A by fact+
+qed
+
+text {*
+  In fact, concluding any (sub-)proof already involves solving any
+  remaining goals by assumption\footnote{This is not a completely
+  trivial operation, as proof by assumption may involve full
+  higher-order unification.}.  Thus we may skip the rather vacuous
+  body of the above proof as well.
+*}
+
+lemma "A --> A"
+proof
+qed
+
+text {*
+  Note that the \isacommand{proof} command refers to the @{text rule}
+  method (without arguments) by default.  Thus it implicitly applies a
+  single rule, as determined from the syntactic form of the statements
+  involved.  The \isacommand{by} command abbreviates any proof with
+  empty body, so the proof may be further pruned.
+*}
+
+lemma "A --> A"
+  by rule
+
+text {*
+  Proof by a single rule may be abbreviated as double-dot.
+*}
+
+lemma "A --> A" ..
+
+text {*
+  Thus we have arrived at an adequate representation of the proof of a
+  tautology that holds by a single standard rule.\footnote{Apparently,
+  the rule here is implication introduction.}
+*}
+
+text {*
+  Let us also reconsider @{text K}.  Its statement is composed of
+  iterated connectives.  Basic decomposition is by a single rule at a
+  time, which is why our first version above was by nesting two
+  proofs.
+
+  The @{text intro} proof method repeatedly decomposes a goal's
+  conclusion.\footnote{The dual method is @{text elim}, acting on a
+  goal's premises.}
+*}
+
+lemma "A --> B --> A"
+proof (intro impI)
+  assume A
+  show A by fact
+qed
+
+text {*
+  Again, the body may be collapsed.
+*}
+
+lemma "A --> B --> A"
+  by (intro impI)
+
+text {*
+  Just like @{text rule}, the @{text intro} and @{text elim} proof
+  methods pick standard structural rules, in case no explicit
+  arguments are given.  While implicit rules are usually just fine for
+  single rule application, this may go too far with iteration.  Thus
+  in practice, @{text intro} and @{text elim} would be typically
+  restricted to certain structures by giving a few rules only, e.g.\
+  \isacommand{proof}~@{text "(intro impI allI)"} to strip implications
+  and universal quantifiers.
+
+  Such well-tuned iterated decomposition of certain structures is the
+  prime application of @{text intro} and @{text elim}.  In contrast,
+  terminal steps that solve a goal completely are usually performed by
+  actual automated proof methods (such as \isacommand{by}~@{text
+  blast}.
+*}
+
+
+subsection {* Variations of backward vs.\ forward reasoning *}
+
+text {*
+  Certainly, any proof may be performed in backward-style only.  On
+  the other hand, small steps of reasoning are often more naturally
+  expressed in forward-style.  Isar supports both backward and forward
+  reasoning as a first-class concept.  In order to demonstrate the
+  difference, we consider several proofs of @{text "A \<and> B \<longrightarrow> B \<and> A"}.
+
+  The first version is purely backward.
+*}
+
+lemma "A & B --> B & A"
+proof
+  assume "A & B"
+  show "B & A"
+  proof
+    show B by (rule conjunct2) fact
+    show A by (rule conjunct1) fact
+  qed
+qed
+
+text {*
+  Above, the @{text "conjunct_1/2"} projection rules had to be named
+  explicitly, since the goals @{text B} and @{text A} did not provide
+  any structural clue.  This may be avoided using \isacommand{from} to
+  focus on the @{text "A \<and> B"} assumption as the current facts,
+  enabling the use of double-dot proofs.  Note that \isacommand{from}
+  already does forward-chaining, involving the \name{conjE} rule here.
+*}
+
+lemma "A & B --> B & A"
+proof
+  assume "A & B"
+  show "B & A"
+  proof
+    from `A & B` show B ..
+    from `A & B` show A ..
+  qed
+qed
+
+text {*
+  In the next version, we move the forward step one level upwards.
+  Forward-chaining from the most recent facts is indicated by the
+  \isacommand{then} command.  Thus the proof of @{text "B \<and> A"} from
+  @{text "A \<and> B"} actually becomes an elimination, rather than an
+  introduction.  The resulting proof structure directly corresponds to
+  that of the @{text conjE} rule, including the repeated goal
+  proposition that is abbreviated as @{text ?thesis} below.
+*}
+
+lemma "A & B --> B & A"
+proof
+  assume "A & B"
+  then show "B & A"
+  proof                    -- {* rule @{text conjE} of @{text "A \<and> B"} *}
+    assume B A
+    then show ?thesis ..   -- {* rule @{text conjI} of @{text "B \<and> A"} *}
+  qed
+qed
+
+text {*
+  In the subsequent version we flatten the structure of the main body
+  by doing forward reasoning all the time.  Only the outermost
+  decomposition step is left as backward.
+*}
+
+lemma "A & B --> B & A"
+proof
+  assume "A & B"
+  from `A & B` have A ..
+  from `A & B` have B ..
+  from `B` `A` show "B & A" ..
+qed
+
+text {*
+  We can still push forward-reasoning a bit further, even at the risk
+  of getting ridiculous.  Note that we force the initial proof step to
+  do nothing here, by referring to the ``-'' proof method.
+*}
+
+lemma "A & B --> B & A"
+proof -
+  {
+    assume "A & B"
+    from `A & B` have A ..
+    from `A & B` have B ..
+    from `B` `A` have "B & A" ..
+  }
+  then show ?thesis ..         -- {* rule \name{impI} *}
+qed
+
+text {*
+  \medskip With these examples we have shifted through a whole range
+  from purely backward to purely forward reasoning.  Apparently, in
+  the extreme ends we get slightly ill-structured proofs, which also
+  require much explicit naming of either rules (backward) or local
+  facts (forward).
+
+  The general lesson learned here is that good proof style would
+  achieve just the \emph{right} balance of top-down backward
+  decomposition, and bottom-up forward composition.  In general, there
+  is no single best way to arrange some pieces of formal reasoning, of
+  course.  Depending on the actual applications, the intended audience
+  etc., rules (and methods) on the one hand vs.\ facts on the other
+  hand have to be emphasized in an appropriate way.  This requires the
+  proof writer to develop good taste, and some practice, of course.
+*}
+
+text {*
+  For our example the most appropriate way of reasoning is probably
+  the middle one, with conjunction introduction done after
+  elimination.
+*}
+
+lemma "A & B --> B & A"
+proof
+  assume "A & B"
+  then show "B & A"
+  proof
+    assume B A
+    then show ?thesis ..
+  qed
+qed
+
+
+
+subsection {* A few examples from ``Introduction to Isabelle'' *}
+
+text {*
+  We rephrase some of the basic reasoning examples of
+  \cite{isabelle-intro}, using HOL rather than FOL.
+*}
+
+subsubsection {* A propositional proof *}
+
+text {*
+  We consider the proposition @{text "P \<or> P \<longrightarrow> P"}.  The proof below
+  involves forward-chaining from @{text "P \<or> P"}, followed by an
+  explicit case-analysis on the two \emph{identical} cases.
+*}
+
+lemma "P | P --> P"
+proof
+  assume "P | P"
+  then show P
+  proof                    -- {*
+    rule @{text disjE}: \smash{$\infer{C}{A \disj B & \infer*{C}{[A]} & \infer*{C}{[B]}}$}
+  *}
+    assume P show P by fact
+  next
+    assume P show P by fact
+  qed
+qed
+
+text {*
+  Case splits are \emph{not} hardwired into the Isar language as a
+  special feature.  The \isacommand{next} command used to separate the
+  cases above is just a short form of managing block structure.
+
+  \medskip In general, applying proof methods may split up a goal into
+  separate ``cases'', i.e.\ new subgoals with individual local
+  assumptions.  The corresponding proof text typically mimics this by
+  establishing results in appropriate contexts, separated by blocks.
+
+  In order to avoid too much explicit parentheses, the Isar system
+  implicitly opens an additional block for any new goal, the
+  \isacommand{next} statement then closes one block level, opening a
+  new one.  The resulting behavior is what one would expect from
+  separating cases, only that it is more flexible.  E.g.\ an induction
+  base case (which does not introduce local assumptions) would
+  \emph{not} require \isacommand{next} to separate the subsequent step
+  case.
+
+  \medskip In our example the situation is even simpler, since the two
+  cases actually coincide.  Consequently the proof may be rephrased as
+  follows.
+*}
+
+lemma "P | P --> P"
+proof
+  assume "P | P"
+  then show P
+  proof
+    assume P
+    show P by fact
+    show P by fact
+  qed
+qed
+
+text {*
+  Again, the rather vacuous body of the proof may be collapsed.  Thus
+  the case analysis degenerates into two assumption steps, which are
+  implicitly performed when concluding the single rule step of the
+  double-dot proof as follows.
+*}
+
+lemma "P | P --> P"
+proof
+  assume "P | P"
+  then show P ..
+qed
+
+
+subsubsection {* A quantifier proof *}
+
+text {*
+  To illustrate quantifier reasoning, let us prove @{text "(\<exists>x. P (f
+  x)) \<longrightarrow> (\<exists>y. P y)"}.  Informally, this holds because any @{text a}
+  with @{text "P (f a)"} may be taken as a witness for the second
+  existential statement.
+
+  The first proof is rather verbose, exhibiting quite a lot of
+  (redundant) detail.  It gives explicit rules, even with some
+  instantiation.  Furthermore, we encounter two new language elements:
+  the \isacommand{fix} command augments the context by some new
+  ``arbitrary, but fixed'' element; the \isacommand{is} annotation
+  binds term abbreviations by higher-order pattern matching.
+*}
+
+lemma "(EX x. P (f x)) --> (EX y. P y)"
+proof
+  assume "EX x. P (f x)"
+  then show "EX y. P y"
+  proof (rule exE)             -- {*
+    rule \name{exE}: \smash{$\infer{B}{\ex x A(x) & \infer*{B}{[A(x)]_x}}$}
+  *}
+    fix a
+    assume "P (f a)" (is "P ?witness")
+    then show ?thesis by (rule exI [of P ?witness])
+  qed
+qed
+
+text {*
+  While explicit rule instantiation may occasionally improve
+  readability of certain aspects of reasoning, it is usually quite
+  redundant.  Above, the basic proof outline gives already enough
+  structural clues for the system to infer both the rules and their
+  instances (by higher-order unification).  Thus we may as well prune
+  the text as follows.
+*}
+
+lemma "(EX x. P (f x)) --> (EX y. P y)"
+proof
+  assume "EX x. P (f x)"
+  then show "EX y. P y"
+  proof
+    fix a
+    assume "P (f a)"
+    then show ?thesis ..
+  qed
+qed
+
+text {*
+  Explicit @{text \<exists>}-elimination as seen above can become quite
+  cumbersome in practice.  The derived Isar language element
+  ``\isakeyword{obtain}'' provides a more handsome way to do
+  generalized existence reasoning.
+*}
+
+lemma "(EX x. P (f x)) --> (EX y. P y)"
+proof
+  assume "EX x. P (f x)"
+  then obtain a where "P (f a)" ..
+  then show "EX y. P y" ..
+qed
+
+text {*
+  Technically, \isakeyword{obtain} is similar to \isakeyword{fix} and
+  \isakeyword{assume} together with a soundness proof of the
+  elimination involved.  Thus it behaves similar to any other forward
+  proof element.  Also note that due to the nature of general
+  existence reasoning involved here, any result exported from the
+  context of an \isakeyword{obtain} statement may \emph{not} refer to
+  the parameters introduced there.
+*}
+
+
+
+subsubsection {* Deriving rules in Isabelle *}
+
+text {*
+  We derive the conjunction elimination rule from the corresponding
+  projections.  The proof is quite straight-forward, since
+  Isabelle/Isar supports non-atomic goals and assumptions fully
+  transparently.
+*}
+
+theorem conjE: "A & B ==> (A ==> B ==> C) ==> C"
+proof -
+  assume "A & B"
+  assume r: "A ==> B ==> C"
+  show C
+  proof (rule r)
+    show A by (rule conjunct1) fact
+    show B by (rule conjunct2) fact
+  qed
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_Examples/Cantor.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,71 @@
+(*  Title:      HOL/Isar_Examples/Cantor.thy
+    Author:     Markus Wenzel, TU Muenchen
+*)
+
+header {* Cantor's Theorem *}
+
+theory Cantor
+imports Main
+begin
+
+text_raw {*
+  \footnote{This is an Isar version of the final example of the
+  Isabelle/HOL manual \cite{isabelle-HOL}.}
+*}
+
+text {*
+  Cantor's Theorem states that every set has more subsets than it has
+  elements.  It has become a favorite basic example in pure
+  higher-order logic since it is so easily expressed: \[\all{f::\alpha
+  \To \alpha \To \idt{bool}} \ex{S::\alpha \To \idt{bool}}
+  \all{x::\alpha} f \ap x \not= S\]
+
+  Viewing types as sets, $\alpha \To \idt{bool}$ represents the
+  powerset of $\alpha$.  This version of the theorem states that for
+  every function from $\alpha$ to its powerset, some subset is outside
+  its range.  The Isabelle/Isar proofs below uses HOL's set theory,
+  with the type $\alpha \ap \idt{set}$ and the operator
+  $\idt{range}::(\alpha \To \beta) \To \beta \ap \idt{set}$.
+*}
+
+theorem "EX S. S ~: range (f :: 'a => 'a set)"
+proof
+  let ?S = "{x. x ~: f x}"
+  show "?S ~: range f"
+  proof
+    assume "?S : range f"
+    then obtain y where "?S = f y" ..
+    then show False
+    proof (rule equalityCE)
+      assume "y : f y"
+      assume "y : ?S" then have "y ~: f y" ..
+      with `y : f y` show ?thesis by contradiction
+    next
+      assume "y ~: ?S"
+      assume "y ~: f y" then have "y : ?S" ..
+      with `y ~: ?S` show ?thesis by contradiction
+    qed
+  qed
+qed
+
+text {*
+  How much creativity is required?  As it happens, Isabelle can prove
+  this theorem automatically using best-first search.  Depth-first
+  search would diverge, but best-first search successfully navigates
+  through the large search space.  The context of Isabelle's classical
+  prover contains rules for the relevant constructs of HOL's set
+  theory.
+*}
+
+theorem "EX S. S ~: range (f :: 'a => 'a set)"
+  by best
+
+text {*
+  While this establishes the same theorem internally, we do not get
+  any idea of how the proof actually works.  There is currently no way
+  to transform internal system-level representations of Isabelle
+  proofs back into Isar text.  Writing intelligible proof documents
+  really is a creative process, after all.
+*}
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_Examples/Drinker.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,54 @@
+(*  Title:      HOL/Isar_Examples/Drinker.thy
+    Author:     Makarius
+*)
+
+header {* The Drinker's Principle *}
+
+theory Drinker
+imports Main
+begin
+
+text {*
+  Here is another example of classical reasoning: the Drinker's
+  Principle says that for some person, if he is drunk, everybody else
+  is drunk!
+
+  We first prove a classical part of de-Morgan's law.
+*}
+
+lemma deMorgan:
+  assumes "\<not> (\<forall>x. P x)"
+  shows "\<exists>x. \<not> P x"
+  using prems
+proof (rule contrapos_np)
+  assume a: "\<not> (\<exists>x. \<not> P x)"
+  show "\<forall>x. P x"
+  proof
+    fix x
+    show "P x"
+    proof (rule classical)
+      assume "\<not> P x"
+      then have "\<exists>x. \<not> P x" ..
+      with a show ?thesis by contradiction
+    qed
+  qed
+qed
+
+theorem Drinker's_Principle: "\<exists>x. drunk x \<longrightarrow> (\<forall>x. drunk x)"
+proof cases
+  fix a assume "\<forall>x. drunk x"
+  then have "drunk a \<longrightarrow> (\<forall>x. drunk x)" ..
+  then show ?thesis ..
+next
+  assume "\<not> (\<forall>x. drunk x)"
+  then have "\<exists>x. \<not> drunk x" by (rule deMorgan)
+  then obtain a where a: "\<not> drunk a" ..
+  have "drunk a \<longrightarrow> (\<forall>x. drunk x)"
+  proof
+    assume "drunk a"
+    with a show "\<forall>x. drunk x" by (contradiction)
+  qed
+  then show ?thesis ..
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_Examples/Expr_Compiler.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,231 @@
+(*  Title:      HOL/Isar_Examples/Expr_Compiler.thy
+    Author:     Markus Wenzel, TU Muenchen
+
+Correctness of a simple expression/stack-machine compiler.
+*)
+
+header {* Correctness of a simple expression compiler *}
+
+theory Expr_Compiler
+imports Main
+begin
+
+text {*
+ This is a (rather trivial) example of program verification.  We model
+ a compiler for translating expressions to stack machine instructions,
+ and prove its correctness wrt.\ some evaluation semantics.
+*}
+
+
+subsection {* Binary operations *}
+
+text {*
+ Binary operations are just functions over some type of values.  This
+ is both for abstract syntax and semantics, i.e.\ we use a ``shallow
+ embedding'' here.
+*}
+
+types
+  'val binop = "'val => 'val => 'val"
+
+
+subsection {* Expressions *}
+
+text {*
+ The language of expressions is defined as an inductive type,
+ consisting of variables, constants, and binary operations on
+ expressions.
+*}
+
+datatype ('adr, 'val) expr =
+  Variable 'adr |
+  Constant 'val |
+  Binop "'val binop" "('adr, 'val) expr" "('adr, 'val) expr"
+
+text {*
+ Evaluation (wrt.\ some environment of variable assignments) is
+ defined by primitive recursion over the structure of expressions.
+*}
+
+consts
+  eval :: "('adr, 'val) expr => ('adr => 'val) => 'val"
+
+primrec
+  "eval (Variable x) env = env x"
+  "eval (Constant c) env = c"
+  "eval (Binop f e1 e2) env = f (eval e1 env) (eval e2 env)"
+
+
+subsection {* Machine *}
+
+text {*
+ Next we model a simple stack machine, with three instructions.
+*}
+
+datatype ('adr, 'val) instr =
+  Const 'val |
+  Load 'adr |
+  Apply "'val binop"
+
+text {*
+ Execution of a list of stack machine instructions is easily defined
+ as follows.
+*}
+
+consts
+  exec :: "(('adr, 'val) instr) list
+    => 'val list => ('adr => 'val) => 'val list"
+
+primrec
+  "exec [] stack env = stack"
+  "exec (instr # instrs) stack env =
+    (case instr of
+      Const c => exec instrs (c # stack) env
+    | Load x => exec instrs (env x # stack) env
+    | Apply f => exec instrs (f (hd stack) (hd (tl stack))
+                   # (tl (tl stack))) env)"
+
+constdefs
+  execute :: "(('adr, 'val) instr) list => ('adr => 'val) => 'val"
+  "execute instrs env == hd (exec instrs [] env)"
+
+
+subsection {* Compiler *}
+
+text {*
+ We are ready to define the compilation function of expressions to
+ lists of stack machine instructions.
+*}
+
+consts
+  compile :: "('adr, 'val) expr => (('adr, 'val) instr) list"
+
+primrec
+  "compile (Variable x) = [Load x]"
+  "compile (Constant c) = [Const c]"
+  "compile (Binop f e1 e2) = compile e2 @ compile e1 @ [Apply f]"
+
+
+text {*
+ The main result of this development is the correctness theorem for
+ $\idt{compile}$.  We first establish a lemma about $\idt{exec}$ and
+ list append.
+*}
+
+lemma exec_append:
+  "exec (xs @ ys) stack env =
+    exec ys (exec xs stack env) env"
+proof (induct xs arbitrary: stack)
+  case Nil
+  show ?case by simp
+next
+  case (Cons x xs)
+  show ?case
+  proof (induct x)
+    case Const
+    from Cons show ?case by simp
+  next
+    case Load
+    from Cons show ?case by simp
+  next
+    case Apply
+    from Cons show ?case by simp
+  qed
+qed
+
+theorem correctness: "execute (compile e) env = eval e env"
+proof -
+  have "\<And>stack. exec (compile e) stack env = eval e env # stack"
+  proof (induct e)
+    case Variable show ?case by simp
+  next
+    case Constant show ?case by simp
+  next
+    case Binop then show ?case by (simp add: exec_append)
+  qed
+  then show ?thesis by (simp add: execute_def)
+qed
+
+
+text {*
+ \bigskip In the proofs above, the \name{simp} method does quite a lot
+ of work behind the scenes (mostly ``functional program execution'').
+ Subsequently, the same reasoning is elaborated in detail --- at most
+ one recursive function definition is used at a time.  Thus we get a
+ better idea of what is actually going on.
+*}
+
+lemma exec_append':
+  "exec (xs @ ys) stack env = exec ys (exec xs stack env) env"
+proof (induct xs arbitrary: stack)
+  case (Nil s)
+  have "exec ([] @ ys) s env = exec ys s env" by simp
+  also have "... = exec ys (exec [] s env) env" by simp
+  finally show ?case .
+next
+  case (Cons x xs s)
+  show ?case
+  proof (induct x)
+    case (Const val)
+    have "exec ((Const val # xs) @ ys) s env = exec (Const val # xs @ ys) s env"
+      by simp
+    also have "... = exec (xs @ ys) (val # s) env" by simp
+    also from Cons have "... = exec ys (exec xs (val # s) env) env" .
+    also have "... = exec ys (exec (Const val # xs) s env) env" by simp
+    finally show ?case .
+  next
+    case (Load adr)
+    from Cons show ?case by simp -- {* same as above *}
+  next
+    case (Apply fn)
+    have "exec ((Apply fn # xs) @ ys) s env =
+        exec (Apply fn # xs @ ys) s env" by simp
+    also have "... =
+        exec (xs @ ys) (fn (hd s) (hd (tl s)) # (tl (tl s))) env" by simp
+    also from Cons have "... =
+        exec ys (exec xs (fn (hd s) (hd (tl s)) # tl (tl s)) env) env" .
+    also have "... = exec ys (exec (Apply fn # xs) s env) env" by simp
+    finally show ?case .
+  qed
+qed
+
+theorem correctness': "execute (compile e) env = eval e env"
+proof -
+  have exec_compile: "\<And>stack. exec (compile e) stack env = eval e env # stack"
+  proof (induct e)
+    case (Variable adr s)
+    have "exec (compile (Variable adr)) s env = exec [Load adr] s env"
+      by simp
+    also have "... = env adr # s" by simp
+    also have "env adr = eval (Variable adr) env" by simp
+    finally show ?case .
+  next
+    case (Constant val s)
+    show ?case by simp -- {* same as above *}
+  next
+    case (Binop fn e1 e2 s)
+    have "exec (compile (Binop fn e1 e2)) s env =
+        exec (compile e2 @ compile e1 @ [Apply fn]) s env" by simp
+    also have "... = exec [Apply fn]
+        (exec (compile e1) (exec (compile e2) s env) env) env"
+      by (simp only: exec_append)
+    also have "exec (compile e2) s env = eval e2 env # s" by fact
+    also have "exec (compile e1) ... env = eval e1 env # ..." by fact
+    also have "exec [Apply fn] ... env =
+        fn (hd ...) (hd (tl ...)) # (tl (tl ...))" by simp
+    also have "... = fn (eval e1 env) (eval e2 env) # s" by simp
+    also have "fn (eval e1 env) (eval e2 env) =
+        eval (Binop fn e1 e2) env"
+      by simp
+    finally show ?case .
+  qed
+
+  have "execute (compile e) env = hd (exec (compile e) [] env)"
+    by (simp add: execute_def)
+  also from exec_compile
+    have "exec (compile e) [] env = [eval e env]" .
+  also have "hd ... = eval e env" by simp
+  finally show ?thesis .
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_Examples/Fibonacci.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,165 @@
+(*  Title:      HOL/Isar_Examples/Fibonacci.thy
+    Author:     Gertrud Bauer
+    Copyright   1999 Technische Universitaet Muenchen
+
+The Fibonacci function.  Demonstrates the use of recdef.  Original
+tactic script by Lawrence C Paulson.
+
+Fibonacci numbers: proofs of laws taken from
+
+  R. L. Graham, D. E. Knuth, O. Patashnik.
+  Concrete Mathematics.
+  (Addison-Wesley, 1989)
+*)
+
+header {* Fib and Gcd commute *}
+
+theory Fibonacci
+imports Primes
+begin
+
+text_raw {*
+ \footnote{Isar version by Gertrud Bauer.  Original tactic script by
+ Larry Paulson.  A few proofs of laws taken from
+ \cite{Concrete-Math}.}
+*}
+
+
+subsection {* Fibonacci numbers *}
+
+fun fib :: "nat \<Rightarrow> nat" where
+  "fib 0 = 0"
+  | "fib (Suc 0) = 1"
+  | "fib (Suc (Suc x)) = fib x + fib (Suc x)"
+
+lemma [simp]: "0 < fib (Suc n)"
+  by (induct n rule: fib.induct) simp_all
+
+
+text {* Alternative induction rule. *}
+
+theorem fib_induct:
+    "P 0 ==> P 1 ==> (!!n. P (n + 1) ==> P n ==> P (n + 2)) ==> P (n::nat)"
+  by (induct rule: fib.induct) simp_all
+
+
+subsection {* Fib and gcd commute *}
+
+text {* A few laws taken from \cite{Concrete-Math}. *}
+
+lemma fib_add:
+  "fib (n + k + 1) = fib (k + 1) * fib (n + 1) + fib k * fib n"
+  (is "?P n")
+  -- {* see \cite[page 280]{Concrete-Math} *}
+proof (induct n rule: fib_induct)
+  show "?P 0" by simp
+  show "?P 1" by simp
+  fix n
+  have "fib (n + 2 + k + 1)
+    = fib (n + k + 1) + fib (n + 1 + k + 1)" by simp
+  also assume "fib (n + k + 1)
+    = fib (k + 1) * fib (n + 1) + fib k * fib n"
+      (is " _ = ?R1")
+  also assume "fib (n + 1 + k + 1)
+    = fib (k + 1) * fib (n + 1 + 1) + fib k * fib (n + 1)"
+      (is " _ = ?R2")
+  also have "?R1 + ?R2
+    = fib (k + 1) * fib (n + 2 + 1) + fib k * fib (n + 2)"
+    by (simp add: add_mult_distrib2)
+  finally show "?P (n + 2)" .
+qed
+
+lemma gcd_fib_Suc_eq_1: "gcd (fib n) (fib (n + 1)) = 1" (is "?P n")
+proof (induct n rule: fib_induct)
+  show "?P 0" by simp
+  show "?P 1" by simp
+  fix n
+  have "fib (n + 2 + 1) = fib (n + 1) + fib (n + 2)"
+    by simp
+  also have "gcd (fib (n + 2)) ... = gcd (fib (n + 2)) (fib (n + 1))"
+    by (simp only: gcd_add2')
+  also have "... = gcd (fib (n + 1)) (fib (n + 1 + 1))"
+    by (simp add: gcd_commute)
+  also assume "... = 1"
+  finally show "?P (n + 2)" .
+qed
+
+lemma gcd_mult_add: "0 < n ==> gcd (n * k + m) n = gcd m n"
+proof -
+  assume "0 < n"
+  then have "gcd (n * k + m) n = gcd n (m mod n)"
+    by (simp add: gcd_non_0 add_commute)
+  also from `0 < n` have "... = gcd m n" by (simp add: gcd_non_0)
+  finally show ?thesis .
+qed
+
+lemma gcd_fib_add: "gcd (fib m) (fib (n + m)) = gcd (fib m) (fib n)"
+proof (cases m)
+  case 0
+  then show ?thesis by simp
+next
+  case (Suc k)
+  then have "gcd (fib m) (fib (n + m)) = gcd (fib (n + k + 1)) (fib (k + 1))"
+    by (simp add: gcd_commute)
+  also have "fib (n + k + 1)
+    = fib (k + 1) * fib (n + 1) + fib k * fib n"
+    by (rule fib_add)
+  also have "gcd ... (fib (k + 1)) = gcd (fib k * fib n) (fib (k + 1))"
+    by (simp add: gcd_mult_add)
+  also have "... = gcd (fib n) (fib (k + 1))"
+    by (simp only: gcd_fib_Suc_eq_1 gcd_mult_cancel)
+  also have "... = gcd (fib m) (fib n)"
+    using Suc by (simp add: gcd_commute)
+  finally show ?thesis .
+qed
+
+lemma gcd_fib_diff:
+  assumes "m <= n"
+  shows "gcd (fib m) (fib (n - m)) = gcd (fib m) (fib n)"
+proof -
+  have "gcd (fib m) (fib (n - m)) = gcd (fib m) (fib (n - m + m))"
+    by (simp add: gcd_fib_add)
+  also from `m <= n` have "n - m + m = n" by simp
+  finally show ?thesis .
+qed
+
+lemma gcd_fib_mod:
+  assumes "0 < m"
+  shows "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
+proof (induct n rule: nat_less_induct)
+  case (1 n) note hyp = this
+  show ?case
+  proof -
+    have "n mod m = (if n < m then n else (n - m) mod m)"
+      by (rule mod_if)
+    also have "gcd (fib m) (fib ...) = gcd (fib m) (fib n)"
+    proof (cases "n < m")
+      case True then show ?thesis by simp
+    next
+      case False then have "m <= n" by simp
+      from `0 < m` and False have "n - m < n" by simp
+      with hyp have "gcd (fib m) (fib ((n - m) mod m))
+        = gcd (fib m) (fib (n - m))" by simp
+      also have "... = gcd (fib m) (fib n)"
+        using `m <= n` by (rule gcd_fib_diff)
+      finally have "gcd (fib m) (fib ((n - m) mod m)) =
+        gcd (fib m) (fib n)" .
+      with False show ?thesis by simp
+    qed
+    finally show ?thesis .
+  qed
+qed
+
+
+theorem fib_gcd: "fib (gcd m n) = gcd (fib m) (fib n)" (is "?P m n")
+proof (induct m n rule: gcd_induct)
+  fix m show "fib (gcd m 0) = gcd (fib m) (fib 0)" by simp
+  fix n :: nat assume n: "0 < n"
+  then have "gcd m n = gcd n (m mod n)" by (rule gcd_non_0)
+  also assume hyp: "fib ... = gcd (fib n) (fib (m mod n))"
+  also from n have "... = gcd (fib n) (fib m)" by (rule gcd_fib_mod)
+  also have "... = gcd (fib m) (fib n)" by (rule gcd_commute)
+  finally show "fib (gcd m n) = gcd (fib m) (fib n)" .
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_Examples/Group.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,267 @@
+(*  Title:      HOL/Isar_Examples/Group.thy
+    Author:     Markus Wenzel, TU Muenchen
+*)
+
+header {* Basic group theory *}
+
+theory Group
+imports Main
+begin
+
+subsection {* Groups and calculational reasoning *} 
+
+text {*
+ Groups over signature $({\times} :: \alpha \To \alpha \To \alpha,
+ \idt{one} :: \alpha, \idt{inverse} :: \alpha \To \alpha)$ are defined
+ as an axiomatic type class as follows.  Note that the parent class
+ $\idt{times}$ is provided by the basic HOL theory.
+*}
+
+consts
+  one :: "'a"
+  inverse :: "'a => 'a"
+
+axclass
+  group < times
+  group_assoc:         "(x * y) * z = x * (y * z)"
+  group_left_one:      "one * x = x"
+  group_left_inverse:  "inverse x * x = one"
+
+text {*
+ The group axioms only state the properties of left one and inverse,
+ the right versions may be derived as follows.
+*}
+
+theorem group_right_inverse: "x * inverse x = (one::'a::group)"
+proof -
+  have "x * inverse x = one * (x * inverse x)"
+    by (simp only: group_left_one)
+  also have "... = one * x * inverse x"
+    by (simp only: group_assoc)
+  also have "... = inverse (inverse x) * inverse x * x * inverse x"
+    by (simp only: group_left_inverse)
+  also have "... = inverse (inverse x) * (inverse x * x) * inverse x"
+    by (simp only: group_assoc)
+  also have "... = inverse (inverse x) * one * inverse x"
+    by (simp only: group_left_inverse)
+  also have "... = inverse (inverse x) * (one * inverse x)"
+    by (simp only: group_assoc)
+  also have "... = inverse (inverse x) * inverse x"
+    by (simp only: group_left_one)
+  also have "... = one"
+    by (simp only: group_left_inverse)
+  finally show ?thesis .
+qed
+
+text {*
+ With \name{group-right-inverse} already available,
+ \name{group-right-one}\label{thm:group-right-one} is now established
+ much easier.
+*}
+
+theorem group_right_one: "x * one = (x::'a::group)"
+proof -
+  have "x * one = x * (inverse x * x)"
+    by (simp only: group_left_inverse)
+  also have "... = x * inverse x * x"
+    by (simp only: group_assoc)
+  also have "... = one * x"
+    by (simp only: group_right_inverse)
+  also have "... = x"
+    by (simp only: group_left_one)
+  finally show ?thesis .
+qed
+
+text {*
+ \medskip The calculational proof style above follows typical
+ presentations given in any introductory course on algebra.  The basic
+ technique is to form a transitive chain of equations, which in turn
+ are established by simplifying with appropriate rules.  The low-level
+ logical details of equational reasoning are left implicit.
+
+ Note that ``$\dots$'' is just a special term variable that is bound
+ automatically to the argument\footnote{The argument of a curried
+ infix expression happens to be its right-hand side.} of the last fact
+ achieved by any local assumption or proven statement.  In contrast to
+ $\var{thesis}$, the ``$\dots$'' variable is bound \emph{after} the
+ proof is finished, though.
+
+ There are only two separate Isar language elements for calculational
+ proofs: ``\isakeyword{also}'' for initial or intermediate
+ calculational steps, and ``\isakeyword{finally}'' for exhibiting the
+ result of a calculation.  These constructs are not hardwired into
+ Isabelle/Isar, but defined on top of the basic Isar/VM interpreter.
+ Expanding the \isakeyword{also} and \isakeyword{finally} derived
+ language elements, calculations may be simulated by hand as
+ demonstrated below.
+*}
+
+theorem "x * one = (x::'a::group)"
+proof -
+  have "x * one = x * (inverse x * x)"
+    by (simp only: group_left_inverse)
+
+  note calculation = this
+    -- {* first calculational step: init calculation register *}
+
+  have "... = x * inverse x * x"
+    by (simp only: group_assoc)
+
+  note calculation = trans [OF calculation this]
+    -- {* general calculational step: compose with transitivity rule *}
+
+  have "... = one * x"
+    by (simp only: group_right_inverse)
+
+  note calculation = trans [OF calculation this]
+    -- {* general calculational step: compose with transitivity rule *}
+
+  have "... = x"
+    by (simp only: group_left_one)
+
+  note calculation = trans [OF calculation this]
+    -- {* final calculational step: compose with transitivity rule ... *}
+  from calculation
+    -- {* ... and pick up the final result *}
+
+  show ?thesis .
+qed
+
+text {*
+ Note that this scheme of calculations is not restricted to plain
+ transitivity.  Rules like anti-symmetry, or even forward and backward
+ substitution work as well.  For the actual implementation of
+ \isacommand{also} and \isacommand{finally}, Isabelle/Isar maintains
+ separate context information of ``transitivity'' rules.  Rule
+ selection takes place automatically by higher-order unification.
+*}
+
+
+subsection {* Groups as monoids *}
+
+text {*
+ Monoids over signature $({\times} :: \alpha \To \alpha \To \alpha,
+ \idt{one} :: \alpha)$ are defined like this.
+*}
+
+axclass monoid < times
+  monoid_assoc:       "(x * y) * z = x * (y * z)"
+  monoid_left_one:   "one * x = x"
+  monoid_right_one:  "x * one = x"
+
+text {*
+ Groups are \emph{not} yet monoids directly from the definition.  For
+ monoids, \name{right-one} had to be included as an axiom, but for
+ groups both \name{right-one} and \name{right-inverse} are derivable
+ from the other axioms.  With \name{group-right-one} derived as a
+ theorem of group theory (see page~\pageref{thm:group-right-one}), we
+ may still instantiate $\idt{group} \subseteq \idt{monoid}$ properly
+ as follows.
+*}
+
+instance group < monoid
+  by (intro_classes,
+       rule group_assoc,
+       rule group_left_one,
+       rule group_right_one)
+
+text {*
+ The \isacommand{instance} command actually is a version of
+ \isacommand{theorem}, setting up a goal that reflects the intended
+ class relation (or type constructor arity).  Thus any Isar proof
+ language element may be involved to establish this statement.  When
+ concluding the proof, the result is transformed into the intended
+ type signature extension behind the scenes.
+*}
+
+subsection {* More theorems of group theory *}
+
+text {*
+ The one element is already uniquely determined by preserving an
+ \emph{arbitrary} group element.
+*}
+
+theorem group_one_equality: "e * x = x ==> one = (e::'a::group)"
+proof -
+  assume eq: "e * x = x"
+  have "one = x * inverse x"
+    by (simp only: group_right_inverse)
+  also have "... = (e * x) * inverse x"
+    by (simp only: eq)
+  also have "... = e * (x * inverse x)"
+    by (simp only: group_assoc)
+  also have "... = e * one"
+    by (simp only: group_right_inverse)
+  also have "... = e"
+    by (simp only: group_right_one)
+  finally show ?thesis .
+qed
+
+text {*
+ Likewise, the inverse is already determined by the cancel property.
+*}
+
+theorem group_inverse_equality:
+  "x' * x = one ==> inverse x = (x'::'a::group)"
+proof -
+  assume eq: "x' * x = one"
+  have "inverse x = one * inverse x"
+    by (simp only: group_left_one)
+  also have "... = (x' * x) * inverse x"
+    by (simp only: eq)
+  also have "... = x' * (x * inverse x)"
+    by (simp only: group_assoc)
+  also have "... = x' * one"
+    by (simp only: group_right_inverse)
+  also have "... = x'"
+    by (simp only: group_right_one)
+  finally show ?thesis .
+qed
+
+text {*
+ The inverse operation has some further characteristic properties.
+*}
+
+theorem group_inverse_times:
+  "inverse (x * y) = inverse y * inverse (x::'a::group)"
+proof (rule group_inverse_equality)
+  show "(inverse y * inverse x) * (x * y) = one"
+  proof -
+    have "(inverse y * inverse x) * (x * y) =
+        (inverse y * (inverse x * x)) * y"
+      by (simp only: group_assoc)
+    also have "... = (inverse y * one) * y"
+      by (simp only: group_left_inverse)
+    also have "... = inverse y * y"
+      by (simp only: group_right_one)
+    also have "... = one"
+      by (simp only: group_left_inverse)
+    finally show ?thesis .
+  qed
+qed
+
+theorem inverse_inverse: "inverse (inverse x) = (x::'a::group)"
+proof (rule group_inverse_equality)
+  show "x * inverse x = one"
+    by (simp only: group_right_inverse)
+qed
+
+theorem inverse_inject: "inverse x = inverse y ==> x = (y::'a::group)"
+proof -
+  assume eq: "inverse x = inverse y"
+  have "x = x * one"
+    by (simp only: group_right_one)
+  also have "... = x * (inverse y * y)"
+    by (simp only: group_left_inverse)
+  also have "... = x * (inverse x * y)"
+    by (simp only: eq)
+  also have "... = (x * inverse x) * y"
+    by (simp only: group_assoc)
+  also have "... = one * y"
+    by (simp only: group_right_inverse)
+  also have "... = y"
+    by (simp only: group_left_one)
+  finally show ?thesis .
+qed
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_Examples/Hoare.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,463 @@
+(*  Title:      HOL/Isar_Examples/Hoare.thy
+    Author:     Markus Wenzel, TU Muenchen
+
+A formulation of Hoare logic suitable for Isar.
+*)
+
+header {* Hoare Logic *}
+
+theory Hoare
+imports Main
+uses ("~~/src/HOL/Hoare/hoare_tac.ML")
+begin
+
+subsection {* Abstract syntax and semantics *}
+
+text {*
+ The following abstract syntax and semantics of Hoare Logic over
+ \texttt{WHILE} programs closely follows the existing tradition in
+ Isabelle/HOL of formalizing the presentation given in
+ \cite[\S6]{Winskel:1993}.  See also
+ \url{http://isabelle.in.tum.de/library/Hoare/} and
+ \cite{Nipkow:1998:Winskel}.
+*}
+
+types
+  'a bexp = "'a set"
+  'a assn = "'a set"
+
+datatype 'a com =
+    Basic "'a => 'a"
+  | Seq "'a com" "'a com"    ("(_;/ _)" [60, 61] 60)
+  | Cond "'a bexp" "'a com" "'a com"
+  | While "'a bexp" "'a assn" "'a com"
+
+abbreviation
+  Skip  ("SKIP") where
+  "SKIP == Basic id"
+
+types
+  'a sem = "'a => 'a => bool"
+
+consts
+  iter :: "nat => 'a bexp => 'a sem => 'a sem"
+primrec
+  "iter 0 b S s s' = (s ~: b & s = s')"
+  "iter (Suc n) b S s s' =
+    (s : b & (EX s''. S s s'' & iter n b S s'' s'))"
+
+consts
+  Sem :: "'a com => 'a sem"
+primrec
+  "Sem (Basic f) s s' = (s' = f s)"
+  "Sem (c1; c2) s s' = (EX s''. Sem c1 s s'' & Sem c2 s'' s')"
+  "Sem (Cond b c1 c2) s s' =
+    (if s : b then Sem c1 s s' else Sem c2 s s')"
+  "Sem (While b x c) s s' = (EX n. iter n b (Sem c) s s')"
+
+constdefs
+  Valid :: "'a bexp => 'a com => 'a bexp => bool"
+    ("(3|- _/ (2_)/ _)" [100, 55, 100] 50)
+  "|- P c Q == ALL s s'. Sem c s s' --> s : P --> s' : Q"
+
+syntax (xsymbols)
+  Valid :: "'a bexp => 'a com => 'a bexp => bool"
+    ("(3\<turnstile> _/ (2_)/ _)" [100, 55, 100] 50)
+
+lemma ValidI [intro?]:
+    "(!!s s'. Sem c s s' ==> s : P ==> s' : Q) ==> |- P c Q"
+  by (simp add: Valid_def)
+
+lemma ValidD [dest?]:
+    "|- P c Q ==> Sem c s s' ==> s : P ==> s' : Q"
+  by (simp add: Valid_def)
+
+
+subsection {* Primitive Hoare rules *}
+
+text {*
+ From the semantics defined above, we derive the standard set of
+ primitive Hoare rules; e.g.\ see \cite[\S6]{Winskel:1993}.  Usually,
+ variant forms of these rules are applied in actual proof, see also
+ \S\ref{sec:hoare-isar} and \S\ref{sec:hoare-vcg}.
+
+ \medskip The \name{basic} rule represents any kind of atomic access
+ to the state space.  This subsumes the common rules of \name{skip}
+ and \name{assign}, as formulated in \S\ref{sec:hoare-isar}.
+*}
+
+theorem basic: "|- {s. f s : P} (Basic f) P"
+proof
+  fix s s' assume s: "s : {s. f s : P}"
+  assume "Sem (Basic f) s s'"
+  hence "s' = f s" by simp
+  with s show "s' : P" by simp
+qed
+
+text {*
+ The rules for sequential commands and semantic consequences are
+ established in a straight forward manner as follows.
+*}
+
+theorem seq: "|- P c1 Q ==> |- Q c2 R ==> |- P (c1; c2) R"
+proof
+  assume cmd1: "|- P c1 Q" and cmd2: "|- Q c2 R"
+  fix s s' assume s: "s : P"
+  assume "Sem (c1; c2) s s'"
+  then obtain s'' where sem1: "Sem c1 s s''" and sem2: "Sem c2 s'' s'"
+    by auto
+  from cmd1 sem1 s have "s'' : Q" ..
+  with cmd2 sem2 show "s' : R" ..
+qed
+
+theorem conseq: "P' <= P ==> |- P c Q ==> Q <= Q' ==> |- P' c Q'"
+proof
+  assume P'P: "P' <= P" and QQ': "Q <= Q'"
+  assume cmd: "|- P c Q"
+  fix s s' :: 'a
+  assume sem: "Sem c s s'"
+  assume "s : P'" with P'P have "s : P" ..
+  with cmd sem have "s' : Q" ..
+  with QQ' show "s' : Q'" ..
+qed
+
+text {*
+ The rule for conditional commands is directly reflected by the
+ corresponding semantics; in the proof we just have to look closely
+ which cases apply.
+*}
+
+theorem cond:
+  "|- (P Int b) c1 Q ==> |- (P Int -b) c2 Q ==> |- P (Cond b c1 c2) Q"
+proof
+  assume case_b: "|- (P Int b) c1 Q" and case_nb: "|- (P Int -b) c2 Q"
+  fix s s' assume s: "s : P"
+  assume sem: "Sem (Cond b c1 c2) s s'"
+  show "s' : Q"
+  proof cases
+    assume b: "s : b"
+    from case_b show ?thesis
+    proof
+      from sem b show "Sem c1 s s'" by simp
+      from s b show "s : P Int b" by simp
+    qed
+  next
+    assume nb: "s ~: b"
+    from case_nb show ?thesis
+    proof
+      from sem nb show "Sem c2 s s'" by simp
+      from s nb show "s : P Int -b" by simp
+    qed
+  qed
+qed
+
+text {*
+ The \name{while} rule is slightly less trivial --- it is the only one
+ based on recursion, which is expressed in the semantics by a
+ Kleene-style least fixed-point construction.  The auxiliary statement
+ below, which is by induction on the number of iterations is the main
+ point to be proven; the rest is by routine application of the
+ semantics of \texttt{WHILE}.
+*}
+
+theorem while:
+  assumes body: "|- (P Int b) c P"
+  shows "|- P (While b X c) (P Int -b)"
+proof
+  fix s s' assume s: "s : P"
+  assume "Sem (While b X c) s s'"
+  then obtain n where "iter n b (Sem c) s s'" by auto
+  from this and s show "s' : P Int -b"
+  proof (induct n arbitrary: s)
+    case 0
+    thus ?case by auto
+  next
+    case (Suc n)
+    then obtain s'' where b: "s : b" and sem: "Sem c s s''"
+      and iter: "iter n b (Sem c) s'' s'"
+      by auto
+    from Suc and b have "s : P Int b" by simp
+    with body sem have "s'' : P" ..
+    with iter show ?case by (rule Suc)
+  qed
+qed
+
+
+subsection {* Concrete syntax for assertions *}
+
+text {*
+ We now introduce concrete syntax for describing commands (with
+ embedded expressions) and assertions. The basic technique is that of
+ semantic ``quote-antiquote''.  A \emph{quotation} is a syntactic
+ entity delimited by an implicit abstraction, say over the state
+ space.  An \emph{antiquotation} is a marked expression within a
+ quotation that refers the implicit argument; a typical antiquotation
+ would select (or even update) components from the state.
+
+ We will see some examples later in the concrete rules and
+ applications.
+*}
+
+text {*
+ The following specification of syntax and translations is for
+ Isabelle experts only; feel free to ignore it.
+
+ While the first part is still a somewhat intelligible specification
+ of the concrete syntactic representation of our Hoare language, the
+ actual ``ML drivers'' is quite involved.  Just note that the we
+ re-use the basic quote/antiquote translations as already defined in
+ Isabelle/Pure (see \verb,Syntax.quote_tr, and
+ \verb,Syntax.quote_tr',).
+*}
+
+syntax
+  "_quote"       :: "'b => ('a => 'b)"       ("(.'(_').)" [0] 1000)
+  "_antiquote"   :: "('a => 'b) => 'b"       ("\<acute>_" [1000] 1000)
+  "_Subst"       :: "'a bexp \<Rightarrow> 'b \<Rightarrow> idt \<Rightarrow> 'a bexp"
+        ("_[_'/\<acute>_]" [1000] 999)
+  "_Assert"      :: "'a => 'a set"           ("(.{_}.)" [0] 1000)
+  "_Assign"      :: "idt => 'b => 'a com"    ("(\<acute>_ :=/ _)" [70, 65] 61)
+  "_Cond"        :: "'a bexp => 'a com => 'a com => 'a com"
+        ("(0IF _/ THEN _/ ELSE _/ FI)" [0, 0, 0] 61)
+  "_While_inv"   :: "'a bexp => 'a assn => 'a com => 'a com"
+        ("(0WHILE _/ INV _ //DO _ /OD)"  [0, 0, 0] 61)
+  "_While"       :: "'a bexp => 'a com => 'a com"
+        ("(0WHILE _ //DO _ /OD)"  [0, 0] 61)
+
+syntax (xsymbols)
+  "_Assert"      :: "'a => 'a set"            ("(\<lbrace>_\<rbrace>)" [0] 1000)
+
+translations
+  ".{b}."                   => "Collect .(b)."
+  "B [a/\<acute>x]"                => ".{\<acute>(_update_name x (\<lambda>_. a)) \<in> B}."
+  "\<acute>x := a"                 => "Basic .(\<acute>(_update_name x (\<lambda>_. a)))."
+  "IF b THEN c1 ELSE c2 FI" => "Cond .{b}. c1 c2"
+  "WHILE b INV i DO c OD"   => "While .{b}. i c"
+  "WHILE b DO c OD"         == "WHILE b INV CONST undefined DO c OD"
+
+parse_translation {*
+  let
+    fun quote_tr [t] = Syntax.quote_tr "_antiquote" t
+      | quote_tr ts = raise TERM ("quote_tr", ts);
+  in [("_quote", quote_tr)] end
+*}
+
+text {*
+ As usual in Isabelle syntax translations, the part for printing is
+ more complicated --- we cannot express parts as macro rules as above.
+ Don't look here, unless you have to do similar things for yourself.
+*}
+
+print_translation {*
+  let
+    fun quote_tr' f (t :: ts) =
+          Term.list_comb (f $ Syntax.quote_tr' "_antiquote" t, ts)
+      | quote_tr' _ _ = raise Match;
+
+    val assert_tr' = quote_tr' (Syntax.const "_Assert");
+
+    fun bexp_tr' name ((Const ("Collect", _) $ t) :: ts) =
+          quote_tr' (Syntax.const name) (t :: ts)
+      | bexp_tr' _ _ = raise Match;
+
+    fun upd_tr' (x_upd, T) =
+      (case try (unsuffix Record.updateN) x_upd of
+        SOME x => (x, if T = dummyT then T else Term.domain_type T)
+      | NONE => raise Match);
+
+    fun update_name_tr' (Free x) = Free (upd_tr' x)
+      | update_name_tr' ((c as Const ("_free", _)) $ Free x) =
+          c $ Free (upd_tr' x)
+      | update_name_tr' (Const x) = Const (upd_tr' x)
+      | update_name_tr' _ = raise Match;
+
+    fun K_tr' (Abs (_,_,t)) = if null (loose_bnos t) then t else raise Match
+      | K_tr' (Abs (_,_,Abs (_,_,t)$Bound 0)) = if null (loose_bnos t) then t else raise Match
+      | K_tr' _ = raise Match;
+
+    fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) =
+          quote_tr' (Syntax.const "_Assign" $ update_name_tr' f)
+            (Abs (x, dummyT, K_tr' k) :: ts)
+      | assign_tr' _ = raise Match;
+  in
+    [("Collect", assert_tr'), ("Basic", assign_tr'),
+      ("Cond", bexp_tr' "_Cond"), ("While", bexp_tr' "_While_inv")]
+  end
+*}
+
+
+subsection {* Rules for single-step proof \label{sec:hoare-isar} *}
+
+text {*
+ We are now ready to introduce a set of Hoare rules to be used in
+ single-step structured proofs in Isabelle/Isar.  We refer to the
+ concrete syntax introduce above.
+
+ \medskip Assertions of Hoare Logic may be manipulated in
+ calculational proofs, with the inclusion expressed in terms of sets
+ or predicates.  Reversed order is supported as well.
+*}
+
+lemma [trans]: "|- P c Q ==> P' <= P ==> |- P' c Q"
+  by (unfold Valid_def) blast
+lemma [trans] : "P' <= P ==> |- P c Q ==> |- P' c Q"
+  by (unfold Valid_def) blast
+
+lemma [trans]: "Q <= Q' ==> |- P c Q ==> |- P c Q'"
+  by (unfold Valid_def) blast
+lemma [trans]: "|- P c Q ==> Q <= Q' ==> |- P c Q'"
+  by (unfold Valid_def) blast
+
+lemma [trans]:
+    "|- .{\<acute>P}. c Q ==> (!!s. P' s --> P s) ==> |- .{\<acute>P'}. c Q"
+  by (simp add: Valid_def)
+lemma [trans]:
+    "(!!s. P' s --> P s) ==> |- .{\<acute>P}. c Q ==> |- .{\<acute>P'}. c Q"
+  by (simp add: Valid_def)
+
+lemma [trans]:
+    "|- P c .{\<acute>Q}. ==> (!!s. Q s --> Q' s) ==> |- P c .{\<acute>Q'}."
+  by (simp add: Valid_def)
+lemma [trans]:
+    "(!!s. Q s --> Q' s) ==> |- P c .{\<acute>Q}. ==> |- P c .{\<acute>Q'}."
+  by (simp add: Valid_def)
+
+
+text {*
+ Identity and basic assignments.\footnote{The $\idt{hoare}$ method
+ introduced in \S\ref{sec:hoare-vcg} is able to provide proper
+ instances for any number of basic assignments, without producing
+ additional verification conditions.}
+*}
+
+lemma skip [intro?]: "|- P SKIP P"
+proof -
+  have "|- {s. id s : P} SKIP P" by (rule basic)
+  thus ?thesis by simp
+qed
+
+lemma assign: "|- P [\<acute>a/\<acute>x] \<acute>x := \<acute>a P"
+  by (rule basic)
+
+text {*
+ Note that above formulation of assignment corresponds to our
+ preferred way to model state spaces, using (extensible) record types
+ in HOL \cite{Naraschewski-Wenzel:1998:HOOL}.  For any record field
+ $x$, Isabelle/HOL provides a functions $x$ (selector) and
+ $\idt{x{\dsh}update}$ (update).  Above, there is only a place-holder
+ appearing for the latter kind of function: due to concrete syntax
+ \isa{\'x := \'a} also contains \isa{x\_update}.\footnote{Note that due
+ to the external nature of HOL record fields, we could not even state
+ a general theorem relating selector and update functions (if this
+ were required here); this would only work for any particular instance
+ of record fields introduced so far.}
+*}
+
+text {*
+ Sequential composition --- normalizing with associativity achieves
+ proper of chunks of code verified separately.
+*}
+
+lemmas [trans, intro?] = seq
+
+lemma seq_assoc [simp]: "( |- P c1;(c2;c3) Q) = ( |- P (c1;c2);c3 Q)"
+  by (auto simp add: Valid_def)
+
+text {*
+ Conditional statements.
+*}
+
+lemmas [trans, intro?] = cond
+
+lemma [trans, intro?]:
+  "|- .{\<acute>P & \<acute>b}. c1 Q
+      ==> |- .{\<acute>P & ~ \<acute>b}. c2 Q
+      ==> |- .{\<acute>P}. IF \<acute>b THEN c1 ELSE c2 FI Q"
+    by (rule cond) (simp_all add: Valid_def)
+
+text {*
+ While statements --- with optional invariant.
+*}
+
+lemma [intro?]:
+    "|- (P Int b) c P ==> |- P (While b P c) (P Int -b)"
+  by (rule while)
+
+lemma [intro?]:
+    "|- (P Int b) c P ==> |- P (While b undefined c) (P Int -b)"
+  by (rule while)
+
+
+lemma [intro?]:
+  "|- .{\<acute>P & \<acute>b}. c .{\<acute>P}.
+    ==> |- .{\<acute>P}. WHILE \<acute>b INV .{\<acute>P}. DO c OD .{\<acute>P & ~ \<acute>b}."
+  by (simp add: while Collect_conj_eq Collect_neg_eq)
+
+lemma [intro?]:
+  "|- .{\<acute>P & \<acute>b}. c .{\<acute>P}.
+    ==> |- .{\<acute>P}. WHILE \<acute>b DO c OD .{\<acute>P & ~ \<acute>b}."
+  by (simp add: while Collect_conj_eq Collect_neg_eq)
+
+
+subsection {* Verification conditions \label{sec:hoare-vcg} *}
+
+text {*
+ We now load the \emph{original} ML file for proof scripts and tactic
+ definition for the Hoare Verification Condition Generator (see
+ \url{http://isabelle.in.tum.de/library/Hoare/}).  As far as we are
+ concerned here, the result is a proof method \name{hoare}, which may
+ be applied to a Hoare Logic assertion to extract purely logical
+ verification conditions.  It is important to note that the method
+ requires \texttt{WHILE} loops to be fully annotated with invariants
+ beforehand.  Furthermore, only \emph{concrete} pieces of code are
+ handled --- the underlying tactic fails ungracefully if supplied with
+ meta-variables or parameters, for example.
+*}
+
+lemma SkipRule: "p \<subseteq> q \<Longrightarrow> Valid p (Basic id) q"
+  by (auto simp add: Valid_def)
+
+lemma BasicRule: "p \<subseteq> {s. f s \<in> q} \<Longrightarrow> Valid p (Basic f) q"
+  by (auto simp: Valid_def)
+
+lemma SeqRule: "Valid P c1 Q \<Longrightarrow> Valid Q c2 R \<Longrightarrow> Valid P (c1;c2) R"
+  by (auto simp: Valid_def)
+
+lemma CondRule:
+  "p \<subseteq> {s. (s \<in> b \<longrightarrow> s \<in> w) \<and> (s \<notin> b \<longrightarrow> s \<in> w')}
+    \<Longrightarrow> Valid w c1 q \<Longrightarrow> Valid w' c2 q \<Longrightarrow> Valid p (Cond b c1 c2) q"
+  by (auto simp: Valid_def)
+
+lemma iter_aux:
+  "\<forall>s s'. Sem c s s' --> s : I & s : b --> s' : I ==>
+       (\<And>s s'. s : I \<Longrightarrow> iter n b (Sem c) s s' \<Longrightarrow> s' : I & s' ~: b)"
+  apply(induct n)
+   apply clarsimp
+   apply (simp (no_asm_use))
+   apply blast
+  done
+
+lemma WhileRule:
+    "p \<subseteq> i \<Longrightarrow> Valid (i \<inter> b) c i \<Longrightarrow> i \<inter> (-b) \<subseteq> q \<Longrightarrow> Valid p (While b i c) q"
+  apply (clarsimp simp: Valid_def)
+  apply (drule iter_aux)
+    prefer 2
+    apply assumption
+   apply blast
+  apply blast
+  done
+
+lemma Compl_Collect: "- Collect b = {x. \<not> b x}"
+  by blast
+
+lemmas AbortRule = SkipRule  -- "dummy version"
+
+use "~~/src/HOL/Hoare/hoare_tac.ML"
+
+method_setup hoare = {*
+  Scan.succeed (fn ctxt =>
+    (SIMPLE_METHOD'
+       (hoare_tac ctxt (simp_tac (HOL_basic_ss addsimps [@{thm "Record.K_record_comp"}] ))))) *}
+  "verification condition generator for Hoare logic"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_Examples/Hoare_Ex.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,329 @@
+header {* Using Hoare Logic *}
+
+theory Hoare_Ex
+imports Hoare
+begin
+
+subsection {* State spaces *}
+
+text {*
+ First of all we provide a store of program variables that
+ occur in any of the programs considered later.  Slightly unexpected
+ things may happen when attempting to work with undeclared variables.
+*}
+
+record vars =
+  I :: nat
+  M :: nat
+  N :: nat
+  S :: nat
+
+text {*
+ While all of our variables happen to have the same type, nothing
+ would prevent us from working with many-sorted programs as well, or
+ even polymorphic ones.  Also note that Isabelle/HOL's extensible
+ record types even provides simple means to extend the state space
+ later.
+*}
+
+
+subsection {* Basic examples *}
+
+text {*
+ We look at few trivialities involving assignment and sequential
+ composition, in order to get an idea of how to work with our
+ formulation of Hoare Logic.
+*}
+
+text {*
+ Using the basic \name{assign} rule directly is a bit cumbersome.
+*}
+
+lemma
+  "|- .{\<acute>(N_update (\<lambda>_. (2 * \<acute>N))) : .{\<acute>N = 10}.}. \<acute>N := 2 * \<acute>N .{\<acute>N = 10}."
+  by (rule assign)
+
+text {*
+ Certainly we want the state modification already done, e.g.\ by
+ simplification.  The \name{hoare} method performs the basic state
+ update for us; we may apply the Simplifier afterwards to achieve
+ ``obvious'' consequences as well.
+*}
+
+lemma "|- .{True}. \<acute>N := 10 .{\<acute>N = 10}."
+  by hoare
+
+lemma "|- .{2 * \<acute>N = 10}. \<acute>N := 2 * \<acute>N .{\<acute>N = 10}."
+  by hoare
+
+lemma "|- .{\<acute>N = 5}. \<acute>N := 2 * \<acute>N .{\<acute>N = 10}."
+  by hoare simp
+
+lemma "|- .{\<acute>N + 1 = a + 1}. \<acute>N := \<acute>N + 1 .{\<acute>N = a + 1}."
+  by hoare
+
+lemma "|- .{\<acute>N = a}. \<acute>N := \<acute>N + 1 .{\<acute>N = a + 1}."
+  by hoare simp
+
+lemma "|- .{a = a & b = b}. \<acute>M := a; \<acute>N := b .{\<acute>M = a & \<acute>N = b}."
+  by hoare
+
+lemma "|- .{True}. \<acute>M := a; \<acute>N := b .{\<acute>M = a & \<acute>N = b}."
+  by hoare simp
+
+lemma
+"|- .{\<acute>M = a & \<acute>N = b}.
+    \<acute>I := \<acute>M; \<acute>M := \<acute>N; \<acute>N := \<acute>I
+    .{\<acute>M = b & \<acute>N = a}."
+  by hoare simp
+
+text {*
+ It is important to note that statements like the following one can
+ only be proven for each individual program variable.  Due to the
+ extra-logical nature of record fields, we cannot formulate a theorem
+ relating record selectors and updates schematically.
+*}
+
+lemma "|- .{\<acute>N = a}. \<acute>N := \<acute>N .{\<acute>N = a}."
+  by hoare
+
+lemma "|- .{\<acute>x = a}. \<acute>x := \<acute>x .{\<acute>x = a}."
+  oops
+
+lemma
+  "Valid {s. x s = a} (Basic (\<lambda>s. x_update (x s) s)) {s. x s = n}"
+  -- {* same statement without concrete syntax *}
+  oops
+
+
+text {*
+ In the following assignments we make use of the consequence rule in
+ order to achieve the intended precondition.  Certainly, the
+ \name{hoare} method is able to handle this case, too.
+*}
+
+lemma "|- .{\<acute>M = \<acute>N}. \<acute>M := \<acute>M + 1 .{\<acute>M ~= \<acute>N}."
+proof -
+  have ".{\<acute>M = \<acute>N}. <= .{\<acute>M + 1 ~= \<acute>N}."
+    by auto
+  also have "|- ... \<acute>M := \<acute>M + 1 .{\<acute>M ~= \<acute>N}."
+    by hoare
+  finally show ?thesis .
+qed
+
+lemma "|- .{\<acute>M = \<acute>N}. \<acute>M := \<acute>M + 1 .{\<acute>M ~= \<acute>N}."
+proof -
+  have "!!m n::nat. m = n --> m + 1 ~= n"
+      -- {* inclusion of assertions expressed in ``pure'' logic, *}
+      -- {* without mentioning the state space *}
+    by simp
+  also have "|- .{\<acute>M + 1 ~= \<acute>N}. \<acute>M := \<acute>M + 1 .{\<acute>M ~= \<acute>N}."
+    by hoare
+  finally show ?thesis .
+qed
+
+lemma "|- .{\<acute>M = \<acute>N}. \<acute>M := \<acute>M + 1 .{\<acute>M ~= \<acute>N}."
+  by hoare simp
+
+
+subsection {* Multiplication by addition *}
+
+text {*
+ We now do some basic examples of actual \texttt{WHILE} programs.
+ This one is a loop for calculating the product of two natural
+ numbers, by iterated addition.  We first give detailed structured
+ proof based on single-step Hoare rules.
+*}
+
+lemma
+  "|- .{\<acute>M = 0 & \<acute>S = 0}.
+      WHILE \<acute>M ~= a
+      DO \<acute>S := \<acute>S + b; \<acute>M := \<acute>M + 1 OD
+      .{\<acute>S = a * b}."
+proof -
+  let "|- _ ?while _" = ?thesis
+  let ".{\<acute>?inv}." = ".{\<acute>S = \<acute>M * b}."
+
+  have ".{\<acute>M = 0 & \<acute>S = 0}. <= .{\<acute>?inv}." by auto
+  also have "|- ... ?while .{\<acute>?inv & ~ (\<acute>M ~= a)}."
+  proof
+    let ?c = "\<acute>S := \<acute>S + b; \<acute>M := \<acute>M + 1"
+    have ".{\<acute>?inv & \<acute>M ~= a}. <= .{\<acute>S + b = (\<acute>M + 1) * b}."
+      by auto
+    also have "|- ... ?c .{\<acute>?inv}." by hoare
+    finally show "|- .{\<acute>?inv & \<acute>M ~= a}. ?c .{\<acute>?inv}." .
+  qed
+  also have "... <= .{\<acute>S = a * b}." by auto
+  finally show ?thesis .
+qed
+
+text {*
+ The subsequent version of the proof applies the \name{hoare} method
+ to reduce the Hoare statement to a purely logical problem that can be
+ solved fully automatically.  Note that we have to specify the
+ \texttt{WHILE} loop invariant in the original statement.
+*}
+
+lemma
+  "|- .{\<acute>M = 0 & \<acute>S = 0}.
+      WHILE \<acute>M ~= a
+      INV .{\<acute>S = \<acute>M * b}.
+      DO \<acute>S := \<acute>S + b; \<acute>M := \<acute>M + 1 OD
+      .{\<acute>S = a * b}."
+  by hoare auto
+
+
+subsection {* Summing natural numbers *}
+
+text {*
+ We verify an imperative program to sum natural numbers up to a given
+ limit.  First some functional definition for proper specification of
+ the problem.
+*}
+
+text {*
+ The following proof is quite explicit in the individual steps taken,
+ with the \name{hoare} method only applied locally to take care of
+ assignment and sequential composition.  Note that we express
+ intermediate proof obligation in pure logic, without referring to the
+ state space.
+*}
+
+declare atLeast0LessThan[symmetric,simp]
+
+theorem
+  "|- .{True}.
+      \<acute>S := 0; \<acute>I := 1;
+      WHILE \<acute>I ~= n
+      DO
+        \<acute>S := \<acute>S + \<acute>I;
+        \<acute>I := \<acute>I + 1
+      OD
+      .{\<acute>S = (SUM j<n. j)}."
+  (is "|- _ (_; ?while) _")
+proof -
+  let ?sum = "\<lambda>k::nat. SUM j<k. j"
+  let ?inv = "\<lambda>s i::nat. s = ?sum i"
+
+  have "|- .{True}. \<acute>S := 0; \<acute>I := 1 .{?inv \<acute>S \<acute>I}."
+  proof -
+    have "True --> 0 = ?sum 1"
+      by simp
+    also have "|- .{...}. \<acute>S := 0; \<acute>I := 1 .{?inv \<acute>S \<acute>I}."
+      by hoare
+    finally show ?thesis .
+  qed
+  also have "|- ... ?while .{?inv \<acute>S \<acute>I & ~ \<acute>I ~= n}."
+  proof
+    let ?body = "\<acute>S := \<acute>S + \<acute>I; \<acute>I := \<acute>I + 1"
+    have "!!s i. ?inv s i & i ~= n -->  ?inv (s + i) (i + 1)"
+      by simp
+    also have "|- .{\<acute>S + \<acute>I = ?sum (\<acute>I + 1)}. ?body .{?inv \<acute>S \<acute>I}."
+      by hoare
+    finally show "|- .{?inv \<acute>S \<acute>I & \<acute>I ~= n}. ?body .{?inv \<acute>S \<acute>I}." .
+  qed
+  also have "!!s i. s = ?sum i & ~ i ~= n --> s = ?sum n"
+    by simp
+  finally show ?thesis .
+qed
+
+text {*
+ The next version uses the \name{hoare} method, while still explaining
+ the resulting proof obligations in an abstract, structured manner.
+*}
+
+theorem
+  "|- .{True}.
+      \<acute>S := 0; \<acute>I := 1;
+      WHILE \<acute>I ~= n
+      INV .{\<acute>S = (SUM j<\<acute>I. j)}.
+      DO
+        \<acute>S := \<acute>S + \<acute>I;
+        \<acute>I := \<acute>I + 1
+      OD
+      .{\<acute>S = (SUM j<n. j)}."
+proof -
+  let ?sum = "\<lambda>k::nat. SUM j<k. j"
+  let ?inv = "\<lambda>s i::nat. s = ?sum i"
+
+  show ?thesis
+  proof hoare
+    show "?inv 0 1" by simp
+  next
+    fix s i assume "?inv s i & i ~= n"
+    thus "?inv (s + i) (i + 1)" by simp
+  next
+    fix s i assume "?inv s i & ~ i ~= n"
+    thus "s = ?sum n" by simp
+  qed
+qed
+
+text {*
+ Certainly, this proof may be done fully automatic as well, provided
+ that the invariant is given beforehand.
+*}
+
+theorem
+  "|- .{True}.
+      \<acute>S := 0; \<acute>I := 1;
+      WHILE \<acute>I ~= n
+      INV .{\<acute>S = (SUM j<\<acute>I. j)}.
+      DO
+        \<acute>S := \<acute>S + \<acute>I;
+        \<acute>I := \<acute>I + 1
+      OD
+      .{\<acute>S = (SUM j<n. j)}."
+  by hoare auto
+
+
+subsection{* Time *}
+
+text{*
+  A simple embedding of time in Hoare logic: function @{text timeit}
+  inserts an extra variable to keep track of the elapsed time.
+*}
+
+record tstate = time :: nat
+
+types 'a time = "\<lparr>time :: nat, \<dots> :: 'a\<rparr>"
+
+consts timeit :: "'a time com \<Rightarrow> 'a time com"
+primrec
+  "timeit (Basic f) = (Basic f; Basic(\<lambda>s. s\<lparr>time := Suc (time s)\<rparr>))"
+  "timeit (c1; c2) = (timeit c1; timeit c2)"
+  "timeit (Cond b c1 c2) = Cond b (timeit c1) (timeit c2)"
+  "timeit (While b iv c) = While b iv (timeit c)"
+
+record tvars = tstate +
+  I :: nat
+  J :: nat
+
+lemma lem: "(0::nat) < n \<Longrightarrow> n + n \<le> Suc (n * n)"
+  by (induct n) simp_all
+
+lemma "|- .{i = \<acute>I & \<acute>time = 0}.
+ timeit(
+ WHILE \<acute>I \<noteq> 0
+ INV .{2*\<acute>time + \<acute>I*\<acute>I + 5*\<acute>I = i*i + 5*i}.
+ DO
+   \<acute>J := \<acute>I;
+   WHILE \<acute>J \<noteq> 0
+   INV .{0 < \<acute>I & 2*\<acute>time + \<acute>I*\<acute>I + 3*\<acute>I + 2*\<acute>J - 2 = i*i + 5*i}.
+   DO \<acute>J := \<acute>J - 1 OD;
+   \<acute>I := \<acute>I - 1
+ OD
+ ) .{2*\<acute>time = i*i + 5*i}."
+  apply simp
+  apply hoare
+      apply simp
+     apply clarsimp
+    apply clarsimp
+   apply arith
+   prefer 2
+   apply clarsimp
+  apply (clarsimp simp: nat_distrib)
+  apply (frule lem)
+  apply arith
+  done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_Examples/Knaster_Tarski.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,111 @@
+(*  Title:      HOL/Isar_Examples/Knaster_Tarski.thy
+    Author:     Markus Wenzel, TU Muenchen
+
+Typical textbook proof example.
+*)
+
+header {* Textbook-style reasoning: the Knaster-Tarski Theorem *}
+
+theory Knaster_Tarski
+imports Main Lattice_Syntax
+begin
+
+
+subsection {* Prose version *}
+
+text {*
+  According to the textbook \cite[pages 93--94]{davey-priestley}, the
+  Knaster-Tarski fixpoint theorem is as follows.\footnote{We have
+  dualized the argument, and tuned the notation a little bit.}
+
+  \textbf{The Knaster-Tarski Fixpoint Theorem.}  Let @{text L} be a
+  complete lattice and @{text "f: L \<rightarrow> L"} an order-preserving map.
+  Then @{text "\<Sqinter>{x \<in> L | f(x) \<le> x}"} is a fixpoint of @{text f}.
+
+  \textbf{Proof.} Let @{text "H = {x \<in> L | f(x) \<le> x}"} and @{text "a =
+  \<Sqinter>H"}.  For all @{text "x \<in> H"} we have @{text "a \<le> x"}, so @{text
+  "f(a) \<le> f(x) \<le> x"}.  Thus @{text "f(a)"} is a lower bound of @{text
+  H}, whence @{text "f(a) \<le> a"}.  We now use this inequality to prove
+  the reverse one (!) and thereby complete the proof that @{text a} is
+  a fixpoint.  Since @{text f} is order-preserving, @{text "f(f(a)) \<le>
+  f(a)"}.  This says @{text "f(a) \<in> H"}, so @{text "a \<le> f(a)"}.
+*}
+
+
+subsection {* Formal versions *}
+
+text {*
+  The Isar proof below closely follows the original presentation.
+  Virtually all of the prose narration has been rephrased in terms of
+  formal Isar language elements.  Just as many textbook-style proofs,
+  there is a strong bias towards forward proof, and several bends in
+  the course of reasoning.
+*}
+
+theorem Knaster_Tarski:
+  fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
+  assumes "mono f"
+  shows "\<exists>a. f a = a"
+proof
+  let ?H = "{u. f u \<le> u}"
+  let ?a = "\<Sqinter>?H"
+  show "f ?a = ?a"
+  proof -
+    {
+      fix x
+      assume "x \<in> ?H"
+      then have "?a \<le> x" by (rule Inf_lower)
+      with `mono f` have "f ?a \<le> f x" ..
+      also from `x \<in> ?H` have "\<dots> \<le> x" ..
+      finally have "f ?a \<le> x" .
+    }
+    then have "f ?a \<le> ?a" by (rule Inf_greatest)
+    {
+      also presume "\<dots> \<le> f ?a"
+      finally (order_antisym) show ?thesis .
+    }
+    from `mono f` and `f ?a \<le> ?a` have "f (f ?a) \<le> f ?a" ..
+    then have "f ?a \<in> ?H" ..
+    then show "?a \<le> f ?a" by (rule Inf_lower)
+  qed
+qed
+
+text {*
+  Above we have used several advanced Isar language elements, such as
+  explicit block structure and weak assumptions.  Thus we have
+  mimicked the particular way of reasoning of the original text.
+
+  In the subsequent version the order of reasoning is changed to
+  achieve structured top-down decomposition of the problem at the
+  outer level, while only the inner steps of reasoning are done in a
+  forward manner.  We are certainly more at ease here, requiring only
+  the most basic features of the Isar language.
+*}
+
+theorem Knaster_Tarski':
+  fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
+  assumes "mono f"
+  shows "\<exists>a. f a = a"
+proof
+  let ?H = "{u. f u \<le> u}"
+  let ?a = "\<Sqinter>?H"
+  show "f ?a = ?a"
+  proof (rule order_antisym)
+    show "f ?a \<le> ?a"
+    proof (rule Inf_greatest)
+      fix x
+      assume "x \<in> ?H"
+      then have "?a \<le> x" by (rule Inf_lower)
+      with `mono f` have "f ?a \<le> f x" ..
+      also from `x \<in> ?H` have "\<dots> \<le> x" ..
+      finally show "f ?a \<le> x" .
+    qed
+    show "?a \<le> f ?a"
+    proof (rule Inf_lower)
+      from `mono f` and `f ?a \<le> ?a` have "f (f ?a) \<le> f ?a" ..
+      then show "f ?a \<in> ?H" ..
+    qed
+  qed
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_Examples/Mutilated_Checkerboard.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,300 @@
+(*  Title:      HOL/Isar_Examples/Mutilated_Checkerboard.thy
+    Author:     Markus Wenzel, TU Muenchen (Isar document)
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory (original scripts)
+*)
+
+header {* The Mutilated Checker Board Problem *}
+
+theory Mutilated_Checkerboard
+imports Main
+begin
+
+text {*
+ The Mutilated Checker Board Problem, formalized inductively.  See
+ \cite{paulson-mutilated-board} and
+ \url{http://isabelle.in.tum.de/library/HOL/Induct/Mutil.html} for the
+ original tactic script version.
+*}
+
+subsection {* Tilings *}
+
+inductive_set
+  tiling :: "'a set set => 'a set set"
+  for A :: "'a set set"
+  where
+    empty: "{} : tiling A"
+  | Un: "a : A ==> t : tiling A ==> a <= - t ==> a Un t : tiling A"
+
+
+text "The union of two disjoint tilings is a tiling."
+
+lemma tiling_Un:
+  assumes "t : tiling A" and "u : tiling A" and "t Int u = {}"
+  shows "t Un u : tiling A"
+proof -
+  let ?T = "tiling A"
+  from `t : ?T` and `t Int u = {}`
+  show "t Un u : ?T"
+  proof (induct t)
+    case empty
+    with `u : ?T` show "{} Un u : ?T" by simp
+  next
+    case (Un a t)
+    show "(a Un t) Un u : ?T"
+    proof -
+      have "a Un (t Un u) : ?T"
+        using `a : A`
+      proof (rule tiling.Un)
+        from `(a Un t) Int u = {}` have "t Int u = {}" by blast
+        then show "t Un u: ?T" by (rule Un)
+        from `a <= - t` and `(a Un t) Int u = {}`
+        show "a <= - (t Un u)" by blast
+      qed
+      also have "a Un (t Un u) = (a Un t) Un u"
+        by (simp only: Un_assoc)
+      finally show ?thesis .
+    qed
+  qed
+qed
+
+
+subsection {* Basic properties of ``below'' *}
+
+constdefs
+  below :: "nat => nat set"
+  "below n == {i. i < n}"
+
+lemma below_less_iff [iff]: "(i: below k) = (i < k)"
+  by (simp add: below_def)
+
+lemma below_0: "below 0 = {}"
+  by (simp add: below_def)
+
+lemma Sigma_Suc1:
+    "m = n + 1 ==> below m <*> B = ({n} <*> B) Un (below n <*> B)"
+  by (simp add: below_def less_Suc_eq) blast
+
+lemma Sigma_Suc2:
+    "m = n + 2 ==> A <*> below m =
+      (A <*> {n}) Un (A <*> {n + 1}) Un (A <*> below n)"
+  by (auto simp add: below_def)
+
+lemmas Sigma_Suc = Sigma_Suc1 Sigma_Suc2
+
+
+subsection {* Basic properties of ``evnodd'' *}
+
+constdefs
+  evnodd :: "(nat * nat) set => nat => (nat * nat) set"
+  "evnodd A b == A Int {(i, j). (i + j) mod 2 = b}"
+
+lemma evnodd_iff:
+    "(i, j): evnodd A b = ((i, j): A  & (i + j) mod 2 = b)"
+  by (simp add: evnodd_def)
+
+lemma evnodd_subset: "evnodd A b <= A"
+  by (unfold evnodd_def, rule Int_lower1)
+
+lemma evnoddD: "x : evnodd A b ==> x : A"
+  by (rule subsetD, rule evnodd_subset)
+
+lemma evnodd_finite: "finite A ==> finite (evnodd A b)"
+  by (rule finite_subset, rule evnodd_subset)
+
+lemma evnodd_Un: "evnodd (A Un B) b = evnodd A b Un evnodd B b"
+  by (unfold evnodd_def) blast
+
+lemma evnodd_Diff: "evnodd (A - B) b = evnodd A b - evnodd B b"
+  by (unfold evnodd_def) blast
+
+lemma evnodd_empty: "evnodd {} b = {}"
+  by (simp add: evnodd_def)
+
+lemma evnodd_insert: "evnodd (insert (i, j) C) b =
+    (if (i + j) mod 2 = b
+      then insert (i, j) (evnodd C b) else evnodd C b)"
+  by (simp add: evnodd_def)
+
+
+subsection {* Dominoes *}
+
+inductive_set
+  domino :: "(nat * nat) set set"
+  where
+    horiz: "{(i, j), (i, j + 1)} : domino"
+  | vertl: "{(i, j), (i + 1, j)} : domino"
+
+lemma dominoes_tile_row:
+  "{i} <*> below (2 * n) : tiling domino"
+  (is "?B n : ?T")
+proof (induct n)
+  case 0
+  show ?case by (simp add: below_0 tiling.empty)
+next
+  case (Suc n)
+  let ?a = "{i} <*> {2 * n + 1} Un {i} <*> {2 * n}"
+  have "?B (Suc n) = ?a Un ?B n"
+    by (auto simp add: Sigma_Suc Un_assoc)
+  moreover have "... : ?T"
+  proof (rule tiling.Un)
+    have "{(i, 2 * n), (i, 2 * n + 1)} : domino"
+      by (rule domino.horiz)
+    also have "{(i, 2 * n), (i, 2 * n + 1)} = ?a" by blast
+    finally show "... : domino" .
+    show "?B n : ?T" by (rule Suc)
+    show "?a <= - ?B n" by blast
+  qed
+  ultimately show ?case by simp
+qed
+
+lemma dominoes_tile_matrix:
+  "below m <*> below (2 * n) : tiling domino"
+  (is "?B m : ?T")
+proof (induct m)
+  case 0
+  show ?case by (simp add: below_0 tiling.empty)
+next
+  case (Suc m)
+  let ?t = "{m} <*> below (2 * n)"
+  have "?B (Suc m) = ?t Un ?B m" by (simp add: Sigma_Suc)
+  moreover have "... : ?T"
+  proof (rule tiling_Un)
+    show "?t : ?T" by (rule dominoes_tile_row)
+    show "?B m : ?T" by (rule Suc)
+    show "?t Int ?B m = {}" by blast
+  qed
+  ultimately show ?case by simp
+qed
+
+lemma domino_singleton:
+  assumes d: "d : domino" and "b < 2"
+  shows "EX i j. evnodd d b = {(i, j)}"  (is "?P d")
+  using d
+proof induct
+  from `b < 2` have b_cases: "b = 0 | b = 1" by arith
+  fix i j
+  note [simp] = evnodd_empty evnodd_insert mod_Suc
+  from b_cases show "?P {(i, j), (i, j + 1)}" by rule auto
+  from b_cases show "?P {(i, j), (i + 1, j)}" by rule auto
+qed
+
+lemma domino_finite:
+  assumes d: "d: domino"
+  shows "finite d"
+  using d
+proof induct
+  fix i j :: nat
+  show "finite {(i, j), (i, j + 1)}" by (intro finite.intros)
+  show "finite {(i, j), (i + 1, j)}" by (intro finite.intros)
+qed
+
+
+subsection {* Tilings of dominoes *}
+
+lemma tiling_domino_finite:
+  assumes t: "t : tiling domino"  (is "t : ?T")
+  shows "finite t"  (is "?F t")
+  using t
+proof induct
+  show "?F {}" by (rule finite.emptyI)
+  fix a t assume "?F t"
+  assume "a : domino" then have "?F a" by (rule domino_finite)
+  from this and `?F t` show "?F (a Un t)" by (rule finite_UnI)
+qed
+
+lemma tiling_domino_01:
+  assumes t: "t : tiling domino"  (is "t : ?T")
+  shows "card (evnodd t 0) = card (evnodd t 1)"
+  using t
+proof induct
+  case empty
+  show ?case by (simp add: evnodd_def)
+next
+  case (Un a t)
+  let ?e = evnodd
+  note hyp = `card (?e t 0) = card (?e t 1)`
+    and at = `a <= - t`
+  have card_suc:
+    "!!b. b < 2 ==> card (?e (a Un t) b) = Suc (card (?e t b))"
+  proof -
+    fix b :: nat assume "b < 2"
+    have "?e (a Un t) b = ?e a b Un ?e t b" by (rule evnodd_Un)
+    also obtain i j where e: "?e a b = {(i, j)}"
+    proof -
+      from `a \<in> domino` and `b < 2`
+      have "EX i j. ?e a b = {(i, j)}" by (rule domino_singleton)
+      then show ?thesis by (blast intro: that)
+    qed
+    moreover have "... Un ?e t b = insert (i, j) (?e t b)" by simp
+    moreover have "card ... = Suc (card (?e t b))"
+    proof (rule card_insert_disjoint)
+      from `t \<in> tiling domino` have "finite t"
+        by (rule tiling_domino_finite)
+      then show "finite (?e t b)"
+        by (rule evnodd_finite)
+      from e have "(i, j) : ?e a b" by simp
+      with at show "(i, j) ~: ?e t b" by (blast dest: evnoddD)
+    qed
+    ultimately show "?thesis b" by simp
+  qed
+  then have "card (?e (a Un t) 0) = Suc (card (?e t 0))" by simp
+  also from hyp have "card (?e t 0) = card (?e t 1)" .
+  also from card_suc have "Suc ... = card (?e (a Un t) 1)"
+    by simp
+  finally show ?case .
+qed
+
+
+subsection {* Main theorem *}
+
+constdefs
+  mutilated_board :: "nat => nat => (nat * nat) set"
+  "mutilated_board m n ==
+    below (2 * (m + 1)) <*> below (2 * (n + 1))
+      - {(0, 0)} - {(2 * m + 1, 2 * n + 1)}"
+
+theorem mutil_not_tiling: "mutilated_board m n ~: tiling domino"
+proof (unfold mutilated_board_def)
+  let ?T = "tiling domino"
+  let ?t = "below (2 * (m + 1)) <*> below (2 * (n + 1))"
+  let ?t' = "?t - {(0, 0)}"
+  let ?t'' = "?t' - {(2 * m + 1, 2 * n + 1)}"
+
+  show "?t'' ~: ?T"
+  proof
+    have t: "?t : ?T" by (rule dominoes_tile_matrix)
+    assume t'': "?t'' : ?T"
+
+    let ?e = evnodd
+    have fin: "finite (?e ?t 0)"
+      by (rule evnodd_finite, rule tiling_domino_finite, rule t)
+
+    note [simp] = evnodd_iff evnodd_empty evnodd_insert evnodd_Diff
+    have "card (?e ?t'' 0) < card (?e ?t' 0)"
+    proof -
+      have "card (?e ?t' 0 - {(2 * m + 1, 2 * n + 1)})
+        < card (?e ?t' 0)"
+      proof (rule card_Diff1_less)
+        from _ fin show "finite (?e ?t' 0)"
+          by (rule finite_subset) auto
+        show "(2 * m + 1, 2 * n + 1) : ?e ?t' 0" by simp
+      qed
+      then show ?thesis by simp
+    qed
+    also have "... < card (?e ?t 0)"
+    proof -
+      have "(0, 0) : ?e ?t 0" by simp
+      with fin have "card (?e ?t 0 - {(0, 0)}) < card (?e ?t 0)"
+        by (rule card_Diff1_less)
+      then show ?thesis by simp
+    qed
+    also from t have "... = card (?e ?t 1)"
+      by (rule tiling_domino_01)
+    also have "?e ?t 1 = ?e ?t'' 1" by simp
+    also from t'' have "card ... = card (?e ?t'' 0)"
+      by (rule tiling_domino_01 [symmetric])
+    finally have "... < ..." . then show False ..
+  qed
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_Examples/Nested_Datatype.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,86 @@
+header {* Nested datatypes *}
+
+theory Nested_Datatype
+imports Main
+begin
+
+subsection {* Terms and substitution *}
+
+datatype ('a, 'b) "term" =
+    Var 'a
+  | App 'b "('a, 'b) term list"
+
+consts
+  subst_term :: "('a => ('a, 'b) term) => ('a, 'b) term => ('a, 'b) term"
+  subst_term_list ::
+    "('a => ('a, 'b) term) => ('a, 'b) term list => ('a, 'b) term list"
+
+primrec (subst)
+  "subst_term f (Var a) = f a"
+  "subst_term f (App b ts) = App b (subst_term_list f ts)"
+  "subst_term_list f [] = []"
+  "subst_term_list f (t # ts) = subst_term f t # subst_term_list f ts"
+
+
+text {*
+ \medskip A simple lemma about composition of substitutions.
+*}
+
+lemma "subst_term (subst_term f1 o f2) t =
+      subst_term f1 (subst_term f2 t)"
+  and "subst_term_list (subst_term f1 o f2) ts =
+      subst_term_list f1 (subst_term_list f2 ts)"
+  by (induct t and ts) simp_all
+
+lemma "subst_term (subst_term f1 o f2) t =
+  subst_term f1 (subst_term f2 t)"
+proof -
+  let "?P t" = ?thesis
+  let ?Q = "\<lambda>ts. subst_term_list (subst_term f1 o f2) ts =
+    subst_term_list f1 (subst_term_list f2 ts)"
+  show ?thesis
+  proof (induct t)
+    fix a show "?P (Var a)" by simp
+  next
+    fix b ts assume "?Q ts"
+    then show "?P (App b ts)"
+      by (simp only: subst.simps)
+  next
+    show "?Q []" by simp
+  next
+    fix t ts
+    assume "?P t" "?Q ts" then show "?Q (t # ts)"
+      by (simp only: subst.simps)
+  qed
+qed
+
+
+subsection {* Alternative induction *}
+
+theorem term_induct' [case_names Var App]:
+  assumes var: "!!a. P (Var a)"
+    and app: "!!b ts. list_all P ts ==> P (App b ts)"
+  shows "P t"
+proof (induct t)
+  fix a show "P (Var a)" by (rule var)
+next
+  fix b t ts assume "list_all P ts"
+  then show "P (App b ts)" by (rule app)
+next
+  show "list_all P []" by simp
+next
+  fix t ts assume "P t" "list_all P ts"
+  then show "list_all P (t # ts)" by simp
+qed
+
+lemma
+  "subst_term (subst_term f1 o f2) t = subst_term f1 (subst_term f2 t)"
+proof (induct t rule: term_induct')
+  case (Var a)
+  show ?case by (simp add: o_def)
+next
+  case (App b ts)
+  then show ?case by (induct ts) simp_all
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_Examples/Peirce.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,90 @@
+(*  Title:      HOL/Isar_Examples/Peirce.thy
+    Author:     Markus Wenzel, TU Muenchen
+*)
+
+header {* Peirce's Law *}
+
+theory Peirce
+imports Main
+begin
+
+text {*
+ We consider Peirce's Law: $((A \impl B) \impl A) \impl A$.  This is
+ an inherently non-intuitionistic statement, so its proof will
+ certainly involve some form of classical contradiction.
+
+ The first proof is again a well-balanced combination of plain
+ backward and forward reasoning.  The actual classical step is where
+ the negated goal may be introduced as additional assumption.  This
+ eventually leads to a contradiction.\footnote{The rule involved there
+ is negation elimination; it holds in intuitionistic logic as well.}
+*}
+
+theorem "((A --> B) --> A) --> A"
+proof
+  assume "(A --> B) --> A"
+  show A
+  proof (rule classical)
+    assume "~ A"
+    have "A --> B"
+    proof
+      assume A
+      with `~ A` show B by contradiction
+    qed
+    with `(A --> B) --> A` show A ..
+  qed
+qed
+
+text {*
+ In the subsequent version the reasoning is rearranged by means of
+ ``weak assumptions'' (as introduced by \isacommand{presume}).  Before
+ assuming the negated goal $\neg A$, its intended consequence $A \impl
+ B$ is put into place in order to solve the main problem.
+ Nevertheless, we do not get anything for free, but have to establish
+ $A \impl B$ later on.  The overall effect is that of a logical
+ \emph{cut}.
+
+ Technically speaking, whenever some goal is solved by
+ \isacommand{show} in the context of weak assumptions then the latter
+ give rise to new subgoals, which may be established separately.  In
+ contrast, strong assumptions (as introduced by \isacommand{assume})
+ are solved immediately.
+*}
+
+theorem "((A --> B) --> A) --> A"
+proof
+  assume "(A --> B) --> A"
+  show A
+  proof (rule classical)
+    presume "A --> B"
+    with `(A --> B) --> A` show A ..
+  next
+    assume "~ A"
+    show "A --> B"
+    proof
+      assume A
+      with `~ A` show B by contradiction
+    qed
+  qed
+qed
+
+text {*
+ Note that the goals stemming from weak assumptions may be even left
+ until qed time, where they get eventually solved ``by assumption'' as
+ well.  In that case there is really no fundamental difference between
+ the two kinds of assumptions, apart from the order of reducing the
+ individual parts of the proof configuration.
+
+ Nevertheless, the ``strong'' mode of plain assumptions is quite
+ important in practice to achieve robustness of proof text
+ interpretation.  By forcing both the conclusion \emph{and} the
+ assumptions to unify with the pending goal to be solved, goal
+ selection becomes quite deterministic.  For example, decomposition
+ with rules of the ``case-analysis'' type usually gives rise to
+ several goals that only differ in there local contexts.  With strong
+ assumptions these may be still solved in any order in a predictable
+ way, while weak ones would quickly lead to great confusion,
+ eventually demanding even some backtracking.
+*}
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_Examples/Puzzle.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,85 @@
+header {* An old chestnut *}
+
+theory Puzzle
+imports Main
+begin
+
+text_raw {*
+  \footnote{A question from ``Bundeswettbewerb Mathematik''.  Original
+  pen-and-paper proof due to Herbert Ehler; Isabelle tactic script by
+  Tobias Nipkow.}
+*}
+
+text {*
+  \textbf{Problem.}  Given some function $f\colon \Nat \to \Nat$ such
+  that $f \ap (f \ap n) < f \ap (\idt{Suc} \ap n)$ for all $n$.
+  Demonstrate that $f$ is the identity.
+*}
+
+theorem
+  assumes f_ax: "\<And>n. f (f n) < f (Suc n)"
+  shows "f n = n"
+proof (rule order_antisym)
+  {
+    fix n show "n \<le> f n"
+    proof (induct k \<equiv> "f n" arbitrary: n rule: less_induct)
+      case (less k n)
+      then have hyp: "\<And>m. f m < f n \<Longrightarrow> m \<le> f m" by (simp only:)
+      show "n \<le> f n"
+      proof (cases n)
+        case (Suc m)
+        from f_ax have "f (f m) < f n" by (simp only: Suc)
+        with hyp have "f m \<le> f (f m)" .
+        also from f_ax have "\<dots> < f n" by (simp only: Suc)
+        finally have "f m < f n" .
+        with hyp have "m \<le> f m" .
+        also note `\<dots> < f n`
+        finally have "m < f n" .
+        then have "n \<le> f n" by (simp only: Suc)
+        then show ?thesis .
+      next
+        case 0
+        then show ?thesis by simp
+      qed
+    qed
+  } note ge = this
+
+  {
+    fix m n :: nat
+    assume "m \<le> n"
+    then have "f m \<le> f n"
+    proof (induct n)
+      case 0
+      then have "m = 0" by simp
+      then show ?case by simp
+    next
+      case (Suc n)
+      from Suc.prems show "f m \<le> f (Suc n)"
+      proof (rule le_SucE)
+        assume "m \<le> n"
+        with Suc.hyps have "f m \<le> f n" .
+        also from ge f_ax have "\<dots> < f (Suc n)"
+          by (rule le_less_trans)
+        finally show ?thesis by simp
+      next
+        assume "m = Suc n"
+        then show ?thesis by simp
+      qed
+    qed
+  } note mono = this
+
+  show "f n \<le> n"
+  proof -
+    have "\<not> n < f n"
+    proof
+      assume "n < f n"
+      then have "Suc n \<le> f n" by simp
+      then have "f (Suc n) \<le> f (f n)" by (rule mono)
+      also have "\<dots> < f (Suc n)" by (rule f_ax)
+      finally have "\<dots> < \<dots>" . then show False ..
+    qed
+    then show ?thesis by simp
+  qed
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_Examples/README.html	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,21 @@
+<!DOCTYPE HTML PUBLIC "-//W3C//DTD HTML 4.01 Transitional//EN" "http://www.w3.org/TR/html4/loose.dtd">
+
+<!-- $Id$ -->
+
+<html>
+
+<head>
+  <meta http-equiv="Content-Type" content="text/html; charset=iso-8859-1">
+  <title>HOL/Isar_Examples</title>
+</head>
+
+<body>
+<h1>HOL/Isar_Examples</h1>
+
+Isar offers a new high-level proof (and theory) language interface to
+Isabelle.  This directory contains some example Isar documents.  See
+also the included document, or the <a
+href="http://isabelle.in.tum.de/Isar/">Isabelle/Isar page</a> for more
+information.
+</body>
+</html>
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_Examples/ROOT.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,18 @@
+(* Miscellaneous Isabelle/Isar examples for Higher-Order Logic. *)
+
+no_document use_thys ["../Old_Number_Theory/Primes", "../Old_Number_Theory/Fibonacci"];
+
+use_thys [
+  "Basic_Logic",
+  "Cantor",
+  "Peirce",
+  "Drinker",
+  "Expr_Compiler",
+  "Group",
+  "Summation",
+  "Knaster_Tarski",
+  "Mutilated_Checkerboard",
+  "Puzzle",
+  "Nested_Datatype",
+  "Hoare_Ex"
+];
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_Examples/Summation.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,158 @@
+(*  Title:      HOL/Isar_Examples/Summation.thy
+    Author:     Markus Wenzel
+*)
+
+header {* Summing natural numbers *}
+
+theory Summation
+imports Main
+begin
+
+text_raw {*
+ \footnote{This example is somewhat reminiscent of the
+ \url{http://isabelle.in.tum.de/library/HOL/ex/NatSum.html}, which is
+ discussed in \cite{isabelle-ref} in the context of permutative
+ rewrite rules and ordered rewriting.}
+*}
+
+text {*
+ Subsequently, we prove some summation laws of natural numbers
+ (including odds, squares, and cubes).  These examples demonstrate how
+ plain natural deduction (including induction) may be combined with
+ calculational proof.
+*}
+
+
+subsection {* Summation laws *}
+
+text {*
+ The sum of natural numbers $0 + \cdots + n$ equals $n \times (n +
+ 1)/2$.  Avoiding formal reasoning about division we prove this
+ equation multiplied by $2$.
+*}
+
+theorem sum_of_naturals:
+  "2 * (\<Sum>i::nat=0..n. i) = n * (n + 1)"
+  (is "?P n" is "?S n = _")
+proof (induct n)
+  show "?P 0" by simp
+next
+  fix n have "?S (n + 1) = ?S n + 2 * (n + 1)" by simp
+  also assume "?S n = n * (n + 1)"
+  also have "... + 2 * (n + 1) = (n + 1) * (n + 2)" by simp
+  finally show "?P (Suc n)" by simp
+qed
+
+text {*
+ The above proof is a typical instance of mathematical induction.  The
+ main statement is viewed as some $\var{P} \ap n$ that is split by the
+ induction method into base case $\var{P} \ap 0$, and step case
+ $\var{P} \ap n \Impl \var{P} \ap (\idt{Suc} \ap n)$ for arbitrary $n$.
+
+ The step case is established by a short calculation in forward
+ manner.  Starting from the left-hand side $\var{S} \ap (n + 1)$ of
+ the thesis, the final result is achieved by transformations involving
+ basic arithmetic reasoning (using the Simplifier).  The main point is
+ where the induction hypothesis $\var{S} \ap n = n \times (n + 1)$ is
+ introduced in order to replace a certain subterm.  So the
+ ``transitivity'' rule involved here is actual \emph{substitution}.
+ Also note how the occurrence of ``\dots'' in the subsequent step
+ documents the position where the right-hand side of the hypothesis
+ got filled in.
+
+ \medskip A further notable point here is integration of calculations
+ with plain natural deduction.  This works so well in Isar for two
+ reasons.
+ \begin{enumerate}
+
+ \item Facts involved in \isakeyword{also}~/ \isakeyword{finally}
+ calculational chains may be just anything.  There is nothing special
+ about \isakeyword{have}, so the natural deduction element
+ \isakeyword{assume} works just as well.
+
+ \item There are two \emph{separate} primitives for building natural
+ deduction contexts: \isakeyword{fix}~$x$ and \isakeyword{assume}~$A$.
+ Thus it is possible to start reasoning with some new ``arbitrary, but
+ fixed'' elements before bringing in the actual assumption.  In
+ contrast, natural deduction is occasionally formalized with basic
+ context elements of the form $x:A$ instead.
+
+ \end{enumerate}
+*}
+
+text {*
+ \medskip We derive further summation laws for odds, squares, and
+ cubes as follows.  The basic technique of induction plus calculation
+ is the same as before.
+*}
+
+theorem sum_of_odds:
+  "(\<Sum>i::nat=0..<n. 2 * i + 1) = n^Suc (Suc 0)"
+  (is "?P n" is "?S n = _")
+proof (induct n)
+  show "?P 0" by simp
+next
+  fix n have "?S (n + 1) = ?S n + 2 * n + 1" by simp
+  also assume "?S n = n^Suc (Suc 0)"
+  also have "... + 2 * n + 1 = (n + 1)^Suc (Suc 0)" by simp
+  finally show "?P (Suc n)" by simp
+qed
+
+text {*
+ Subsequently we require some additional tweaking of Isabelle built-in
+ arithmetic simplifications, such as bringing in distributivity by
+ hand.
+*}
+
+lemmas distrib = add_mult_distrib add_mult_distrib2
+
+theorem sum_of_squares:
+  "6 * (\<Sum>i::nat=0..n. i^Suc (Suc 0)) = n * (n + 1) * (2 * n + 1)"
+  (is "?P n" is "?S n = _")
+proof (induct n)
+  show "?P 0" by simp
+next
+  fix n have "?S (n + 1) = ?S n + 6 * (n + 1)^Suc (Suc 0)"
+    by (simp add: distrib)
+  also assume "?S n = n * (n + 1) * (2 * n + 1)"
+  also have "... + 6 * (n + 1)^Suc (Suc 0) =
+    (n + 1) * (n + 2) * (2 * (n + 1) + 1)" by (simp add: distrib)
+  finally show "?P (Suc n)" by simp
+qed
+
+theorem sum_of_cubes:
+  "4 * (\<Sum>i::nat=0..n. i^3) = (n * (n + 1))^Suc (Suc 0)"
+  (is "?P n" is "?S n = _")
+proof (induct n)
+  show "?P 0" by (simp add: power_eq_if)
+next
+  fix n have "?S (n + 1) = ?S n + 4 * (n + 1)^3"
+    by (simp add: power_eq_if distrib)
+  also assume "?S n = (n * (n + 1))^Suc (Suc 0)"
+  also have "... + 4 * (n + 1)^3 = ((n + 1) * ((n + 1) + 1))^Suc (Suc 0)"
+    by (simp add: power_eq_if distrib)
+  finally show "?P (Suc n)" by simp
+qed
+
+text {*
+ Comparing these examples with the tactic script version
+ \url{http://isabelle.in.tum.de/library/HOL/ex/NatSum.html}, we note
+ an important difference of how induction vs.\ simplification is
+ applied.  While \cite[\S10]{isabelle-ref} advises for these examples
+ that ``induction should not be applied until the goal is in the
+ simplest form'' this would be a very bad idea in our setting.
+
+ Simplification normalizes all arithmetic expressions involved,
+ producing huge intermediate goals.  With applying induction
+ afterwards, the Isar proof text would have to reflect the emerging
+ configuration by appropriate sub-proofs.  This would result in badly
+ structured, low-level technical reasoning, without any good idea of
+ the actual point.
+
+ \medskip As a general rule of good proof style, automatic methods
+ such as $\idt{simp}$ or $\idt{auto}$ should normally be never used as
+ initial proof methods, but only as terminal ones, solving certain
+ goals completely.
+*}
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_Examples/document/proof.sty	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,254 @@
+%       proof.sty       (Proof Figure Macros)
+%
+%       version 1.0
+%       October 13, 1990
+%       Copyright (C) 1990 Makoto Tatsuta (tatsuta@riec.tohoku.ac.jp)
+%
+% This program is free software; you can redistribute it or modify
+% it under the terms of the GNU General Public License as published by
+% the Free Software Foundation; either versions 1, or (at your option)
+% any later version.
+%
+% This program is distributed in the hope that it will be useful
+% but WITHOUT ANY WARRANTY; without even the implied warranty of
+% MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE.  See the
+% GNU General Public License for more details.
+%
+%       Usage:
+%               In \documentstyle, specify an optional style `proof', say,
+%                       \documentstyle[proof]{article}.
+%
+%       The following macros are available:
+%
+%       In all the following macros, all the arguments such as
+%       <Lowers> and <Uppers> are processed in math mode.
+%
+%       \infer<Lower><Uppers>
+%               draws an inference.
+%
+%               Use & in <Uppers> to delimit upper formulae.
+%               <Uppers> consists more than 0 formulae.
+%
+%               \infer returns \hbox{ ... } or \vbox{ ... } and
+%               sets \@LeftOffset and \@RightOffset globally.
+%
+%       \infer[<Label>]<Lower><Uppers>
+%               draws an inference labeled with <Label>.
+%
+%       \infer*<Lower><Uppers>
+%               draws a many step deduction.
+%
+%       \infer*[<Label>]<Lower><Uppers>
+%               draws a many step deduction labeled with <Label>.
+%
+%       \deduce<Lower><Uppers>
+%               draws an inference without a rule.
+%
+%       \deduce[<Proof>]<Lower><Uppers>
+%               draws a many step deduction with a proof name.
+%
+%       Example:
+%               If you want to write
+%                           B C
+%                          -----
+%                      A     D
+%                     ----------
+%                         E
+%       use
+%               \infer{E}{
+%                       A
+%                       &
+%                       \infer{D}{B & C}
+%               }
+%
+
+%       Style Parameters
+
+\newdimen\inferLineSkip         \inferLineSkip=2pt
+\newdimen\inferLabelSkip        \inferLabelSkip=5pt
+\def\inferTabSkip{\quad}
+
+%       Variables
+
+\newdimen\@LeftOffset   % global
+\newdimen\@RightOffset  % global
+\newdimen\@SavedLeftOffset      % safe from users
+
+\newdimen\UpperWidth
+\newdimen\LowerWidth
+\newdimen\LowerHeight
+\newdimen\UpperLeftOffset
+\newdimen\UpperRightOffset
+\newdimen\UpperCenter
+\newdimen\LowerCenter
+\newdimen\UpperAdjust
+\newdimen\RuleAdjust
+\newdimen\LowerAdjust
+\newdimen\RuleWidth
+\newdimen\HLabelAdjust
+\newdimen\VLabelAdjust
+\newdimen\WidthAdjust
+
+\newbox\@UpperPart
+\newbox\@LowerPart
+\newbox\@LabelPart
+\newbox\ResultBox
+
+%       Flags
+
+\newif\if@inferRule     % whether \@infer draws a rule.
+\newif\if@ReturnLeftOffset      % whether \@infer returns \@LeftOffset.
+\newif\if@MathSaved     % whether inner math mode where \infer or
+                        % \deduce appears.
+
+%       Special Fonts
+
+\def\DeduceSym{\vtop{\baselineskip4\p@ \lineskiplimit\z@
+    \vbox{\hbox{.}\hbox{.}\hbox{.}}\hbox{.}}}
+
+%       Math Save Macros
+%
+%       \@SaveMath is called in the very begining of toplevel macros
+%       which are \infer and \deduce.
+%       \@RestoreMath is called in the very last before toplevel macros end.
+%       Remark \infer and \deduce ends calling \@infer.
+
+%\def\@SaveMath{\@MathSavedfalse \ifmmode \ifinner
+%        \relax $\relax \@MathSavedtrue \fi\fi }
+%
+%\def\@RestoreMath{\if@MathSaved \relax $\relax\fi }
+
+\def\@SaveMath{\relax}
+\def\@RestoreMath{\relax}
+
+
+%       Macros
+
+\def\@ifEmpty#1#2#3{\def\@tempa{\@empty}\def\@tempb{#1}\relax
+        \ifx \@tempa \@tempb #2\else #3\fi }
+
+\def\infer{\@SaveMath \@ifnextchar *{\@inferSteps}{\@inferOneStep}}
+
+\def\@inferOneStep{\@inferRuletrue
+        \@ifnextchar [{\@infer}{\@infer[\@empty]}}
+
+\def\@inferSteps*{\@ifnextchar [{\@@inferSteps}{\@@inferSteps[\@empty]}}
+
+\def\@@inferSteps[#1]{\@deduce{#1}[\DeduceSym]}
+
+\def\deduce{\@SaveMath \@ifnextchar [{\@deduce{\@empty}}
+        {\@inferRulefalse \@infer[\@empty]}}
+
+%       \@deduce<Proof Label>[<Proof>]<Lower><Uppers>
+
+\def\@deduce#1[#2]#3#4{\@inferRulefalse
+        \@infer[\@empty]{#3}{\@SaveMath \@infer[{#1}]{#2}{#4}}}
+
+%       \@infer[<Label>]<Lower><Uppers>
+%               If \@inferRuletrue, draws a rule and <Label> is right to
+%               a rule.
+%               Otherwise, draws no rule and <Label> is right to <Lower>.
+
+\def\@infer[#1]#2#3{\relax
+% Get parameters
+        \if@ReturnLeftOffset \else \@SavedLeftOffset=\@LeftOffset \fi
+        \setbox\@LabelPart=\hbox{$#1$}\relax
+        \setbox\@LowerPart=\hbox{$#2$}\relax
+%
+        \global\@LeftOffset=0pt
+        \setbox\@UpperPart=\vbox{\tabskip=0pt \halign{\relax
+                \global\@RightOffset=0pt \@ReturnLeftOffsettrue $##$&&
+                \inferTabSkip
+                \global\@RightOffset=0pt \@ReturnLeftOffsetfalse $##$\cr
+                #3\cr}}\relax
+%                       Here is a little trick.
+%                       \@ReturnLeftOffsettrue(false) influences on \infer or
+%                       \deduce placed in ## locally
+%                       because of \@SaveMath and \@RestoreMath.
+        \UpperLeftOffset=\@LeftOffset
+        \UpperRightOffset=\@RightOffset
+% Calculate Adjustments
+        \LowerWidth=\wd\@LowerPart
+        \LowerHeight=\ht\@LowerPart
+        \LowerCenter=0.5\LowerWidth
+%
+        \UpperWidth=\wd\@UpperPart \advance\UpperWidth by -\UpperLeftOffset
+        \advance\UpperWidth by -\UpperRightOffset
+        \UpperCenter=\UpperLeftOffset
+        \advance\UpperCenter by 0.5\UpperWidth
+%
+        \ifdim \UpperWidth > \LowerWidth
+                % \UpperCenter > \LowerCenter
+        \UpperAdjust=0pt
+        \RuleAdjust=\UpperLeftOffset
+        \LowerAdjust=\UpperCenter \advance\LowerAdjust by -\LowerCenter
+        \RuleWidth=\UpperWidth
+        \global\@LeftOffset=\LowerAdjust
+%
+        \else   % \UpperWidth <= \LowerWidth
+        \ifdim \UpperCenter > \LowerCenter
+%
+        \UpperAdjust=0pt
+        \RuleAdjust=\UpperCenter \advance\RuleAdjust by -\LowerCenter
+        \LowerAdjust=\RuleAdjust
+        \RuleWidth=\LowerWidth
+        \global\@LeftOffset=\LowerAdjust
+%
+        \else   % \UpperWidth <= \LowerWidth
+                % \UpperCenter <= \LowerCenter
+%
+        \UpperAdjust=\LowerCenter \advance\UpperAdjust by -\UpperCenter
+        \RuleAdjust=0pt
+        \LowerAdjust=0pt
+        \RuleWidth=\LowerWidth
+        \global\@LeftOffset=0pt
+%
+        \fi\fi
+% Make a box
+        \if@inferRule
+%
+        \setbox\ResultBox=\vbox{
+                \moveright \UpperAdjust \box\@UpperPart
+                \nointerlineskip \kern\inferLineSkip
+                \moveright \RuleAdjust \vbox{\hrule width\RuleWidth}\relax
+                \nointerlineskip \kern\inferLineSkip
+                \moveright \LowerAdjust \box\@LowerPart }\relax
+%
+        \@ifEmpty{#1}{}{\relax
+%
+        \HLabelAdjust=\wd\ResultBox     \advance\HLabelAdjust by -\RuleAdjust
+        \advance\HLabelAdjust by -\RuleWidth
+        \WidthAdjust=\HLabelAdjust
+        \advance\WidthAdjust by -\inferLabelSkip
+        \advance\WidthAdjust by -\wd\@LabelPart
+        \ifdim \WidthAdjust < 0pt \WidthAdjust=0pt \fi
+%
+        \VLabelAdjust=\dp\@LabelPart
+        \advance\VLabelAdjust by -\ht\@LabelPart
+        \VLabelAdjust=0.5\VLabelAdjust  \advance\VLabelAdjust by \LowerHeight
+        \advance\VLabelAdjust by \inferLineSkip
+%
+        \setbox\ResultBox=\hbox{\box\ResultBox
+                \kern -\HLabelAdjust \kern\inferLabelSkip
+                \raise\VLabelAdjust \box\@LabelPart \kern\WidthAdjust}\relax
+%
+        }\relax % end @ifEmpty
+%
+        \else % \@inferRulefalse
+%
+        \setbox\ResultBox=\vbox{
+                \moveright \UpperAdjust \box\@UpperPart
+                \nointerlineskip \kern\inferLineSkip
+                \moveright \LowerAdjust \hbox{\unhbox\@LowerPart
+                        \@ifEmpty{#1}{}{\relax
+                        \kern\inferLabelSkip \unhbox\@LabelPart}}}\relax
+        \fi
+%
+        \global\@RightOffset=\wd\ResultBox
+        \global\advance\@RightOffset by -\@LeftOffset
+        \global\advance\@RightOffset by -\LowerWidth
+        \if@ReturnLeftOffset \else \global\@LeftOffset=\@SavedLeftOffset \fi
+%
+        \box\ResultBox
+        \@RestoreMath
+}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_Examples/document/root.bib	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,91 @@
+
+@string{CUCL="Comp. Lab., Univ. Camb."}
+@string{CUP="Cambridge University Press"}
+@string{Springer="Springer-Verlag"}
+@string{TUM="TU Munich"}
+
+@Book{Concrete-Math,
+  author = 	 {R. L. Graham and D. E. Knuth and O. Patashnik},
+  title = 	 {Concrete Mathematics},
+  publisher = 	 {Addison-Wesley},
+  year = 	 1989
+}
+
+@InProceedings{Naraschewski-Wenzel:1998:HOOL,
+  author	= {Wolfgang Naraschewski and Markus Wenzel},
+  title		= {Object-Oriented Verification based on Record Subtyping in
+                  {H}igher-{O}rder {L}ogic},
+  crossref      = {tphols98}}
+
+@Article{Nipkow:1998:Winskel,
+  author = 	 {Tobias Nipkow},
+  title = 	 {Winskel is (almost) Right: Towards a Mechanized Semantics Textbook},
+  journal = 	 {Formal Aspects of Computing},
+  year = 	 1998,
+  volume =	 10,
+  pages =	 {171--186}
+}
+
+@InProceedings{Wenzel:1999:TPHOL,
+  author = 	 {Markus Wenzel},
+  title = 	 {{Isar} --- a Generic Interpretative Approach to Readable Formal Proof Documents},
+  crossref =     {tphols99}}
+
+@Book{Winskel:1993,
+  author = 	 {G. Winskel},
+  title = 	 {The Formal Semantics of Programming Languages},
+  publisher = 	 {MIT Press},
+  year = 	 1993
+}
+
+@Book{davey-priestley,
+  author	= {B. A. Davey and H. A. Priestley},
+  title		= {Introduction to Lattices and Order},
+  publisher	= CUP,
+  year		= 1990}
+
+@manual{isabelle-HOL,
+  author	= {Tobias Nipkow and Lawrence C. Paulson and Markus Wenzel},
+  title		= {{Isabelle}'s Logics: {HOL}},
+  institution	= {Institut f\"ur Informatik, Technische Universi\"at
+                  M\"unchen and Computer Laboratory, University of Cambridge}}
+
+@manual{isabelle-intro,
+  author	= {Lawrence C. Paulson},
+  title		= {Introduction to {Isabelle}},
+  institution	= CUCL}
+
+@manual{isabelle-isar-ref,
+  author	= {Markus Wenzel},
+  title		= {The {Isabelle/Isar} Reference Manual},
+  institution	= TUM}
+
+@manual{isabelle-ref,
+  author	= {Lawrence C. Paulson},
+  title		= {The {Isabelle} Reference Manual},
+  institution	= CUCL}
+
+@TechReport{paulson-mutilated-board,
+  author = 	 {Lawrence C. Paulson},
+  title = 	 {A Simple Formalization and Proof for the Mutilated Chess Board},
+  institution =  CUCL,
+  year = 	 1996,
+  number =	 394,
+  note = {\url{http://www.cl.cam.ac.uk/users/lcp/papers/Reports/mutil.pdf}}
+}
+
+@Proceedings{tphols98,
+  title		= {Theorem Proving in Higher Order Logics: {TPHOLs} '98},
+  booktitle	= {Theorem Proving in Higher Order Logics: {TPHOLs} '98},
+  editor	= {Jim Grundy and Malcom Newey},
+  series	= {LNCS},
+  volume        = 1479,
+  year		= 1998}
+
+@Proceedings{tphols99,
+  title		= {Theorem Proving in Higher Order Logics: {TPHOLs} '99},
+  booktitle	= {Theorem Proving in Higher Order Logics: {TPHOLs} '99},
+  editor	= {Bertot, Y. and Dowek, G. and Hirschowitz, A. and
+                  Paulin, C. and Thery, L.},
+  series	= {LNCS 1690},
+  year		= 1999}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_Examples/document/root.tex	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,30 @@
+\input{style}
+
+\hyphenation{Isabelle}
+
+\begin{document}
+
+\title{Miscellaneous Isabelle/Isar examples for Higher-Order Logic}
+\author{Markus Wenzel \\ \url{http://www.in.tum.de/~wenzelm/} \\[2ex]
+  With contributions by Gertrud Bauer and Tobias Nipkow}
+\maketitle
+
+\begin{abstract}
+  Isar offers a high-level proof (and theory) language for Isabelle.
+  We give various examples of Isabelle/Isar proof developments,
+  ranging from simple demonstrations of certain language features to a
+  bit more advanced applications.  The ``real'' applications of
+  Isabelle/Isar are found elsewhere.
+\end{abstract}
+
+\tableofcontents
+
+\parindent 0pt \parskip 0.5ex
+
+\input{session}
+
+\nocite{isabelle-isar-ref,Wenzel:1999:TPHOL}
+\bibliographystyle{abbrv}
+\bibliography{root}
+
+\end{document}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Isar_Examples/document/style.tex	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,40 @@
+\documentclass[11pt,a4paper]{article}
+\usepackage[only,bigsqcap]{stmaryrd}
+\usepackage{ifthen,proof,amssymb,isabelle,isabellesym}
+\isabellestyle{it}
+\usepackage{pdfsetup}\urlstyle{rm}
+
+\renewcommand{\isamarkupheader}[1]{\section{#1}}
+
+\renewcommand{\isacommand}[1]
+{\ifthenelse{\equal{sorry}{#1}}{$\;$\dummyproof}
+  {\ifthenelse{\equal{oops}{#1}}{$\vdots$}{\isakeyword{#1}}}}
+
+\newcommand{\DUMMYPROOF}{{\langle\idt{proof}\rangle}}
+\newcommand{\dummyproof}{$\DUMMYPROOF$}
+
+\newcommand{\name}[1]{\textsl{#1}}
+
+\newcommand{\idt}[1]{{\mathord{\mathit{#1}}}}
+\newcommand{\var}[1]{{?\!\idt{#1}}}
+\DeclareMathSymbol{\dshsym}{\mathalpha}{letters}{"2D}
+\newcommand{\dsh}{\dshsym}
+
+\newcommand{\To}{\to}
+\newcommand{\dt}{{\mathpunct.}}
+\newcommand{\ap}{\mathbin{\!}}
+\newcommand{\lam}[1]{\mathop{\lambda} #1\dt\;}
+\newcommand{\all}[1]{\forall #1\dt\;}
+\newcommand{\ex}[1]{\exists #1\dt\;}
+\newcommand{\impl}{\to}
+\newcommand{\conj}{\land}
+\newcommand{\disj}{\lor}
+\newcommand{\Impl}{\Longrightarrow}
+
+\newcommand{\Nat}{\mathord{\mathrm{I}\mkern-3.8mu\mathrm{N}}}
+
+
+%%% Local Variables: 
+%%% mode: latex
+%%% TeX-master: "root"
+%%% End: 
--- a/src/HOL/Isar_examples/Basic_Logic.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,448 +0,0 @@
-(*  Title:      HOL/Isar_examples/Basic_Logic.thy
-    Author:     Markus Wenzel, TU Muenchen
-
-Basic propositional and quantifier reasoning.
-*)
-
-header {* Basic logical reasoning *}
-
-theory Basic_Logic
-imports Main
-begin
-
-
-subsection {* Pure backward reasoning *}
-
-text {*
-  In order to get a first idea of how Isabelle/Isar proof documents
-  may look like, we consider the propositions @{text I}, @{text K},
-  and @{text S}.  The following (rather explicit) proofs should
-  require little extra explanations.
-*}
-
-lemma I: "A --> A"
-proof
-  assume A
-  show A by fact
-qed
-
-lemma K: "A --> B --> A"
-proof
-  assume A
-  show "B --> A"
-  proof
-    show A by fact
-  qed
-qed
-
-lemma S: "(A --> B --> C) --> (A --> B) --> A --> C"
-proof
-  assume "A --> B --> C"
-  show "(A --> B) --> A --> C"
-  proof
-    assume "A --> B"
-    show "A --> C"
-    proof
-      assume A
-      show C
-      proof (rule mp)
-        show "B --> C" by (rule mp) fact+
-        show B by (rule mp) fact+
-      qed
-    qed
-  qed
-qed
-
-text {*
-  Isar provides several ways to fine-tune the reasoning, avoiding
-  excessive detail.  Several abbreviated language elements are
-  available, enabling the writer to express proofs in a more concise
-  way, even without referring to any automated proof tools yet.
-
-  First of all, proof by assumption may be abbreviated as a single
-  dot.
-*}
-
-lemma "A --> A"
-proof
-  assume A
-  show A by fact+
-qed
-
-text {*
-  In fact, concluding any (sub-)proof already involves solving any
-  remaining goals by assumption\footnote{This is not a completely
-  trivial operation, as proof by assumption may involve full
-  higher-order unification.}.  Thus we may skip the rather vacuous
-  body of the above proof as well.
-*}
-
-lemma "A --> A"
-proof
-qed
-
-text {*
-  Note that the \isacommand{proof} command refers to the @{text rule}
-  method (without arguments) by default.  Thus it implicitly applies a
-  single rule, as determined from the syntactic form of the statements
-  involved.  The \isacommand{by} command abbreviates any proof with
-  empty body, so the proof may be further pruned.
-*}
-
-lemma "A --> A"
-  by rule
-
-text {*
-  Proof by a single rule may be abbreviated as double-dot.
-*}
-
-lemma "A --> A" ..
-
-text {*
-  Thus we have arrived at an adequate representation of the proof of a
-  tautology that holds by a single standard rule.\footnote{Apparently,
-  the rule here is implication introduction.}
-*}
-
-text {*
-  Let us also reconsider @{text K}.  Its statement is composed of
-  iterated connectives.  Basic decomposition is by a single rule at a
-  time, which is why our first version above was by nesting two
-  proofs.
-
-  The @{text intro} proof method repeatedly decomposes a goal's
-  conclusion.\footnote{The dual method is @{text elim}, acting on a
-  goal's premises.}
-*}
-
-lemma "A --> B --> A"
-proof (intro impI)
-  assume A
-  show A by fact
-qed
-
-text {*
-  Again, the body may be collapsed.
-*}
-
-lemma "A --> B --> A"
-  by (intro impI)
-
-text {*
-  Just like @{text rule}, the @{text intro} and @{text elim} proof
-  methods pick standard structural rules, in case no explicit
-  arguments are given.  While implicit rules are usually just fine for
-  single rule application, this may go too far with iteration.  Thus
-  in practice, @{text intro} and @{text elim} would be typically
-  restricted to certain structures by giving a few rules only, e.g.\
-  \isacommand{proof}~@{text "(intro impI allI)"} to strip implications
-  and universal quantifiers.
-
-  Such well-tuned iterated decomposition of certain structures is the
-  prime application of @{text intro} and @{text elim}.  In contrast,
-  terminal steps that solve a goal completely are usually performed by
-  actual automated proof methods (such as \isacommand{by}~@{text
-  blast}.
-*}
-
-
-subsection {* Variations of backward vs.\ forward reasoning *}
-
-text {*
-  Certainly, any proof may be performed in backward-style only.  On
-  the other hand, small steps of reasoning are often more naturally
-  expressed in forward-style.  Isar supports both backward and forward
-  reasoning as a first-class concept.  In order to demonstrate the
-  difference, we consider several proofs of @{text "A \<and> B \<longrightarrow> B \<and> A"}.
-
-  The first version is purely backward.
-*}
-
-lemma "A & B --> B & A"
-proof
-  assume "A & B"
-  show "B & A"
-  proof
-    show B by (rule conjunct2) fact
-    show A by (rule conjunct1) fact
-  qed
-qed
-
-text {*
-  Above, the @{text "conjunct_1/2"} projection rules had to be named
-  explicitly, since the goals @{text B} and @{text A} did not provide
-  any structural clue.  This may be avoided using \isacommand{from} to
-  focus on the @{text "A \<and> B"} assumption as the current facts,
-  enabling the use of double-dot proofs.  Note that \isacommand{from}
-  already does forward-chaining, involving the \name{conjE} rule here.
-*}
-
-lemma "A & B --> B & A"
-proof
-  assume "A & B"
-  show "B & A"
-  proof
-    from `A & B` show B ..
-    from `A & B` show A ..
-  qed
-qed
-
-text {*
-  In the next version, we move the forward step one level upwards.
-  Forward-chaining from the most recent facts is indicated by the
-  \isacommand{then} command.  Thus the proof of @{text "B \<and> A"} from
-  @{text "A \<and> B"} actually becomes an elimination, rather than an
-  introduction.  The resulting proof structure directly corresponds to
-  that of the @{text conjE} rule, including the repeated goal
-  proposition that is abbreviated as @{text ?thesis} below.
-*}
-
-lemma "A & B --> B & A"
-proof
-  assume "A & B"
-  then show "B & A"
-  proof                    -- {* rule @{text conjE} of @{text "A \<and> B"} *}
-    assume B A
-    then show ?thesis ..   -- {* rule @{text conjI} of @{text "B \<and> A"} *}
-  qed
-qed
-
-text {*
-  In the subsequent version we flatten the structure of the main body
-  by doing forward reasoning all the time.  Only the outermost
-  decomposition step is left as backward.
-*}
-
-lemma "A & B --> B & A"
-proof
-  assume "A & B"
-  from `A & B` have A ..
-  from `A & B` have B ..
-  from `B` `A` show "B & A" ..
-qed
-
-text {*
-  We can still push forward-reasoning a bit further, even at the risk
-  of getting ridiculous.  Note that we force the initial proof step to
-  do nothing here, by referring to the ``-'' proof method.
-*}
-
-lemma "A & B --> B & A"
-proof -
-  {
-    assume "A & B"
-    from `A & B` have A ..
-    from `A & B` have B ..
-    from `B` `A` have "B & A" ..
-  }
-  then show ?thesis ..         -- {* rule \name{impI} *}
-qed
-
-text {*
-  \medskip With these examples we have shifted through a whole range
-  from purely backward to purely forward reasoning.  Apparently, in
-  the extreme ends we get slightly ill-structured proofs, which also
-  require much explicit naming of either rules (backward) or local
-  facts (forward).
-
-  The general lesson learned here is that good proof style would
-  achieve just the \emph{right} balance of top-down backward
-  decomposition, and bottom-up forward composition.  In general, there
-  is no single best way to arrange some pieces of formal reasoning, of
-  course.  Depending on the actual applications, the intended audience
-  etc., rules (and methods) on the one hand vs.\ facts on the other
-  hand have to be emphasized in an appropriate way.  This requires the
-  proof writer to develop good taste, and some practice, of course.
-*}
-
-text {*
-  For our example the most appropriate way of reasoning is probably
-  the middle one, with conjunction introduction done after
-  elimination.
-*}
-
-lemma "A & B --> B & A"
-proof
-  assume "A & B"
-  then show "B & A"
-  proof
-    assume B A
-    then show ?thesis ..
-  qed
-qed
-
-
-
-subsection {* A few examples from ``Introduction to Isabelle'' *}
-
-text {*
-  We rephrase some of the basic reasoning examples of
-  \cite{isabelle-intro}, using HOL rather than FOL.
-*}
-
-subsubsection {* A propositional proof *}
-
-text {*
-  We consider the proposition @{text "P \<or> P \<longrightarrow> P"}.  The proof below
-  involves forward-chaining from @{text "P \<or> P"}, followed by an
-  explicit case-analysis on the two \emph{identical} cases.
-*}
-
-lemma "P | P --> P"
-proof
-  assume "P | P"
-  then show P
-  proof                    -- {*
-    rule @{text disjE}: \smash{$\infer{C}{A \disj B & \infer*{C}{[A]} & \infer*{C}{[B]}}$}
-  *}
-    assume P show P by fact
-  next
-    assume P show P by fact
-  qed
-qed
-
-text {*
-  Case splits are \emph{not} hardwired into the Isar language as a
-  special feature.  The \isacommand{next} command used to separate the
-  cases above is just a short form of managing block structure.
-
-  \medskip In general, applying proof methods may split up a goal into
-  separate ``cases'', i.e.\ new subgoals with individual local
-  assumptions.  The corresponding proof text typically mimics this by
-  establishing results in appropriate contexts, separated by blocks.
-
-  In order to avoid too much explicit parentheses, the Isar system
-  implicitly opens an additional block for any new goal, the
-  \isacommand{next} statement then closes one block level, opening a
-  new one.  The resulting behavior is what one would expect from
-  separating cases, only that it is more flexible.  E.g.\ an induction
-  base case (which does not introduce local assumptions) would
-  \emph{not} require \isacommand{next} to separate the subsequent step
-  case.
-
-  \medskip In our example the situation is even simpler, since the two
-  cases actually coincide.  Consequently the proof may be rephrased as
-  follows.
-*}
-
-lemma "P | P --> P"
-proof
-  assume "P | P"
-  then show P
-  proof
-    assume P
-    show P by fact
-    show P by fact
-  qed
-qed
-
-text {*
-  Again, the rather vacuous body of the proof may be collapsed.  Thus
-  the case analysis degenerates into two assumption steps, which are
-  implicitly performed when concluding the single rule step of the
-  double-dot proof as follows.
-*}
-
-lemma "P | P --> P"
-proof
-  assume "P | P"
-  then show P ..
-qed
-
-
-subsubsection {* A quantifier proof *}
-
-text {*
-  To illustrate quantifier reasoning, let us prove @{text "(\<exists>x. P (f
-  x)) \<longrightarrow> (\<exists>y. P y)"}.  Informally, this holds because any @{text a}
-  with @{text "P (f a)"} may be taken as a witness for the second
-  existential statement.
-
-  The first proof is rather verbose, exhibiting quite a lot of
-  (redundant) detail.  It gives explicit rules, even with some
-  instantiation.  Furthermore, we encounter two new language elements:
-  the \isacommand{fix} command augments the context by some new
-  ``arbitrary, but fixed'' element; the \isacommand{is} annotation
-  binds term abbreviations by higher-order pattern matching.
-*}
-
-lemma "(EX x. P (f x)) --> (EX y. P y)"
-proof
-  assume "EX x. P (f x)"
-  then show "EX y. P y"
-  proof (rule exE)             -- {*
-    rule \name{exE}: \smash{$\infer{B}{\ex x A(x) & \infer*{B}{[A(x)]_x}}$}
-  *}
-    fix a
-    assume "P (f a)" (is "P ?witness")
-    then show ?thesis by (rule exI [of P ?witness])
-  qed
-qed
-
-text {*
-  While explicit rule instantiation may occasionally improve
-  readability of certain aspects of reasoning, it is usually quite
-  redundant.  Above, the basic proof outline gives already enough
-  structural clues for the system to infer both the rules and their
-  instances (by higher-order unification).  Thus we may as well prune
-  the text as follows.
-*}
-
-lemma "(EX x. P (f x)) --> (EX y. P y)"
-proof
-  assume "EX x. P (f x)"
-  then show "EX y. P y"
-  proof
-    fix a
-    assume "P (f a)"
-    then show ?thesis ..
-  qed
-qed
-
-text {*
-  Explicit @{text \<exists>}-elimination as seen above can become quite
-  cumbersome in practice.  The derived Isar language element
-  ``\isakeyword{obtain}'' provides a more handsome way to do
-  generalized existence reasoning.
-*}
-
-lemma "(EX x. P (f x)) --> (EX y. P y)"
-proof
-  assume "EX x. P (f x)"
-  then obtain a where "P (f a)" ..
-  then show "EX y. P y" ..
-qed
-
-text {*
-  Technically, \isakeyword{obtain} is similar to \isakeyword{fix} and
-  \isakeyword{assume} together with a soundness proof of the
-  elimination involved.  Thus it behaves similar to any other forward
-  proof element.  Also note that due to the nature of general
-  existence reasoning involved here, any result exported from the
-  context of an \isakeyword{obtain} statement may \emph{not} refer to
-  the parameters introduced there.
-*}
-
-
-
-subsubsection {* Deriving rules in Isabelle *}
-
-text {*
-  We derive the conjunction elimination rule from the corresponding
-  projections.  The proof is quite straight-forward, since
-  Isabelle/Isar supports non-atomic goals and assumptions fully
-  transparently.
-*}
-
-theorem conjE: "A & B ==> (A ==> B ==> C) ==> C"
-proof -
-  assume "A & B"
-  assume r: "A ==> B ==> C"
-  show C
-  proof (rule r)
-    show A by (rule conjunct1) fact
-    show B by (rule conjunct2) fact
-  qed
-qed
-
-end
--- a/src/HOL/Isar_examples/Cantor.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,71 +0,0 @@
-(*  Title:      HOL/Isar_examples/Cantor.thy
-    Author:     Markus Wenzel, TU Muenchen
-*)
-
-header {* Cantor's Theorem *}
-
-theory Cantor
-imports Main
-begin
-
-text_raw {*
-  \footnote{This is an Isar version of the final example of the
-  Isabelle/HOL manual \cite{isabelle-HOL}.}
-*}
-
-text {*
-  Cantor's Theorem states that every set has more subsets than it has
-  elements.  It has become a favorite basic example in pure
-  higher-order logic since it is so easily expressed: \[\all{f::\alpha
-  \To \alpha \To \idt{bool}} \ex{S::\alpha \To \idt{bool}}
-  \all{x::\alpha} f \ap x \not= S\]
-
-  Viewing types as sets, $\alpha \To \idt{bool}$ represents the
-  powerset of $\alpha$.  This version of the theorem states that for
-  every function from $\alpha$ to its powerset, some subset is outside
-  its range.  The Isabelle/Isar proofs below uses HOL's set theory,
-  with the type $\alpha \ap \idt{set}$ and the operator
-  $\idt{range}::(\alpha \To \beta) \To \beta \ap \idt{set}$.
-*}
-
-theorem "EX S. S ~: range (f :: 'a => 'a set)"
-proof
-  let ?S = "{x. x ~: f x}"
-  show "?S ~: range f"
-  proof
-    assume "?S : range f"
-    then obtain y where "?S = f y" ..
-    then show False
-    proof (rule equalityCE)
-      assume "y : f y"
-      assume "y : ?S" then have "y ~: f y" ..
-      with `y : f y` show ?thesis by contradiction
-    next
-      assume "y ~: ?S"
-      assume "y ~: f y" then have "y : ?S" ..
-      with `y ~: ?S` show ?thesis by contradiction
-    qed
-  qed
-qed
-
-text {*
-  How much creativity is required?  As it happens, Isabelle can prove
-  this theorem automatically using best-first search.  Depth-first
-  search would diverge, but best-first search successfully navigates
-  through the large search space.  The context of Isabelle's classical
-  prover contains rules for the relevant constructs of HOL's set
-  theory.
-*}
-
-theorem "EX S. S ~: range (f :: 'a => 'a set)"
-  by best
-
-text {*
-  While this establishes the same theorem internally, we do not get
-  any idea of how the proof actually works.  There is currently no way
-  to transform internal system-level representations of Isabelle
-  proofs back into Isar text.  Writing intelligible proof documents
-  really is a creative process, after all.
-*}
-
-end
--- a/src/HOL/Isar_examples/Drinker.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,54 +0,0 @@
-(*  Title:      HOL/Isar_examples/Drinker.thy
-    Author:     Makarius
-*)
-
-header {* The Drinker's Principle *}
-
-theory Drinker
-imports Main
-begin
-
-text {*
-  Here is another example of classical reasoning: the Drinker's
-  Principle says that for some person, if he is drunk, everybody else
-  is drunk!
-
-  We first prove a classical part of de-Morgan's law.
-*}
-
-lemma deMorgan:
-  assumes "\<not> (\<forall>x. P x)"
-  shows "\<exists>x. \<not> P x"
-  using prems
-proof (rule contrapos_np)
-  assume a: "\<not> (\<exists>x. \<not> P x)"
-  show "\<forall>x. P x"
-  proof
-    fix x
-    show "P x"
-    proof (rule classical)
-      assume "\<not> P x"
-      then have "\<exists>x. \<not> P x" ..
-      with a show ?thesis by contradiction
-    qed
-  qed
-qed
-
-theorem Drinker's_Principle: "\<exists>x. drunk x \<longrightarrow> (\<forall>x. drunk x)"
-proof cases
-  fix a assume "\<forall>x. drunk x"
-  then have "drunk a \<longrightarrow> (\<forall>x. drunk x)" ..
-  then show ?thesis ..
-next
-  assume "\<not> (\<forall>x. drunk x)"
-  then have "\<exists>x. \<not> drunk x" by (rule deMorgan)
-  then obtain a where a: "\<not> drunk a" ..
-  have "drunk a \<longrightarrow> (\<forall>x. drunk x)"
-  proof
-    assume "drunk a"
-    with a show "\<forall>x. drunk x" by (contradiction)
-  qed
-  then show ?thesis ..
-qed
-
-end
--- a/src/HOL/Isar_examples/Expr_Compiler.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,231 +0,0 @@
-(*  Title:      HOL/Isar_examples/Expr_Compiler.thy
-    Author:     Markus Wenzel, TU Muenchen
-
-Correctness of a simple expression/stack-machine compiler.
-*)
-
-header {* Correctness of a simple expression compiler *}
-
-theory Expr_Compiler
-imports Main
-begin
-
-text {*
- This is a (rather trivial) example of program verification.  We model
- a compiler for translating expressions to stack machine instructions,
- and prove its correctness wrt.\ some evaluation semantics.
-*}
-
-
-subsection {* Binary operations *}
-
-text {*
- Binary operations are just functions over some type of values.  This
- is both for abstract syntax and semantics, i.e.\ we use a ``shallow
- embedding'' here.
-*}
-
-types
-  'val binop = "'val => 'val => 'val"
-
-
-subsection {* Expressions *}
-
-text {*
- The language of expressions is defined as an inductive type,
- consisting of variables, constants, and binary operations on
- expressions.
-*}
-
-datatype ('adr, 'val) expr =
-  Variable 'adr |
-  Constant 'val |
-  Binop "'val binop" "('adr, 'val) expr" "('adr, 'val) expr"
-
-text {*
- Evaluation (wrt.\ some environment of variable assignments) is
- defined by primitive recursion over the structure of expressions.
-*}
-
-consts
-  eval :: "('adr, 'val) expr => ('adr => 'val) => 'val"
-
-primrec
-  "eval (Variable x) env = env x"
-  "eval (Constant c) env = c"
-  "eval (Binop f e1 e2) env = f (eval e1 env) (eval e2 env)"
-
-
-subsection {* Machine *}
-
-text {*
- Next we model a simple stack machine, with three instructions.
-*}
-
-datatype ('adr, 'val) instr =
-  Const 'val |
-  Load 'adr |
-  Apply "'val binop"
-
-text {*
- Execution of a list of stack machine instructions is easily defined
- as follows.
-*}
-
-consts
-  exec :: "(('adr, 'val) instr) list
-    => 'val list => ('adr => 'val) => 'val list"
-
-primrec
-  "exec [] stack env = stack"
-  "exec (instr # instrs) stack env =
-    (case instr of
-      Const c => exec instrs (c # stack) env
-    | Load x => exec instrs (env x # stack) env
-    | Apply f => exec instrs (f (hd stack) (hd (tl stack))
-                   # (tl (tl stack))) env)"
-
-constdefs
-  execute :: "(('adr, 'val) instr) list => ('adr => 'val) => 'val"
-  "execute instrs env == hd (exec instrs [] env)"
-
-
-subsection {* Compiler *}
-
-text {*
- We are ready to define the compilation function of expressions to
- lists of stack machine instructions.
-*}
-
-consts
-  compile :: "('adr, 'val) expr => (('adr, 'val) instr) list"
-
-primrec
-  "compile (Variable x) = [Load x]"
-  "compile (Constant c) = [Const c]"
-  "compile (Binop f e1 e2) = compile e2 @ compile e1 @ [Apply f]"
-
-
-text {*
- The main result of this development is the correctness theorem for
- $\idt{compile}$.  We first establish a lemma about $\idt{exec}$ and
- list append.
-*}
-
-lemma exec_append:
-  "exec (xs @ ys) stack env =
-    exec ys (exec xs stack env) env"
-proof (induct xs arbitrary: stack)
-  case Nil
-  show ?case by simp
-next
-  case (Cons x xs)
-  show ?case
-  proof (induct x)
-    case Const
-    from Cons show ?case by simp
-  next
-    case Load
-    from Cons show ?case by simp
-  next
-    case Apply
-    from Cons show ?case by simp
-  qed
-qed
-
-theorem correctness: "execute (compile e) env = eval e env"
-proof -
-  have "\<And>stack. exec (compile e) stack env = eval e env # stack"
-  proof (induct e)
-    case Variable show ?case by simp
-  next
-    case Constant show ?case by simp
-  next
-    case Binop then show ?case by (simp add: exec_append)
-  qed
-  then show ?thesis by (simp add: execute_def)
-qed
-
-
-text {*
- \bigskip In the proofs above, the \name{simp} method does quite a lot
- of work behind the scenes (mostly ``functional program execution'').
- Subsequently, the same reasoning is elaborated in detail --- at most
- one recursive function definition is used at a time.  Thus we get a
- better idea of what is actually going on.
-*}
-
-lemma exec_append':
-  "exec (xs @ ys) stack env = exec ys (exec xs stack env) env"
-proof (induct xs arbitrary: stack)
-  case (Nil s)
-  have "exec ([] @ ys) s env = exec ys s env" by simp
-  also have "... = exec ys (exec [] s env) env" by simp
-  finally show ?case .
-next
-  case (Cons x xs s)
-  show ?case
-  proof (induct x)
-    case (Const val)
-    have "exec ((Const val # xs) @ ys) s env = exec (Const val # xs @ ys) s env"
-      by simp
-    also have "... = exec (xs @ ys) (val # s) env" by simp
-    also from Cons have "... = exec ys (exec xs (val # s) env) env" .
-    also have "... = exec ys (exec (Const val # xs) s env) env" by simp
-    finally show ?case .
-  next
-    case (Load adr)
-    from Cons show ?case by simp -- {* same as above *}
-  next
-    case (Apply fn)
-    have "exec ((Apply fn # xs) @ ys) s env =
-        exec (Apply fn # xs @ ys) s env" by simp
-    also have "... =
-        exec (xs @ ys) (fn (hd s) (hd (tl s)) # (tl (tl s))) env" by simp
-    also from Cons have "... =
-        exec ys (exec xs (fn (hd s) (hd (tl s)) # tl (tl s)) env) env" .
-    also have "... = exec ys (exec (Apply fn # xs) s env) env" by simp
-    finally show ?case .
-  qed
-qed
-
-theorem correctness': "execute (compile e) env = eval e env"
-proof -
-  have exec_compile: "\<And>stack. exec (compile e) stack env = eval e env # stack"
-  proof (induct e)
-    case (Variable adr s)
-    have "exec (compile (Variable adr)) s env = exec [Load adr] s env"
-      by simp
-    also have "... = env adr # s" by simp
-    also have "env adr = eval (Variable adr) env" by simp
-    finally show ?case .
-  next
-    case (Constant val s)
-    show ?case by simp -- {* same as above *}
-  next
-    case (Binop fn e1 e2 s)
-    have "exec (compile (Binop fn e1 e2)) s env =
-        exec (compile e2 @ compile e1 @ [Apply fn]) s env" by simp
-    also have "... = exec [Apply fn]
-        (exec (compile e1) (exec (compile e2) s env) env) env"
-      by (simp only: exec_append)
-    also have "exec (compile e2) s env = eval e2 env # s" by fact
-    also have "exec (compile e1) ... env = eval e1 env # ..." by fact
-    also have "exec [Apply fn] ... env =
-        fn (hd ...) (hd (tl ...)) # (tl (tl ...))" by simp
-    also have "... = fn (eval e1 env) (eval e2 env) # s" by simp
-    also have "fn (eval e1 env) (eval e2 env) =
-        eval (Binop fn e1 e2) env"
-      by simp
-    finally show ?case .
-  qed
-
-  have "execute (compile e) env = hd (exec (compile e) [] env)"
-    by (simp add: execute_def)
-  also from exec_compile
-    have "exec (compile e) [] env = [eval e env]" .
-  also have "hd ... = eval e env" by simp
-  finally show ?thesis .
-qed
-
-end
--- a/src/HOL/Isar_examples/Fibonacci.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,165 +0,0 @@
-(*  Title:      HOL/Isar_examples/Fibonacci.thy
-    Author:     Gertrud Bauer
-    Copyright   1999 Technische Universitaet Muenchen
-
-The Fibonacci function.  Demonstrates the use of recdef.  Original
-tactic script by Lawrence C Paulson.
-
-Fibonacci numbers: proofs of laws taken from
-
-  R. L. Graham, D. E. Knuth, O. Patashnik.
-  Concrete Mathematics.
-  (Addison-Wesley, 1989)
-*)
-
-header {* Fib and Gcd commute *}
-
-theory Fibonacci
-imports Primes
-begin
-
-text_raw {*
- \footnote{Isar version by Gertrud Bauer.  Original tactic script by
- Larry Paulson.  A few proofs of laws taken from
- \cite{Concrete-Math}.}
-*}
-
-
-subsection {* Fibonacci numbers *}
-
-fun fib :: "nat \<Rightarrow> nat" where
-  "fib 0 = 0"
-  | "fib (Suc 0) = 1"
-  | "fib (Suc (Suc x)) = fib x + fib (Suc x)"
-
-lemma [simp]: "0 < fib (Suc n)"
-  by (induct n rule: fib.induct) simp_all
-
-
-text {* Alternative induction rule. *}
-
-theorem fib_induct:
-    "P 0 ==> P 1 ==> (!!n. P (n + 1) ==> P n ==> P (n + 2)) ==> P (n::nat)"
-  by (induct rule: fib.induct) simp_all
-
-
-subsection {* Fib and gcd commute *}
-
-text {* A few laws taken from \cite{Concrete-Math}. *}
-
-lemma fib_add:
-  "fib (n + k + 1) = fib (k + 1) * fib (n + 1) + fib k * fib n"
-  (is "?P n")
-  -- {* see \cite[page 280]{Concrete-Math} *}
-proof (induct n rule: fib_induct)
-  show "?P 0" by simp
-  show "?P 1" by simp
-  fix n
-  have "fib (n + 2 + k + 1)
-    = fib (n + k + 1) + fib (n + 1 + k + 1)" by simp
-  also assume "fib (n + k + 1)
-    = fib (k + 1) * fib (n + 1) + fib k * fib n"
-      (is " _ = ?R1")
-  also assume "fib (n + 1 + k + 1)
-    = fib (k + 1) * fib (n + 1 + 1) + fib k * fib (n + 1)"
-      (is " _ = ?R2")
-  also have "?R1 + ?R2
-    = fib (k + 1) * fib (n + 2 + 1) + fib k * fib (n + 2)"
-    by (simp add: add_mult_distrib2)
-  finally show "?P (n + 2)" .
-qed
-
-lemma gcd_fib_Suc_eq_1: "gcd (fib n) (fib (n + 1)) = 1" (is "?P n")
-proof (induct n rule: fib_induct)
-  show "?P 0" by simp
-  show "?P 1" by simp
-  fix n
-  have "fib (n + 2 + 1) = fib (n + 1) + fib (n + 2)"
-    by simp
-  also have "gcd (fib (n + 2)) ... = gcd (fib (n + 2)) (fib (n + 1))"
-    by (simp only: gcd_add2')
-  also have "... = gcd (fib (n + 1)) (fib (n + 1 + 1))"
-    by (simp add: gcd_commute)
-  also assume "... = 1"
-  finally show "?P (n + 2)" .
-qed
-
-lemma gcd_mult_add: "0 < n ==> gcd (n * k + m) n = gcd m n"
-proof -
-  assume "0 < n"
-  then have "gcd (n * k + m) n = gcd n (m mod n)"
-    by (simp add: gcd_non_0 add_commute)
-  also from `0 < n` have "... = gcd m n" by (simp add: gcd_non_0)
-  finally show ?thesis .
-qed
-
-lemma gcd_fib_add: "gcd (fib m) (fib (n + m)) = gcd (fib m) (fib n)"
-proof (cases m)
-  case 0
-  then show ?thesis by simp
-next
-  case (Suc k)
-  then have "gcd (fib m) (fib (n + m)) = gcd (fib (n + k + 1)) (fib (k + 1))"
-    by (simp add: gcd_commute)
-  also have "fib (n + k + 1)
-    = fib (k + 1) * fib (n + 1) + fib k * fib n"
-    by (rule fib_add)
-  also have "gcd ... (fib (k + 1)) = gcd (fib k * fib n) (fib (k + 1))"
-    by (simp add: gcd_mult_add)
-  also have "... = gcd (fib n) (fib (k + 1))"
-    by (simp only: gcd_fib_Suc_eq_1 gcd_mult_cancel)
-  also have "... = gcd (fib m) (fib n)"
-    using Suc by (simp add: gcd_commute)
-  finally show ?thesis .
-qed
-
-lemma gcd_fib_diff:
-  assumes "m <= n"
-  shows "gcd (fib m) (fib (n - m)) = gcd (fib m) (fib n)"
-proof -
-  have "gcd (fib m) (fib (n - m)) = gcd (fib m) (fib (n - m + m))"
-    by (simp add: gcd_fib_add)
-  also from `m <= n` have "n - m + m = n" by simp
-  finally show ?thesis .
-qed
-
-lemma gcd_fib_mod:
-  assumes "0 < m"
-  shows "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
-proof (induct n rule: nat_less_induct)
-  case (1 n) note hyp = this
-  show ?case
-  proof -
-    have "n mod m = (if n < m then n else (n - m) mod m)"
-      by (rule mod_if)
-    also have "gcd (fib m) (fib ...) = gcd (fib m) (fib n)"
-    proof (cases "n < m")
-      case True then show ?thesis by simp
-    next
-      case False then have "m <= n" by simp
-      from `0 < m` and False have "n - m < n" by simp
-      with hyp have "gcd (fib m) (fib ((n - m) mod m))
-        = gcd (fib m) (fib (n - m))" by simp
-      also have "... = gcd (fib m) (fib n)"
-        using `m <= n` by (rule gcd_fib_diff)
-      finally have "gcd (fib m) (fib ((n - m) mod m)) =
-        gcd (fib m) (fib n)" .
-      with False show ?thesis by simp
-    qed
-    finally show ?thesis .
-  qed
-qed
-
-
-theorem fib_gcd: "fib (gcd m n) = gcd (fib m) (fib n)" (is "?P m n")
-proof (induct m n rule: gcd_induct)
-  fix m show "fib (gcd m 0) = gcd (fib m) (fib 0)" by simp
-  fix n :: nat assume n: "0 < n"
-  then have "gcd m n = gcd n (m mod n)" by (rule gcd_non_0)
-  also assume hyp: "fib ... = gcd (fib n) (fib (m mod n))"
-  also from n have "... = gcd (fib n) (fib m)" by (rule gcd_fib_mod)
-  also have "... = gcd (fib m) (fib n)" by (rule gcd_commute)
-  finally show "fib (gcd m n) = gcd (fib m) (fib n)" .
-qed
-
-end
--- a/src/HOL/Isar_examples/Group.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,267 +0,0 @@
-(*  Title:      HOL/Isar_examples/Group.thy
-    Author:     Markus Wenzel, TU Muenchen
-*)
-
-header {* Basic group theory *}
-
-theory Group
-imports Main
-begin
-
-subsection {* Groups and calculational reasoning *} 
-
-text {*
- Groups over signature $({\times} :: \alpha \To \alpha \To \alpha,
- \idt{one} :: \alpha, \idt{inverse} :: \alpha \To \alpha)$ are defined
- as an axiomatic type class as follows.  Note that the parent class
- $\idt{times}$ is provided by the basic HOL theory.
-*}
-
-consts
-  one :: "'a"
-  inverse :: "'a => 'a"
-
-axclass
-  group < times
-  group_assoc:         "(x * y) * z = x * (y * z)"
-  group_left_one:      "one * x = x"
-  group_left_inverse:  "inverse x * x = one"
-
-text {*
- The group axioms only state the properties of left one and inverse,
- the right versions may be derived as follows.
-*}
-
-theorem group_right_inverse: "x * inverse x = (one::'a::group)"
-proof -
-  have "x * inverse x = one * (x * inverse x)"
-    by (simp only: group_left_one)
-  also have "... = one * x * inverse x"
-    by (simp only: group_assoc)
-  also have "... = inverse (inverse x) * inverse x * x * inverse x"
-    by (simp only: group_left_inverse)
-  also have "... = inverse (inverse x) * (inverse x * x) * inverse x"
-    by (simp only: group_assoc)
-  also have "... = inverse (inverse x) * one * inverse x"
-    by (simp only: group_left_inverse)
-  also have "... = inverse (inverse x) * (one * inverse x)"
-    by (simp only: group_assoc)
-  also have "... = inverse (inverse x) * inverse x"
-    by (simp only: group_left_one)
-  also have "... = one"
-    by (simp only: group_left_inverse)
-  finally show ?thesis .
-qed
-
-text {*
- With \name{group-right-inverse} already available,
- \name{group-right-one}\label{thm:group-right-one} is now established
- much easier.
-*}
-
-theorem group_right_one: "x * one = (x::'a::group)"
-proof -
-  have "x * one = x * (inverse x * x)"
-    by (simp only: group_left_inverse)
-  also have "... = x * inverse x * x"
-    by (simp only: group_assoc)
-  also have "... = one * x"
-    by (simp only: group_right_inverse)
-  also have "... = x"
-    by (simp only: group_left_one)
-  finally show ?thesis .
-qed
-
-text {*
- \medskip The calculational proof style above follows typical
- presentations given in any introductory course on algebra.  The basic
- technique is to form a transitive chain of equations, which in turn
- are established by simplifying with appropriate rules.  The low-level
- logical details of equational reasoning are left implicit.
-
- Note that ``$\dots$'' is just a special term variable that is bound
- automatically to the argument\footnote{The argument of a curried
- infix expression happens to be its right-hand side.} of the last fact
- achieved by any local assumption or proven statement.  In contrast to
- $\var{thesis}$, the ``$\dots$'' variable is bound \emph{after} the
- proof is finished, though.
-
- There are only two separate Isar language elements for calculational
- proofs: ``\isakeyword{also}'' for initial or intermediate
- calculational steps, and ``\isakeyword{finally}'' for exhibiting the
- result of a calculation.  These constructs are not hardwired into
- Isabelle/Isar, but defined on top of the basic Isar/VM interpreter.
- Expanding the \isakeyword{also} and \isakeyword{finally} derived
- language elements, calculations may be simulated by hand as
- demonstrated below.
-*}
-
-theorem "x * one = (x::'a::group)"
-proof -
-  have "x * one = x * (inverse x * x)"
-    by (simp only: group_left_inverse)
-
-  note calculation = this
-    -- {* first calculational step: init calculation register *}
-
-  have "... = x * inverse x * x"
-    by (simp only: group_assoc)
-
-  note calculation = trans [OF calculation this]
-    -- {* general calculational step: compose with transitivity rule *}
-
-  have "... = one * x"
-    by (simp only: group_right_inverse)
-
-  note calculation = trans [OF calculation this]
-    -- {* general calculational step: compose with transitivity rule *}
-
-  have "... = x"
-    by (simp only: group_left_one)
-
-  note calculation = trans [OF calculation this]
-    -- {* final calculational step: compose with transitivity rule ... *}
-  from calculation
-    -- {* ... and pick up the final result *}
-
-  show ?thesis .
-qed
-
-text {*
- Note that this scheme of calculations is not restricted to plain
- transitivity.  Rules like anti-symmetry, or even forward and backward
- substitution work as well.  For the actual implementation of
- \isacommand{also} and \isacommand{finally}, Isabelle/Isar maintains
- separate context information of ``transitivity'' rules.  Rule
- selection takes place automatically by higher-order unification.
-*}
-
-
-subsection {* Groups as monoids *}
-
-text {*
- Monoids over signature $({\times} :: \alpha \To \alpha \To \alpha,
- \idt{one} :: \alpha)$ are defined like this.
-*}
-
-axclass monoid < times
-  monoid_assoc:       "(x * y) * z = x * (y * z)"
-  monoid_left_one:   "one * x = x"
-  monoid_right_one:  "x * one = x"
-
-text {*
- Groups are \emph{not} yet monoids directly from the definition.  For
- monoids, \name{right-one} had to be included as an axiom, but for
- groups both \name{right-one} and \name{right-inverse} are derivable
- from the other axioms.  With \name{group-right-one} derived as a
- theorem of group theory (see page~\pageref{thm:group-right-one}), we
- may still instantiate $\idt{group} \subseteq \idt{monoid}$ properly
- as follows.
-*}
-
-instance group < monoid
-  by (intro_classes,
-       rule group_assoc,
-       rule group_left_one,
-       rule group_right_one)
-
-text {*
- The \isacommand{instance} command actually is a version of
- \isacommand{theorem}, setting up a goal that reflects the intended
- class relation (or type constructor arity).  Thus any Isar proof
- language element may be involved to establish this statement.  When
- concluding the proof, the result is transformed into the intended
- type signature extension behind the scenes.
-*}
-
-subsection {* More theorems of group theory *}
-
-text {*
- The one element is already uniquely determined by preserving an
- \emph{arbitrary} group element.
-*}
-
-theorem group_one_equality: "e * x = x ==> one = (e::'a::group)"
-proof -
-  assume eq: "e * x = x"
-  have "one = x * inverse x"
-    by (simp only: group_right_inverse)
-  also have "... = (e * x) * inverse x"
-    by (simp only: eq)
-  also have "... = e * (x * inverse x)"
-    by (simp only: group_assoc)
-  also have "... = e * one"
-    by (simp only: group_right_inverse)
-  also have "... = e"
-    by (simp only: group_right_one)
-  finally show ?thesis .
-qed
-
-text {*
- Likewise, the inverse is already determined by the cancel property.
-*}
-
-theorem group_inverse_equality:
-  "x' * x = one ==> inverse x = (x'::'a::group)"
-proof -
-  assume eq: "x' * x = one"
-  have "inverse x = one * inverse x"
-    by (simp only: group_left_one)
-  also have "... = (x' * x) * inverse x"
-    by (simp only: eq)
-  also have "... = x' * (x * inverse x)"
-    by (simp only: group_assoc)
-  also have "... = x' * one"
-    by (simp only: group_right_inverse)
-  also have "... = x'"
-    by (simp only: group_right_one)
-  finally show ?thesis .
-qed
-
-text {*
- The inverse operation has some further characteristic properties.
-*}
-
-theorem group_inverse_times:
-  "inverse (x * y) = inverse y * inverse (x::'a::group)"
-proof (rule group_inverse_equality)
-  show "(inverse y * inverse x) * (x * y) = one"
-  proof -
-    have "(inverse y * inverse x) * (x * y) =
-        (inverse y * (inverse x * x)) * y"
-      by (simp only: group_assoc)
-    also have "... = (inverse y * one) * y"
-      by (simp only: group_left_inverse)
-    also have "... = inverse y * y"
-      by (simp only: group_right_one)
-    also have "... = one"
-      by (simp only: group_left_inverse)
-    finally show ?thesis .
-  qed
-qed
-
-theorem inverse_inverse: "inverse (inverse x) = (x::'a::group)"
-proof (rule group_inverse_equality)
-  show "x * inverse x = one"
-    by (simp only: group_right_inverse)
-qed
-
-theorem inverse_inject: "inverse x = inverse y ==> x = (y::'a::group)"
-proof -
-  assume eq: "inverse x = inverse y"
-  have "x = x * one"
-    by (simp only: group_right_one)
-  also have "... = x * (inverse y * y)"
-    by (simp only: group_left_inverse)
-  also have "... = x * (inverse x * y)"
-    by (simp only: eq)
-  also have "... = (x * inverse x) * y"
-    by (simp only: group_assoc)
-  also have "... = one * y"
-    by (simp only: group_right_inverse)
-  also have "... = y"
-    by (simp only: group_left_one)
-  finally show ?thesis .
-qed
-
-end
\ No newline at end of file
--- a/src/HOL/Isar_examples/Hoare.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,463 +0,0 @@
-(*  Title:      HOL/Isar_examples/Hoare.thy
-    Author:     Markus Wenzel, TU Muenchen
-
-A formulation of Hoare logic suitable for Isar.
-*)
-
-header {* Hoare Logic *}
-
-theory Hoare
-imports Main
-uses ("~~/src/HOL/Hoare/hoare_tac.ML")
-begin
-
-subsection {* Abstract syntax and semantics *}
-
-text {*
- The following abstract syntax and semantics of Hoare Logic over
- \texttt{WHILE} programs closely follows the existing tradition in
- Isabelle/HOL of formalizing the presentation given in
- \cite[\S6]{Winskel:1993}.  See also
- \url{http://isabelle.in.tum.de/library/Hoare/} and
- \cite{Nipkow:1998:Winskel}.
-*}
-
-types
-  'a bexp = "'a set"
-  'a assn = "'a set"
-
-datatype 'a com =
-    Basic "'a => 'a"
-  | Seq "'a com" "'a com"    ("(_;/ _)" [60, 61] 60)
-  | Cond "'a bexp" "'a com" "'a com"
-  | While "'a bexp" "'a assn" "'a com"
-
-abbreviation
-  Skip  ("SKIP") where
-  "SKIP == Basic id"
-
-types
-  'a sem = "'a => 'a => bool"
-
-consts
-  iter :: "nat => 'a bexp => 'a sem => 'a sem"
-primrec
-  "iter 0 b S s s' = (s ~: b & s = s')"
-  "iter (Suc n) b S s s' =
-    (s : b & (EX s''. S s s'' & iter n b S s'' s'))"
-
-consts
-  Sem :: "'a com => 'a sem"
-primrec
-  "Sem (Basic f) s s' = (s' = f s)"
-  "Sem (c1; c2) s s' = (EX s''. Sem c1 s s'' & Sem c2 s'' s')"
-  "Sem (Cond b c1 c2) s s' =
-    (if s : b then Sem c1 s s' else Sem c2 s s')"
-  "Sem (While b x c) s s' = (EX n. iter n b (Sem c) s s')"
-
-constdefs
-  Valid :: "'a bexp => 'a com => 'a bexp => bool"
-    ("(3|- _/ (2_)/ _)" [100, 55, 100] 50)
-  "|- P c Q == ALL s s'. Sem c s s' --> s : P --> s' : Q"
-
-syntax (xsymbols)
-  Valid :: "'a bexp => 'a com => 'a bexp => bool"
-    ("(3\<turnstile> _/ (2_)/ _)" [100, 55, 100] 50)
-
-lemma ValidI [intro?]:
-    "(!!s s'. Sem c s s' ==> s : P ==> s' : Q) ==> |- P c Q"
-  by (simp add: Valid_def)
-
-lemma ValidD [dest?]:
-    "|- P c Q ==> Sem c s s' ==> s : P ==> s' : Q"
-  by (simp add: Valid_def)
-
-
-subsection {* Primitive Hoare rules *}
-
-text {*
- From the semantics defined above, we derive the standard set of
- primitive Hoare rules; e.g.\ see \cite[\S6]{Winskel:1993}.  Usually,
- variant forms of these rules are applied in actual proof, see also
- \S\ref{sec:hoare-isar} and \S\ref{sec:hoare-vcg}.
-
- \medskip The \name{basic} rule represents any kind of atomic access
- to the state space.  This subsumes the common rules of \name{skip}
- and \name{assign}, as formulated in \S\ref{sec:hoare-isar}.
-*}
-
-theorem basic: "|- {s. f s : P} (Basic f) P"
-proof
-  fix s s' assume s: "s : {s. f s : P}"
-  assume "Sem (Basic f) s s'"
-  hence "s' = f s" by simp
-  with s show "s' : P" by simp
-qed
-
-text {*
- The rules for sequential commands and semantic consequences are
- established in a straight forward manner as follows.
-*}
-
-theorem seq: "|- P c1 Q ==> |- Q c2 R ==> |- P (c1; c2) R"
-proof
-  assume cmd1: "|- P c1 Q" and cmd2: "|- Q c2 R"
-  fix s s' assume s: "s : P"
-  assume "Sem (c1; c2) s s'"
-  then obtain s'' where sem1: "Sem c1 s s''" and sem2: "Sem c2 s'' s'"
-    by auto
-  from cmd1 sem1 s have "s'' : Q" ..
-  with cmd2 sem2 show "s' : R" ..
-qed
-
-theorem conseq: "P' <= P ==> |- P c Q ==> Q <= Q' ==> |- P' c Q'"
-proof
-  assume P'P: "P' <= P" and QQ': "Q <= Q'"
-  assume cmd: "|- P c Q"
-  fix s s' :: 'a
-  assume sem: "Sem c s s'"
-  assume "s : P'" with P'P have "s : P" ..
-  with cmd sem have "s' : Q" ..
-  with QQ' show "s' : Q'" ..
-qed
-
-text {*
- The rule for conditional commands is directly reflected by the
- corresponding semantics; in the proof we just have to look closely
- which cases apply.
-*}
-
-theorem cond:
-  "|- (P Int b) c1 Q ==> |- (P Int -b) c2 Q ==> |- P (Cond b c1 c2) Q"
-proof
-  assume case_b: "|- (P Int b) c1 Q" and case_nb: "|- (P Int -b) c2 Q"
-  fix s s' assume s: "s : P"
-  assume sem: "Sem (Cond b c1 c2) s s'"
-  show "s' : Q"
-  proof cases
-    assume b: "s : b"
-    from case_b show ?thesis
-    proof
-      from sem b show "Sem c1 s s'" by simp
-      from s b show "s : P Int b" by simp
-    qed
-  next
-    assume nb: "s ~: b"
-    from case_nb show ?thesis
-    proof
-      from sem nb show "Sem c2 s s'" by simp
-      from s nb show "s : P Int -b" by simp
-    qed
-  qed
-qed
-
-text {*
- The \name{while} rule is slightly less trivial --- it is the only one
- based on recursion, which is expressed in the semantics by a
- Kleene-style least fixed-point construction.  The auxiliary statement
- below, which is by induction on the number of iterations is the main
- point to be proven; the rest is by routine application of the
- semantics of \texttt{WHILE}.
-*}
-
-theorem while:
-  assumes body: "|- (P Int b) c P"
-  shows "|- P (While b X c) (P Int -b)"
-proof
-  fix s s' assume s: "s : P"
-  assume "Sem (While b X c) s s'"
-  then obtain n where "iter n b (Sem c) s s'" by auto
-  from this and s show "s' : P Int -b"
-  proof (induct n arbitrary: s)
-    case 0
-    thus ?case by auto
-  next
-    case (Suc n)
-    then obtain s'' where b: "s : b" and sem: "Sem c s s''"
-      and iter: "iter n b (Sem c) s'' s'"
-      by auto
-    from Suc and b have "s : P Int b" by simp
-    with body sem have "s'' : P" ..
-    with iter show ?case by (rule Suc)
-  qed
-qed
-
-
-subsection {* Concrete syntax for assertions *}
-
-text {*
- We now introduce concrete syntax for describing commands (with
- embedded expressions) and assertions. The basic technique is that of
- semantic ``quote-antiquote''.  A \emph{quotation} is a syntactic
- entity delimited by an implicit abstraction, say over the state
- space.  An \emph{antiquotation} is a marked expression within a
- quotation that refers the implicit argument; a typical antiquotation
- would select (or even update) components from the state.
-
- We will see some examples later in the concrete rules and
- applications.
-*}
-
-text {*
- The following specification of syntax and translations is for
- Isabelle experts only; feel free to ignore it.
-
- While the first part is still a somewhat intelligible specification
- of the concrete syntactic representation of our Hoare language, the
- actual ``ML drivers'' is quite involved.  Just note that the we
- re-use the basic quote/antiquote translations as already defined in
- Isabelle/Pure (see \verb,Syntax.quote_tr, and
- \verb,Syntax.quote_tr',).
-*}
-
-syntax
-  "_quote"       :: "'b => ('a => 'b)"       ("(.'(_').)" [0] 1000)
-  "_antiquote"   :: "('a => 'b) => 'b"       ("\<acute>_" [1000] 1000)
-  "_Subst"       :: "'a bexp \<Rightarrow> 'b \<Rightarrow> idt \<Rightarrow> 'a bexp"
-        ("_[_'/\<acute>_]" [1000] 999)
-  "_Assert"      :: "'a => 'a set"           ("(.{_}.)" [0] 1000)
-  "_Assign"      :: "idt => 'b => 'a com"    ("(\<acute>_ :=/ _)" [70, 65] 61)
-  "_Cond"        :: "'a bexp => 'a com => 'a com => 'a com"
-        ("(0IF _/ THEN _/ ELSE _/ FI)" [0, 0, 0] 61)
-  "_While_inv"   :: "'a bexp => 'a assn => 'a com => 'a com"
-        ("(0WHILE _/ INV _ //DO _ /OD)"  [0, 0, 0] 61)
-  "_While"       :: "'a bexp => 'a com => 'a com"
-        ("(0WHILE _ //DO _ /OD)"  [0, 0] 61)
-
-syntax (xsymbols)
-  "_Assert"      :: "'a => 'a set"            ("(\<lbrace>_\<rbrace>)" [0] 1000)
-
-translations
-  ".{b}."                   => "Collect .(b)."
-  "B [a/\<acute>x]"                => ".{\<acute>(_update_name x (\<lambda>_. a)) \<in> B}."
-  "\<acute>x := a"                 => "Basic .(\<acute>(_update_name x (\<lambda>_. a)))."
-  "IF b THEN c1 ELSE c2 FI" => "Cond .{b}. c1 c2"
-  "WHILE b INV i DO c OD"   => "While .{b}. i c"
-  "WHILE b DO c OD"         == "WHILE b INV CONST undefined DO c OD"
-
-parse_translation {*
-  let
-    fun quote_tr [t] = Syntax.quote_tr "_antiquote" t
-      | quote_tr ts = raise TERM ("quote_tr", ts);
-  in [("_quote", quote_tr)] end
-*}
-
-text {*
- As usual in Isabelle syntax translations, the part for printing is
- more complicated --- we cannot express parts as macro rules as above.
- Don't look here, unless you have to do similar things for yourself.
-*}
-
-print_translation {*
-  let
-    fun quote_tr' f (t :: ts) =
-          Term.list_comb (f $ Syntax.quote_tr' "_antiquote" t, ts)
-      | quote_tr' _ _ = raise Match;
-
-    val assert_tr' = quote_tr' (Syntax.const "_Assert");
-
-    fun bexp_tr' name ((Const ("Collect", _) $ t) :: ts) =
-          quote_tr' (Syntax.const name) (t :: ts)
-      | bexp_tr' _ _ = raise Match;
-
-    fun upd_tr' (x_upd, T) =
-      (case try (unsuffix Record.updateN) x_upd of
-        SOME x => (x, if T = dummyT then T else Term.domain_type T)
-      | NONE => raise Match);
-
-    fun update_name_tr' (Free x) = Free (upd_tr' x)
-      | update_name_tr' ((c as Const ("_free", _)) $ Free x) =
-          c $ Free (upd_tr' x)
-      | update_name_tr' (Const x) = Const (upd_tr' x)
-      | update_name_tr' _ = raise Match;
-
-    fun K_tr' (Abs (_,_,t)) = if null (loose_bnos t) then t else raise Match
-      | K_tr' (Abs (_,_,Abs (_,_,t)$Bound 0)) = if null (loose_bnos t) then t else raise Match
-      | K_tr' _ = raise Match;
-
-    fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) =
-          quote_tr' (Syntax.const "_Assign" $ update_name_tr' f)
-            (Abs (x, dummyT, K_tr' k) :: ts)
-      | assign_tr' _ = raise Match;
-  in
-    [("Collect", assert_tr'), ("Basic", assign_tr'),
-      ("Cond", bexp_tr' "_Cond"), ("While", bexp_tr' "_While_inv")]
-  end
-*}
-
-
-subsection {* Rules for single-step proof \label{sec:hoare-isar} *}
-
-text {*
- We are now ready to introduce a set of Hoare rules to be used in
- single-step structured proofs in Isabelle/Isar.  We refer to the
- concrete syntax introduce above.
-
- \medskip Assertions of Hoare Logic may be manipulated in
- calculational proofs, with the inclusion expressed in terms of sets
- or predicates.  Reversed order is supported as well.
-*}
-
-lemma [trans]: "|- P c Q ==> P' <= P ==> |- P' c Q"
-  by (unfold Valid_def) blast
-lemma [trans] : "P' <= P ==> |- P c Q ==> |- P' c Q"
-  by (unfold Valid_def) blast
-
-lemma [trans]: "Q <= Q' ==> |- P c Q ==> |- P c Q'"
-  by (unfold Valid_def) blast
-lemma [trans]: "|- P c Q ==> Q <= Q' ==> |- P c Q'"
-  by (unfold Valid_def) blast
-
-lemma [trans]:
-    "|- .{\<acute>P}. c Q ==> (!!s. P' s --> P s) ==> |- .{\<acute>P'}. c Q"
-  by (simp add: Valid_def)
-lemma [trans]:
-    "(!!s. P' s --> P s) ==> |- .{\<acute>P}. c Q ==> |- .{\<acute>P'}. c Q"
-  by (simp add: Valid_def)
-
-lemma [trans]:
-    "|- P c .{\<acute>Q}. ==> (!!s. Q s --> Q' s) ==> |- P c .{\<acute>Q'}."
-  by (simp add: Valid_def)
-lemma [trans]:
-    "(!!s. Q s --> Q' s) ==> |- P c .{\<acute>Q}. ==> |- P c .{\<acute>Q'}."
-  by (simp add: Valid_def)
-
-
-text {*
- Identity and basic assignments.\footnote{The $\idt{hoare}$ method
- introduced in \S\ref{sec:hoare-vcg} is able to provide proper
- instances for any number of basic assignments, without producing
- additional verification conditions.}
-*}
-
-lemma skip [intro?]: "|- P SKIP P"
-proof -
-  have "|- {s. id s : P} SKIP P" by (rule basic)
-  thus ?thesis by simp
-qed
-
-lemma assign: "|- P [\<acute>a/\<acute>x] \<acute>x := \<acute>a P"
-  by (rule basic)
-
-text {*
- Note that above formulation of assignment corresponds to our
- preferred way to model state spaces, using (extensible) record types
- in HOL \cite{Naraschewski-Wenzel:1998:HOOL}.  For any record field
- $x$, Isabelle/HOL provides a functions $x$ (selector) and
- $\idt{x{\dsh}update}$ (update).  Above, there is only a place-holder
- appearing for the latter kind of function: due to concrete syntax
- \isa{\'x := \'a} also contains \isa{x\_update}.\footnote{Note that due
- to the external nature of HOL record fields, we could not even state
- a general theorem relating selector and update functions (if this
- were required here); this would only work for any particular instance
- of record fields introduced so far.}
-*}
-
-text {*
- Sequential composition --- normalizing with associativity achieves
- proper of chunks of code verified separately.
-*}
-
-lemmas [trans, intro?] = seq
-
-lemma seq_assoc [simp]: "( |- P c1;(c2;c3) Q) = ( |- P (c1;c2);c3 Q)"
-  by (auto simp add: Valid_def)
-
-text {*
- Conditional statements.
-*}
-
-lemmas [trans, intro?] = cond
-
-lemma [trans, intro?]:
-  "|- .{\<acute>P & \<acute>b}. c1 Q
-      ==> |- .{\<acute>P & ~ \<acute>b}. c2 Q
-      ==> |- .{\<acute>P}. IF \<acute>b THEN c1 ELSE c2 FI Q"
-    by (rule cond) (simp_all add: Valid_def)
-
-text {*
- While statements --- with optional invariant.
-*}
-
-lemma [intro?]:
-    "|- (P Int b) c P ==> |- P (While b P c) (P Int -b)"
-  by (rule while)
-
-lemma [intro?]:
-    "|- (P Int b) c P ==> |- P (While b undefined c) (P Int -b)"
-  by (rule while)
-
-
-lemma [intro?]:
-  "|- .{\<acute>P & \<acute>b}. c .{\<acute>P}.
-    ==> |- .{\<acute>P}. WHILE \<acute>b INV .{\<acute>P}. DO c OD .{\<acute>P & ~ \<acute>b}."
-  by (simp add: while Collect_conj_eq Collect_neg_eq)
-
-lemma [intro?]:
-  "|- .{\<acute>P & \<acute>b}. c .{\<acute>P}.
-    ==> |- .{\<acute>P}. WHILE \<acute>b DO c OD .{\<acute>P & ~ \<acute>b}."
-  by (simp add: while Collect_conj_eq Collect_neg_eq)
-
-
-subsection {* Verification conditions \label{sec:hoare-vcg} *}
-
-text {*
- We now load the \emph{original} ML file for proof scripts and tactic
- definition for the Hoare Verification Condition Generator (see
- \url{http://isabelle.in.tum.de/library/Hoare/}).  As far as we are
- concerned here, the result is a proof method \name{hoare}, which may
- be applied to a Hoare Logic assertion to extract purely logical
- verification conditions.  It is important to note that the method
- requires \texttt{WHILE} loops to be fully annotated with invariants
- beforehand.  Furthermore, only \emph{concrete} pieces of code are
- handled --- the underlying tactic fails ungracefully if supplied with
- meta-variables or parameters, for example.
-*}
-
-lemma SkipRule: "p \<subseteq> q \<Longrightarrow> Valid p (Basic id) q"
-  by (auto simp add: Valid_def)
-
-lemma BasicRule: "p \<subseteq> {s. f s \<in> q} \<Longrightarrow> Valid p (Basic f) q"
-  by (auto simp: Valid_def)
-
-lemma SeqRule: "Valid P c1 Q \<Longrightarrow> Valid Q c2 R \<Longrightarrow> Valid P (c1;c2) R"
-  by (auto simp: Valid_def)
-
-lemma CondRule:
-  "p \<subseteq> {s. (s \<in> b \<longrightarrow> s \<in> w) \<and> (s \<notin> b \<longrightarrow> s \<in> w')}
-    \<Longrightarrow> Valid w c1 q \<Longrightarrow> Valid w' c2 q \<Longrightarrow> Valid p (Cond b c1 c2) q"
-  by (auto simp: Valid_def)
-
-lemma iter_aux:
-  "\<forall>s s'. Sem c s s' --> s : I & s : b --> s' : I ==>
-       (\<And>s s'. s : I \<Longrightarrow> iter n b (Sem c) s s' \<Longrightarrow> s' : I & s' ~: b)"
-  apply(induct n)
-   apply clarsimp
-   apply (simp (no_asm_use))
-   apply blast
-  done
-
-lemma WhileRule:
-    "p \<subseteq> i \<Longrightarrow> Valid (i \<inter> b) c i \<Longrightarrow> i \<inter> (-b) \<subseteq> q \<Longrightarrow> Valid p (While b i c) q"
-  apply (clarsimp simp: Valid_def)
-  apply (drule iter_aux)
-    prefer 2
-    apply assumption
-   apply blast
-  apply blast
-  done
-
-lemma Compl_Collect: "- Collect b = {x. \<not> b x}"
-  by blast
-
-lemmas AbortRule = SkipRule  -- "dummy version"
-
-use "~~/src/HOL/Hoare/hoare_tac.ML"
-
-method_setup hoare = {*
-  Scan.succeed (fn ctxt =>
-    (SIMPLE_METHOD'
-       (hoare_tac ctxt (simp_tac (HOL_basic_ss addsimps [@{thm "Record.K_record_comp"}] ))))) *}
-  "verification condition generator for Hoare logic"
-
-end
--- a/src/HOL/Isar_examples/Hoare_Ex.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,329 +0,0 @@
-header {* Using Hoare Logic *}
-
-theory Hoare_Ex
-imports Hoare
-begin
-
-subsection {* State spaces *}
-
-text {*
- First of all we provide a store of program variables that
- occur in any of the programs considered later.  Slightly unexpected
- things may happen when attempting to work with undeclared variables.
-*}
-
-record vars =
-  I :: nat
-  M :: nat
-  N :: nat
-  S :: nat
-
-text {*
- While all of our variables happen to have the same type, nothing
- would prevent us from working with many-sorted programs as well, or
- even polymorphic ones.  Also note that Isabelle/HOL's extensible
- record types even provides simple means to extend the state space
- later.
-*}
-
-
-subsection {* Basic examples *}
-
-text {*
- We look at few trivialities involving assignment and sequential
- composition, in order to get an idea of how to work with our
- formulation of Hoare Logic.
-*}
-
-text {*
- Using the basic \name{assign} rule directly is a bit cumbersome.
-*}
-
-lemma
-  "|- .{\<acute>(N_update (\<lambda>_. (2 * \<acute>N))) : .{\<acute>N = 10}.}. \<acute>N := 2 * \<acute>N .{\<acute>N = 10}."
-  by (rule assign)
-
-text {*
- Certainly we want the state modification already done, e.g.\ by
- simplification.  The \name{hoare} method performs the basic state
- update for us; we may apply the Simplifier afterwards to achieve
- ``obvious'' consequences as well.
-*}
-
-lemma "|- .{True}. \<acute>N := 10 .{\<acute>N = 10}."
-  by hoare
-
-lemma "|- .{2 * \<acute>N = 10}. \<acute>N := 2 * \<acute>N .{\<acute>N = 10}."
-  by hoare
-
-lemma "|- .{\<acute>N = 5}. \<acute>N := 2 * \<acute>N .{\<acute>N = 10}."
-  by hoare simp
-
-lemma "|- .{\<acute>N + 1 = a + 1}. \<acute>N := \<acute>N + 1 .{\<acute>N = a + 1}."
-  by hoare
-
-lemma "|- .{\<acute>N = a}. \<acute>N := \<acute>N + 1 .{\<acute>N = a + 1}."
-  by hoare simp
-
-lemma "|- .{a = a & b = b}. \<acute>M := a; \<acute>N := b .{\<acute>M = a & \<acute>N = b}."
-  by hoare
-
-lemma "|- .{True}. \<acute>M := a; \<acute>N := b .{\<acute>M = a & \<acute>N = b}."
-  by hoare simp
-
-lemma
-"|- .{\<acute>M = a & \<acute>N = b}.
-    \<acute>I := \<acute>M; \<acute>M := \<acute>N; \<acute>N := \<acute>I
-    .{\<acute>M = b & \<acute>N = a}."
-  by hoare simp
-
-text {*
- It is important to note that statements like the following one can
- only be proven for each individual program variable.  Due to the
- extra-logical nature of record fields, we cannot formulate a theorem
- relating record selectors and updates schematically.
-*}
-
-lemma "|- .{\<acute>N = a}. \<acute>N := \<acute>N .{\<acute>N = a}."
-  by hoare
-
-lemma "|- .{\<acute>x = a}. \<acute>x := \<acute>x .{\<acute>x = a}."
-  oops
-
-lemma
-  "Valid {s. x s = a} (Basic (\<lambda>s. x_update (x s) s)) {s. x s = n}"
-  -- {* same statement without concrete syntax *}
-  oops
-
-
-text {*
- In the following assignments we make use of the consequence rule in
- order to achieve the intended precondition.  Certainly, the
- \name{hoare} method is able to handle this case, too.
-*}
-
-lemma "|- .{\<acute>M = \<acute>N}. \<acute>M := \<acute>M + 1 .{\<acute>M ~= \<acute>N}."
-proof -
-  have ".{\<acute>M = \<acute>N}. <= .{\<acute>M + 1 ~= \<acute>N}."
-    by auto
-  also have "|- ... \<acute>M := \<acute>M + 1 .{\<acute>M ~= \<acute>N}."
-    by hoare
-  finally show ?thesis .
-qed
-
-lemma "|- .{\<acute>M = \<acute>N}. \<acute>M := \<acute>M + 1 .{\<acute>M ~= \<acute>N}."
-proof -
-  have "!!m n::nat. m = n --> m + 1 ~= n"
-      -- {* inclusion of assertions expressed in ``pure'' logic, *}
-      -- {* without mentioning the state space *}
-    by simp
-  also have "|- .{\<acute>M + 1 ~= \<acute>N}. \<acute>M := \<acute>M + 1 .{\<acute>M ~= \<acute>N}."
-    by hoare
-  finally show ?thesis .
-qed
-
-lemma "|- .{\<acute>M = \<acute>N}. \<acute>M := \<acute>M + 1 .{\<acute>M ~= \<acute>N}."
-  by hoare simp
-
-
-subsection {* Multiplication by addition *}
-
-text {*
- We now do some basic examples of actual \texttt{WHILE} programs.
- This one is a loop for calculating the product of two natural
- numbers, by iterated addition.  We first give detailed structured
- proof based on single-step Hoare rules.
-*}
-
-lemma
-  "|- .{\<acute>M = 0 & \<acute>S = 0}.
-      WHILE \<acute>M ~= a
-      DO \<acute>S := \<acute>S + b; \<acute>M := \<acute>M + 1 OD
-      .{\<acute>S = a * b}."
-proof -
-  let "|- _ ?while _" = ?thesis
-  let ".{\<acute>?inv}." = ".{\<acute>S = \<acute>M * b}."
-
-  have ".{\<acute>M = 0 & \<acute>S = 0}. <= .{\<acute>?inv}." by auto
-  also have "|- ... ?while .{\<acute>?inv & ~ (\<acute>M ~= a)}."
-  proof
-    let ?c = "\<acute>S := \<acute>S + b; \<acute>M := \<acute>M + 1"
-    have ".{\<acute>?inv & \<acute>M ~= a}. <= .{\<acute>S + b = (\<acute>M + 1) * b}."
-      by auto
-    also have "|- ... ?c .{\<acute>?inv}." by hoare
-    finally show "|- .{\<acute>?inv & \<acute>M ~= a}. ?c .{\<acute>?inv}." .
-  qed
-  also have "... <= .{\<acute>S = a * b}." by auto
-  finally show ?thesis .
-qed
-
-text {*
- The subsequent version of the proof applies the \name{hoare} method
- to reduce the Hoare statement to a purely logical problem that can be
- solved fully automatically.  Note that we have to specify the
- \texttt{WHILE} loop invariant in the original statement.
-*}
-
-lemma
-  "|- .{\<acute>M = 0 & \<acute>S = 0}.
-      WHILE \<acute>M ~= a
-      INV .{\<acute>S = \<acute>M * b}.
-      DO \<acute>S := \<acute>S + b; \<acute>M := \<acute>M + 1 OD
-      .{\<acute>S = a * b}."
-  by hoare auto
-
-
-subsection {* Summing natural numbers *}
-
-text {*
- We verify an imperative program to sum natural numbers up to a given
- limit.  First some functional definition for proper specification of
- the problem.
-*}
-
-text {*
- The following proof is quite explicit in the individual steps taken,
- with the \name{hoare} method only applied locally to take care of
- assignment and sequential composition.  Note that we express
- intermediate proof obligation in pure logic, without referring to the
- state space.
-*}
-
-declare atLeast0LessThan[symmetric,simp]
-
-theorem
-  "|- .{True}.
-      \<acute>S := 0; \<acute>I := 1;
-      WHILE \<acute>I ~= n
-      DO
-        \<acute>S := \<acute>S + \<acute>I;
-        \<acute>I := \<acute>I + 1
-      OD
-      .{\<acute>S = (SUM j<n. j)}."
-  (is "|- _ (_; ?while) _")
-proof -
-  let ?sum = "\<lambda>k::nat. SUM j<k. j"
-  let ?inv = "\<lambda>s i::nat. s = ?sum i"
-
-  have "|- .{True}. \<acute>S := 0; \<acute>I := 1 .{?inv \<acute>S \<acute>I}."
-  proof -
-    have "True --> 0 = ?sum 1"
-      by simp
-    also have "|- .{...}. \<acute>S := 0; \<acute>I := 1 .{?inv \<acute>S \<acute>I}."
-      by hoare
-    finally show ?thesis .
-  qed
-  also have "|- ... ?while .{?inv \<acute>S \<acute>I & ~ \<acute>I ~= n}."
-  proof
-    let ?body = "\<acute>S := \<acute>S + \<acute>I; \<acute>I := \<acute>I + 1"
-    have "!!s i. ?inv s i & i ~= n -->  ?inv (s + i) (i + 1)"
-      by simp
-    also have "|- .{\<acute>S + \<acute>I = ?sum (\<acute>I + 1)}. ?body .{?inv \<acute>S \<acute>I}."
-      by hoare
-    finally show "|- .{?inv \<acute>S \<acute>I & \<acute>I ~= n}. ?body .{?inv \<acute>S \<acute>I}." .
-  qed
-  also have "!!s i. s = ?sum i & ~ i ~= n --> s = ?sum n"
-    by simp
-  finally show ?thesis .
-qed
-
-text {*
- The next version uses the \name{hoare} method, while still explaining
- the resulting proof obligations in an abstract, structured manner.
-*}
-
-theorem
-  "|- .{True}.
-      \<acute>S := 0; \<acute>I := 1;
-      WHILE \<acute>I ~= n
-      INV .{\<acute>S = (SUM j<\<acute>I. j)}.
-      DO
-        \<acute>S := \<acute>S + \<acute>I;
-        \<acute>I := \<acute>I + 1
-      OD
-      .{\<acute>S = (SUM j<n. j)}."
-proof -
-  let ?sum = "\<lambda>k::nat. SUM j<k. j"
-  let ?inv = "\<lambda>s i::nat. s = ?sum i"
-
-  show ?thesis
-  proof hoare
-    show "?inv 0 1" by simp
-  next
-    fix s i assume "?inv s i & i ~= n"
-    thus "?inv (s + i) (i + 1)" by simp
-  next
-    fix s i assume "?inv s i & ~ i ~= n"
-    thus "s = ?sum n" by simp
-  qed
-qed
-
-text {*
- Certainly, this proof may be done fully automatic as well, provided
- that the invariant is given beforehand.
-*}
-
-theorem
-  "|- .{True}.
-      \<acute>S := 0; \<acute>I := 1;
-      WHILE \<acute>I ~= n
-      INV .{\<acute>S = (SUM j<\<acute>I. j)}.
-      DO
-        \<acute>S := \<acute>S + \<acute>I;
-        \<acute>I := \<acute>I + 1
-      OD
-      .{\<acute>S = (SUM j<n. j)}."
-  by hoare auto
-
-
-subsection{* Time *}
-
-text{*
-  A simple embedding of time in Hoare logic: function @{text timeit}
-  inserts an extra variable to keep track of the elapsed time.
-*}
-
-record tstate = time :: nat
-
-types 'a time = "\<lparr>time :: nat, \<dots> :: 'a\<rparr>"
-
-consts timeit :: "'a time com \<Rightarrow> 'a time com"
-primrec
-  "timeit (Basic f) = (Basic f; Basic(\<lambda>s. s\<lparr>time := Suc (time s)\<rparr>))"
-  "timeit (c1; c2) = (timeit c1; timeit c2)"
-  "timeit (Cond b c1 c2) = Cond b (timeit c1) (timeit c2)"
-  "timeit (While b iv c) = While b iv (timeit c)"
-
-record tvars = tstate +
-  I :: nat
-  J :: nat
-
-lemma lem: "(0::nat) < n \<Longrightarrow> n + n \<le> Suc (n * n)"
-  by (induct n) simp_all
-
-lemma "|- .{i = \<acute>I & \<acute>time = 0}.
- timeit(
- WHILE \<acute>I \<noteq> 0
- INV .{2*\<acute>time + \<acute>I*\<acute>I + 5*\<acute>I = i*i + 5*i}.
- DO
-   \<acute>J := \<acute>I;
-   WHILE \<acute>J \<noteq> 0
-   INV .{0 < \<acute>I & 2*\<acute>time + \<acute>I*\<acute>I + 3*\<acute>I + 2*\<acute>J - 2 = i*i + 5*i}.
-   DO \<acute>J := \<acute>J - 1 OD;
-   \<acute>I := \<acute>I - 1
- OD
- ) .{2*\<acute>time = i*i + 5*i}."
-  apply simp
-  apply hoare
-      apply simp
-     apply clarsimp
-    apply clarsimp
-   apply arith
-   prefer 2
-   apply clarsimp
-  apply (clarsimp simp: nat_distrib)
-  apply (frule lem)
-  apply arith
-  done
-
-end
--- a/src/HOL/Isar_examples/Knaster_Tarski.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,111 +0,0 @@
-(*  Title:      HOL/Isar_examples/Knaster_Tarski.thy
-    Author:     Markus Wenzel, TU Muenchen
-
-Typical textbook proof example.
-*)
-
-header {* Textbook-style reasoning: the Knaster-Tarski Theorem *}
-
-theory Knaster_Tarski
-imports Main Lattice_Syntax
-begin
-
-
-subsection {* Prose version *}
-
-text {*
-  According to the textbook \cite[pages 93--94]{davey-priestley}, the
-  Knaster-Tarski fixpoint theorem is as follows.\footnote{We have
-  dualized the argument, and tuned the notation a little bit.}
-
-  \textbf{The Knaster-Tarski Fixpoint Theorem.}  Let @{text L} be a
-  complete lattice and @{text "f: L \<rightarrow> L"} an order-preserving map.
-  Then @{text "\<Sqinter>{x \<in> L | f(x) \<le> x}"} is a fixpoint of @{text f}.
-
-  \textbf{Proof.} Let @{text "H = {x \<in> L | f(x) \<le> x}"} and @{text "a =
-  \<Sqinter>H"}.  For all @{text "x \<in> H"} we have @{text "a \<le> x"}, so @{text
-  "f(a) \<le> f(x) \<le> x"}.  Thus @{text "f(a)"} is a lower bound of @{text
-  H}, whence @{text "f(a) \<le> a"}.  We now use this inequality to prove
-  the reverse one (!) and thereby complete the proof that @{text a} is
-  a fixpoint.  Since @{text f} is order-preserving, @{text "f(f(a)) \<le>
-  f(a)"}.  This says @{text "f(a) \<in> H"}, so @{text "a \<le> f(a)"}.
-*}
-
-
-subsection {* Formal versions *}
-
-text {*
-  The Isar proof below closely follows the original presentation.
-  Virtually all of the prose narration has been rephrased in terms of
-  formal Isar language elements.  Just as many textbook-style proofs,
-  there is a strong bias towards forward proof, and several bends in
-  the course of reasoning.
-*}
-
-theorem Knaster_Tarski:
-  fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
-  assumes "mono f"
-  shows "\<exists>a. f a = a"
-proof
-  let ?H = "{u. f u \<le> u}"
-  let ?a = "\<Sqinter>?H"
-  show "f ?a = ?a"
-  proof -
-    {
-      fix x
-      assume "x \<in> ?H"
-      then have "?a \<le> x" by (rule Inf_lower)
-      with `mono f` have "f ?a \<le> f x" ..
-      also from `x \<in> ?H` have "\<dots> \<le> x" ..
-      finally have "f ?a \<le> x" .
-    }
-    then have "f ?a \<le> ?a" by (rule Inf_greatest)
-    {
-      also presume "\<dots> \<le> f ?a"
-      finally (order_antisym) show ?thesis .
-    }
-    from `mono f` and `f ?a \<le> ?a` have "f (f ?a) \<le> f ?a" ..
-    then have "f ?a \<in> ?H" ..
-    then show "?a \<le> f ?a" by (rule Inf_lower)
-  qed
-qed
-
-text {*
-  Above we have used several advanced Isar language elements, such as
-  explicit block structure and weak assumptions.  Thus we have
-  mimicked the particular way of reasoning of the original text.
-
-  In the subsequent version the order of reasoning is changed to
-  achieve structured top-down decomposition of the problem at the
-  outer level, while only the inner steps of reasoning are done in a
-  forward manner.  We are certainly more at ease here, requiring only
-  the most basic features of the Isar language.
-*}
-
-theorem Knaster_Tarski':
-  fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
-  assumes "mono f"
-  shows "\<exists>a. f a = a"
-proof
-  let ?H = "{u. f u \<le> u}"
-  let ?a = "\<Sqinter>?H"
-  show "f ?a = ?a"
-  proof (rule order_antisym)
-    show "f ?a \<le> ?a"
-    proof (rule Inf_greatest)
-      fix x
-      assume "x \<in> ?H"
-      then have "?a \<le> x" by (rule Inf_lower)
-      with `mono f` have "f ?a \<le> f x" ..
-      also from `x \<in> ?H` have "\<dots> \<le> x" ..
-      finally show "f ?a \<le> x" .
-    qed
-    show "?a \<le> f ?a"
-    proof (rule Inf_lower)
-      from `mono f` and `f ?a \<le> ?a` have "f (f ?a) \<le> f ?a" ..
-      then show "f ?a \<in> ?H" ..
-    qed
-  qed
-qed
-
-end
--- a/src/HOL/Isar_examples/Mutilated_Checkerboard.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,300 +0,0 @@
-(*  Title:      HOL/Isar_examples/Mutilated_Checkerboard.thy
-    Author:     Markus Wenzel, TU Muenchen (Isar document)
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory (original scripts)
-*)
-
-header {* The Mutilated Checker Board Problem *}
-
-theory Mutilated_Checkerboard
-imports Main
-begin
-
-text {*
- The Mutilated Checker Board Problem, formalized inductively.  See
- \cite{paulson-mutilated-board} and
- \url{http://isabelle.in.tum.de/library/HOL/Induct/Mutil.html} for the
- original tactic script version.
-*}
-
-subsection {* Tilings *}
-
-inductive_set
-  tiling :: "'a set set => 'a set set"
-  for A :: "'a set set"
-  where
-    empty: "{} : tiling A"
-  | Un: "a : A ==> t : tiling A ==> a <= - t ==> a Un t : tiling A"
-
-
-text "The union of two disjoint tilings is a tiling."
-
-lemma tiling_Un:
-  assumes "t : tiling A" and "u : tiling A" and "t Int u = {}"
-  shows "t Un u : tiling A"
-proof -
-  let ?T = "tiling A"
-  from `t : ?T` and `t Int u = {}`
-  show "t Un u : ?T"
-  proof (induct t)
-    case empty
-    with `u : ?T` show "{} Un u : ?T" by simp
-  next
-    case (Un a t)
-    show "(a Un t) Un u : ?T"
-    proof -
-      have "a Un (t Un u) : ?T"
-        using `a : A`
-      proof (rule tiling.Un)
-        from `(a Un t) Int u = {}` have "t Int u = {}" by blast
-        then show "t Un u: ?T" by (rule Un)
-        from `a <= - t` and `(a Un t) Int u = {}`
-        show "a <= - (t Un u)" by blast
-      qed
-      also have "a Un (t Un u) = (a Un t) Un u"
-        by (simp only: Un_assoc)
-      finally show ?thesis .
-    qed
-  qed
-qed
-
-
-subsection {* Basic properties of ``below'' *}
-
-constdefs
-  below :: "nat => nat set"
-  "below n == {i. i < n}"
-
-lemma below_less_iff [iff]: "(i: below k) = (i < k)"
-  by (simp add: below_def)
-
-lemma below_0: "below 0 = {}"
-  by (simp add: below_def)
-
-lemma Sigma_Suc1:
-    "m = n + 1 ==> below m <*> B = ({n} <*> B) Un (below n <*> B)"
-  by (simp add: below_def less_Suc_eq) blast
-
-lemma Sigma_Suc2:
-    "m = n + 2 ==> A <*> below m =
-      (A <*> {n}) Un (A <*> {n + 1}) Un (A <*> below n)"
-  by (auto simp add: below_def)
-
-lemmas Sigma_Suc = Sigma_Suc1 Sigma_Suc2
-
-
-subsection {* Basic properties of ``evnodd'' *}
-
-constdefs
-  evnodd :: "(nat * nat) set => nat => (nat * nat) set"
-  "evnodd A b == A Int {(i, j). (i + j) mod 2 = b}"
-
-lemma evnodd_iff:
-    "(i, j): evnodd A b = ((i, j): A  & (i + j) mod 2 = b)"
-  by (simp add: evnodd_def)
-
-lemma evnodd_subset: "evnodd A b <= A"
-  by (unfold evnodd_def, rule Int_lower1)
-
-lemma evnoddD: "x : evnodd A b ==> x : A"
-  by (rule subsetD, rule evnodd_subset)
-
-lemma evnodd_finite: "finite A ==> finite (evnodd A b)"
-  by (rule finite_subset, rule evnodd_subset)
-
-lemma evnodd_Un: "evnodd (A Un B) b = evnodd A b Un evnodd B b"
-  by (unfold evnodd_def) blast
-
-lemma evnodd_Diff: "evnodd (A - B) b = evnodd A b - evnodd B b"
-  by (unfold evnodd_def) blast
-
-lemma evnodd_empty: "evnodd {} b = {}"
-  by (simp add: evnodd_def)
-
-lemma evnodd_insert: "evnodd (insert (i, j) C) b =
-    (if (i + j) mod 2 = b
-      then insert (i, j) (evnodd C b) else evnodd C b)"
-  by (simp add: evnodd_def)
-
-
-subsection {* Dominoes *}
-
-inductive_set
-  domino :: "(nat * nat) set set"
-  where
-    horiz: "{(i, j), (i, j + 1)} : domino"
-  | vertl: "{(i, j), (i + 1, j)} : domino"
-
-lemma dominoes_tile_row:
-  "{i} <*> below (2 * n) : tiling domino"
-  (is "?B n : ?T")
-proof (induct n)
-  case 0
-  show ?case by (simp add: below_0 tiling.empty)
-next
-  case (Suc n)
-  let ?a = "{i} <*> {2 * n + 1} Un {i} <*> {2 * n}"
-  have "?B (Suc n) = ?a Un ?B n"
-    by (auto simp add: Sigma_Suc Un_assoc)
-  moreover have "... : ?T"
-  proof (rule tiling.Un)
-    have "{(i, 2 * n), (i, 2 * n + 1)} : domino"
-      by (rule domino.horiz)
-    also have "{(i, 2 * n), (i, 2 * n + 1)} = ?a" by blast
-    finally show "... : domino" .
-    show "?B n : ?T" by (rule Suc)
-    show "?a <= - ?B n" by blast
-  qed
-  ultimately show ?case by simp
-qed
-
-lemma dominoes_tile_matrix:
-  "below m <*> below (2 * n) : tiling domino"
-  (is "?B m : ?T")
-proof (induct m)
-  case 0
-  show ?case by (simp add: below_0 tiling.empty)
-next
-  case (Suc m)
-  let ?t = "{m} <*> below (2 * n)"
-  have "?B (Suc m) = ?t Un ?B m" by (simp add: Sigma_Suc)
-  moreover have "... : ?T"
-  proof (rule tiling_Un)
-    show "?t : ?T" by (rule dominoes_tile_row)
-    show "?B m : ?T" by (rule Suc)
-    show "?t Int ?B m = {}" by blast
-  qed
-  ultimately show ?case by simp
-qed
-
-lemma domino_singleton:
-  assumes d: "d : domino" and "b < 2"
-  shows "EX i j. evnodd d b = {(i, j)}"  (is "?P d")
-  using d
-proof induct
-  from `b < 2` have b_cases: "b = 0 | b = 1" by arith
-  fix i j
-  note [simp] = evnodd_empty evnodd_insert mod_Suc
-  from b_cases show "?P {(i, j), (i, j + 1)}" by rule auto
-  from b_cases show "?P {(i, j), (i + 1, j)}" by rule auto
-qed
-
-lemma domino_finite:
-  assumes d: "d: domino"
-  shows "finite d"
-  using d
-proof induct
-  fix i j :: nat
-  show "finite {(i, j), (i, j + 1)}" by (intro finite.intros)
-  show "finite {(i, j), (i + 1, j)}" by (intro finite.intros)
-qed
-
-
-subsection {* Tilings of dominoes *}
-
-lemma tiling_domino_finite:
-  assumes t: "t : tiling domino"  (is "t : ?T")
-  shows "finite t"  (is "?F t")
-  using t
-proof induct
-  show "?F {}" by (rule finite.emptyI)
-  fix a t assume "?F t"
-  assume "a : domino" then have "?F a" by (rule domino_finite)
-  from this and `?F t` show "?F (a Un t)" by (rule finite_UnI)
-qed
-
-lemma tiling_domino_01:
-  assumes t: "t : tiling domino"  (is "t : ?T")
-  shows "card (evnodd t 0) = card (evnodd t 1)"
-  using t
-proof induct
-  case empty
-  show ?case by (simp add: evnodd_def)
-next
-  case (Un a t)
-  let ?e = evnodd
-  note hyp = `card (?e t 0) = card (?e t 1)`
-    and at = `a <= - t`
-  have card_suc:
-    "!!b. b < 2 ==> card (?e (a Un t) b) = Suc (card (?e t b))"
-  proof -
-    fix b :: nat assume "b < 2"
-    have "?e (a Un t) b = ?e a b Un ?e t b" by (rule evnodd_Un)
-    also obtain i j where e: "?e a b = {(i, j)}"
-    proof -
-      from `a \<in> domino` and `b < 2`
-      have "EX i j. ?e a b = {(i, j)}" by (rule domino_singleton)
-      then show ?thesis by (blast intro: that)
-    qed
-    moreover have "... Un ?e t b = insert (i, j) (?e t b)" by simp
-    moreover have "card ... = Suc (card (?e t b))"
-    proof (rule card_insert_disjoint)
-      from `t \<in> tiling domino` have "finite t"
-        by (rule tiling_domino_finite)
-      then show "finite (?e t b)"
-        by (rule evnodd_finite)
-      from e have "(i, j) : ?e a b" by simp
-      with at show "(i, j) ~: ?e t b" by (blast dest: evnoddD)
-    qed
-    ultimately show "?thesis b" by simp
-  qed
-  then have "card (?e (a Un t) 0) = Suc (card (?e t 0))" by simp
-  also from hyp have "card (?e t 0) = card (?e t 1)" .
-  also from card_suc have "Suc ... = card (?e (a Un t) 1)"
-    by simp
-  finally show ?case .
-qed
-
-
-subsection {* Main theorem *}
-
-constdefs
-  mutilated_board :: "nat => nat => (nat * nat) set"
-  "mutilated_board m n ==
-    below (2 * (m + 1)) <*> below (2 * (n + 1))
-      - {(0, 0)} - {(2 * m + 1, 2 * n + 1)}"
-
-theorem mutil_not_tiling: "mutilated_board m n ~: tiling domino"
-proof (unfold mutilated_board_def)
-  let ?T = "tiling domino"
-  let ?t = "below (2 * (m + 1)) <*> below (2 * (n + 1))"
-  let ?t' = "?t - {(0, 0)}"
-  let ?t'' = "?t' - {(2 * m + 1, 2 * n + 1)}"
-
-  show "?t'' ~: ?T"
-  proof
-    have t: "?t : ?T" by (rule dominoes_tile_matrix)
-    assume t'': "?t'' : ?T"
-
-    let ?e = evnodd
-    have fin: "finite (?e ?t 0)"
-      by (rule evnodd_finite, rule tiling_domino_finite, rule t)
-
-    note [simp] = evnodd_iff evnodd_empty evnodd_insert evnodd_Diff
-    have "card (?e ?t'' 0) < card (?e ?t' 0)"
-    proof -
-      have "card (?e ?t' 0 - {(2 * m + 1, 2 * n + 1)})
-        < card (?e ?t' 0)"
-      proof (rule card_Diff1_less)
-        from _ fin show "finite (?e ?t' 0)"
-          by (rule finite_subset) auto
-        show "(2 * m + 1, 2 * n + 1) : ?e ?t' 0" by simp
-      qed
-      then show ?thesis by simp
-    qed
-    also have "... < card (?e ?t 0)"
-    proof -
-      have "(0, 0) : ?e ?t 0" by simp
-      with fin have "card (?e ?t 0 - {(0, 0)}) < card (?e ?t 0)"
-        by (rule card_Diff1_less)
-      then show ?thesis by simp
-    qed
-    also from t have "... = card (?e ?t 1)"
-      by (rule tiling_domino_01)
-    also have "?e ?t 1 = ?e ?t'' 1" by simp
-    also from t'' have "card ... = card (?e ?t'' 0)"
-      by (rule tiling_domino_01 [symmetric])
-    finally have "... < ..." . then show False ..
-  qed
-qed
-
-end
--- a/src/HOL/Isar_examples/Nested_Datatype.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,86 +0,0 @@
-header {* Nested datatypes *}
-
-theory Nested_Datatype
-imports Main
-begin
-
-subsection {* Terms and substitution *}
-
-datatype ('a, 'b) "term" =
-    Var 'a
-  | App 'b "('a, 'b) term list"
-
-consts
-  subst_term :: "('a => ('a, 'b) term) => ('a, 'b) term => ('a, 'b) term"
-  subst_term_list ::
-    "('a => ('a, 'b) term) => ('a, 'b) term list => ('a, 'b) term list"
-
-primrec (subst)
-  "subst_term f (Var a) = f a"
-  "subst_term f (App b ts) = App b (subst_term_list f ts)"
-  "subst_term_list f [] = []"
-  "subst_term_list f (t # ts) = subst_term f t # subst_term_list f ts"
-
-
-text {*
- \medskip A simple lemma about composition of substitutions.
-*}
-
-lemma "subst_term (subst_term f1 o f2) t =
-      subst_term f1 (subst_term f2 t)"
-  and "subst_term_list (subst_term f1 o f2) ts =
-      subst_term_list f1 (subst_term_list f2 ts)"
-  by (induct t and ts) simp_all
-
-lemma "subst_term (subst_term f1 o f2) t =
-  subst_term f1 (subst_term f2 t)"
-proof -
-  let "?P t" = ?thesis
-  let ?Q = "\<lambda>ts. subst_term_list (subst_term f1 o f2) ts =
-    subst_term_list f1 (subst_term_list f2 ts)"
-  show ?thesis
-  proof (induct t)
-    fix a show "?P (Var a)" by simp
-  next
-    fix b ts assume "?Q ts"
-    then show "?P (App b ts)"
-      by (simp only: subst.simps)
-  next
-    show "?Q []" by simp
-  next
-    fix t ts
-    assume "?P t" "?Q ts" then show "?Q (t # ts)"
-      by (simp only: subst.simps)
-  qed
-qed
-
-
-subsection {* Alternative induction *}
-
-theorem term_induct' [case_names Var App]:
-  assumes var: "!!a. P (Var a)"
-    and app: "!!b ts. list_all P ts ==> P (App b ts)"
-  shows "P t"
-proof (induct t)
-  fix a show "P (Var a)" by (rule var)
-next
-  fix b t ts assume "list_all P ts"
-  then show "P (App b ts)" by (rule app)
-next
-  show "list_all P []" by simp
-next
-  fix t ts assume "P t" "list_all P ts"
-  then show "list_all P (t # ts)" by simp
-qed
-
-lemma
-  "subst_term (subst_term f1 o f2) t = subst_term f1 (subst_term f2 t)"
-proof (induct t rule: term_induct')
-  case (Var a)
-  show ?case by (simp add: o_def)
-next
-  case (App b ts)
-  then show ?case by (induct ts) simp_all
-qed
-
-end
--- a/src/HOL/Isar_examples/Peirce.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,90 +0,0 @@
-(*  Title:      HOL/Isar_examples/Peirce.thy
-    Author:     Markus Wenzel, TU Muenchen
-*)
-
-header {* Peirce's Law *}
-
-theory Peirce
-imports Main
-begin
-
-text {*
- We consider Peirce's Law: $((A \impl B) \impl A) \impl A$.  This is
- an inherently non-intuitionistic statement, so its proof will
- certainly involve some form of classical contradiction.
-
- The first proof is again a well-balanced combination of plain
- backward and forward reasoning.  The actual classical step is where
- the negated goal may be introduced as additional assumption.  This
- eventually leads to a contradiction.\footnote{The rule involved there
- is negation elimination; it holds in intuitionistic logic as well.}
-*}
-
-theorem "((A --> B) --> A) --> A"
-proof
-  assume "(A --> B) --> A"
-  show A
-  proof (rule classical)
-    assume "~ A"
-    have "A --> B"
-    proof
-      assume A
-      with `~ A` show B by contradiction
-    qed
-    with `(A --> B) --> A` show A ..
-  qed
-qed
-
-text {*
- In the subsequent version the reasoning is rearranged by means of
- ``weak assumptions'' (as introduced by \isacommand{presume}).  Before
- assuming the negated goal $\neg A$, its intended consequence $A \impl
- B$ is put into place in order to solve the main problem.
- Nevertheless, we do not get anything for free, but have to establish
- $A \impl B$ later on.  The overall effect is that of a logical
- \emph{cut}.
-
- Technically speaking, whenever some goal is solved by
- \isacommand{show} in the context of weak assumptions then the latter
- give rise to new subgoals, which may be established separately.  In
- contrast, strong assumptions (as introduced by \isacommand{assume})
- are solved immediately.
-*}
-
-theorem "((A --> B) --> A) --> A"
-proof
-  assume "(A --> B) --> A"
-  show A
-  proof (rule classical)
-    presume "A --> B"
-    with `(A --> B) --> A` show A ..
-  next
-    assume "~ A"
-    show "A --> B"
-    proof
-      assume A
-      with `~ A` show B by contradiction
-    qed
-  qed
-qed
-
-text {*
- Note that the goals stemming from weak assumptions may be even left
- until qed time, where they get eventually solved ``by assumption'' as
- well.  In that case there is really no fundamental difference between
- the two kinds of assumptions, apart from the order of reducing the
- individual parts of the proof configuration.
-
- Nevertheless, the ``strong'' mode of plain assumptions is quite
- important in practice to achieve robustness of proof text
- interpretation.  By forcing both the conclusion \emph{and} the
- assumptions to unify with the pending goal to be solved, goal
- selection becomes quite deterministic.  For example, decomposition
- with rules of the ``case-analysis'' type usually gives rise to
- several goals that only differ in there local contexts.  With strong
- assumptions these may be still solved in any order in a predictable
- way, while weak ones would quickly lead to great confusion,
- eventually demanding even some backtracking.
-*}
-
-end
--- a/src/HOL/Isar_examples/Puzzle.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,85 +0,0 @@
-header {* An old chestnut *}
-
-theory Puzzle
-imports Main
-begin
-
-text_raw {*
-  \footnote{A question from ``Bundeswettbewerb Mathematik''.  Original
-  pen-and-paper proof due to Herbert Ehler; Isabelle tactic script by
-  Tobias Nipkow.}
-*}
-
-text {*
-  \textbf{Problem.}  Given some function $f\colon \Nat \to \Nat$ such
-  that $f \ap (f \ap n) < f \ap (\idt{Suc} \ap n)$ for all $n$.
-  Demonstrate that $f$ is the identity.
-*}
-
-theorem
-  assumes f_ax: "\<And>n. f (f n) < f (Suc n)"
-  shows "f n = n"
-proof (rule order_antisym)
-  {
-    fix n show "n \<le> f n"
-    proof (induct k \<equiv> "f n" arbitrary: n rule: less_induct)
-      case (less k n)
-      then have hyp: "\<And>m. f m < f n \<Longrightarrow> m \<le> f m" by (simp only:)
-      show "n \<le> f n"
-      proof (cases n)
-        case (Suc m)
-        from f_ax have "f (f m) < f n" by (simp only: Suc)
-        with hyp have "f m \<le> f (f m)" .
-        also from f_ax have "\<dots> < f n" by (simp only: Suc)
-        finally have "f m < f n" .
-        with hyp have "m \<le> f m" .
-        also note `\<dots> < f n`
-        finally have "m < f n" .
-        then have "n \<le> f n" by (simp only: Suc)
-        then show ?thesis .
-      next
-        case 0
-        then show ?thesis by simp
-      qed
-    qed
-  } note ge = this
-
-  {
-    fix m n :: nat
-    assume "m \<le> n"
-    then have "f m \<le> f n"
-    proof (induct n)
-      case 0
-      then have "m = 0" by simp
-      then show ?case by simp
-    next
-      case (Suc n)
-      from Suc.prems show "f m \<le> f (Suc n)"
-      proof (rule le_SucE)
-        assume "m \<le> n"
-        with Suc.hyps have "f m \<le> f n" .
-        also from ge f_ax have "\<dots> < f (Suc n)"
-          by (rule le_less_trans)
-        finally show ?thesis by simp
-      next
-        assume "m = Suc n"
-        then show ?thesis by simp
-      qed
-    qed
-  } note mono = this
-
-  show "f n \<le> n"
-  proof -
-    have "\<not> n < f n"
-    proof
-      assume "n < f n"
-      then have "Suc n \<le> f n" by simp
-      then have "f (Suc n) \<le> f (f n)" by (rule mono)
-      also have "\<dots> < f (Suc n)" by (rule f_ax)
-      finally have "\<dots> < \<dots>" . then show False ..
-    qed
-    then show ?thesis by simp
-  qed
-qed
-
-end
--- a/src/HOL/Isar_examples/README.html	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,21 +0,0 @@
-<!DOCTYPE HTML PUBLIC "-//W3C//DTD HTML 4.01 Transitional//EN" "http://www.w3.org/TR/html4/loose.dtd">
-
-<!-- $Id$ -->
-
-<html>
-
-<head>
-  <meta http-equiv="Content-Type" content="text/html; charset=iso-8859-1">
-  <title>HOL/Isar_examples</title>
-</head>
-
-<body>
-<h1>HOL/Isar_examples</h1>
-
-Isar offers a new high-level proof (and theory) language interface to
-Isabelle.  This directory contains some example Isar documents.  See
-also the included document, or the <a
-href="http://isabelle.in.tum.de/Isar/">Isabelle/Isar page</a> for more
-information.
-</body>
-</html>
--- a/src/HOL/Isar_examples/ROOT.ML	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,22 +0,0 @@
-(*  Title:      HOL/Isar_examples/ROOT.ML
-    Author:     Markus Wenzel, TU Muenchen
-
-Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
-*)
-
-no_document use_thys ["../Old_Number_Theory/Primes", "../Old_Number_Theory/Fibonacci"];
-
-use_thys [
-  "Basic_Logic",
-  "Cantor",
-  "Peirce",
-  "Drinker",
-  "Expr_Compiler",
-  "Group",
-  "Summation",
-  "Knaster_Tarski",
-  "Mutilated_Checkerboard",
-  "Puzzle",
-  "Nested_Datatype",
-  "Hoare_Ex"
-];
--- a/src/HOL/Isar_examples/Summation.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,158 +0,0 @@
-(*  Title:      HOL/Isar_examples/Summation.thy
-    Author:     Markus Wenzel
-*)
-
-header {* Summing natural numbers *}
-
-theory Summation
-imports Main
-begin
-
-text_raw {*
- \footnote{This example is somewhat reminiscent of the
- \url{http://isabelle.in.tum.de/library/HOL/ex/NatSum.html}, which is
- discussed in \cite{isabelle-ref} in the context of permutative
- rewrite rules and ordered rewriting.}
-*}
-
-text {*
- Subsequently, we prove some summation laws of natural numbers
- (including odds, squares, and cubes).  These examples demonstrate how
- plain natural deduction (including induction) may be combined with
- calculational proof.
-*}
-
-
-subsection {* Summation laws *}
-
-text {*
- The sum of natural numbers $0 + \cdots + n$ equals $n \times (n +
- 1)/2$.  Avoiding formal reasoning about division we prove this
- equation multiplied by $2$.
-*}
-
-theorem sum_of_naturals:
-  "2 * (\<Sum>i::nat=0..n. i) = n * (n + 1)"
-  (is "?P n" is "?S n = _")
-proof (induct n)
-  show "?P 0" by simp
-next
-  fix n have "?S (n + 1) = ?S n + 2 * (n + 1)" by simp
-  also assume "?S n = n * (n + 1)"
-  also have "... + 2 * (n + 1) = (n + 1) * (n + 2)" by simp
-  finally show "?P (Suc n)" by simp
-qed
-
-text {*
- The above proof is a typical instance of mathematical induction.  The
- main statement is viewed as some $\var{P} \ap n$ that is split by the
- induction method into base case $\var{P} \ap 0$, and step case
- $\var{P} \ap n \Impl \var{P} \ap (\idt{Suc} \ap n)$ for arbitrary $n$.
-
- The step case is established by a short calculation in forward
- manner.  Starting from the left-hand side $\var{S} \ap (n + 1)$ of
- the thesis, the final result is achieved by transformations involving
- basic arithmetic reasoning (using the Simplifier).  The main point is
- where the induction hypothesis $\var{S} \ap n = n \times (n + 1)$ is
- introduced in order to replace a certain subterm.  So the
- ``transitivity'' rule involved here is actual \emph{substitution}.
- Also note how the occurrence of ``\dots'' in the subsequent step
- documents the position where the right-hand side of the hypothesis
- got filled in.
-
- \medskip A further notable point here is integration of calculations
- with plain natural deduction.  This works so well in Isar for two
- reasons.
- \begin{enumerate}
-
- \item Facts involved in \isakeyword{also}~/ \isakeyword{finally}
- calculational chains may be just anything.  There is nothing special
- about \isakeyword{have}, so the natural deduction element
- \isakeyword{assume} works just as well.
-
- \item There are two \emph{separate} primitives for building natural
- deduction contexts: \isakeyword{fix}~$x$ and \isakeyword{assume}~$A$.
- Thus it is possible to start reasoning with some new ``arbitrary, but
- fixed'' elements before bringing in the actual assumption.  In
- contrast, natural deduction is occasionally formalized with basic
- context elements of the form $x:A$ instead.
-
- \end{enumerate}
-*}
-
-text {*
- \medskip We derive further summation laws for odds, squares, and
- cubes as follows.  The basic technique of induction plus calculation
- is the same as before.
-*}
-
-theorem sum_of_odds:
-  "(\<Sum>i::nat=0..<n. 2 * i + 1) = n^Suc (Suc 0)"
-  (is "?P n" is "?S n = _")
-proof (induct n)
-  show "?P 0" by simp
-next
-  fix n have "?S (n + 1) = ?S n + 2 * n + 1" by simp
-  also assume "?S n = n^Suc (Suc 0)"
-  also have "... + 2 * n + 1 = (n + 1)^Suc (Suc 0)" by simp
-  finally show "?P (Suc n)" by simp
-qed
-
-text {*
- Subsequently we require some additional tweaking of Isabelle built-in
- arithmetic simplifications, such as bringing in distributivity by
- hand.
-*}
-
-lemmas distrib = add_mult_distrib add_mult_distrib2
-
-theorem sum_of_squares:
-  "6 * (\<Sum>i::nat=0..n. i^Suc (Suc 0)) = n * (n + 1) * (2 * n + 1)"
-  (is "?P n" is "?S n = _")
-proof (induct n)
-  show "?P 0" by simp
-next
-  fix n have "?S (n + 1) = ?S n + 6 * (n + 1)^Suc (Suc 0)"
-    by (simp add: distrib)
-  also assume "?S n = n * (n + 1) * (2 * n + 1)"
-  also have "... + 6 * (n + 1)^Suc (Suc 0) =
-    (n + 1) * (n + 2) * (2 * (n + 1) + 1)" by (simp add: distrib)
-  finally show "?P (Suc n)" by simp
-qed
-
-theorem sum_of_cubes:
-  "4 * (\<Sum>i::nat=0..n. i^3) = (n * (n + 1))^Suc (Suc 0)"
-  (is "?P n" is "?S n = _")
-proof (induct n)
-  show "?P 0" by (simp add: power_eq_if)
-next
-  fix n have "?S (n + 1) = ?S n + 4 * (n + 1)^3"
-    by (simp add: power_eq_if distrib)
-  also assume "?S n = (n * (n + 1))^Suc (Suc 0)"
-  also have "... + 4 * (n + 1)^3 = ((n + 1) * ((n + 1) + 1))^Suc (Suc 0)"
-    by (simp add: power_eq_if distrib)
-  finally show "?P (Suc n)" by simp
-qed
-
-text {*
- Comparing these examples with the tactic script version
- \url{http://isabelle.in.tum.de/library/HOL/ex/NatSum.html}, we note
- an important difference of how induction vs.\ simplification is
- applied.  While \cite[\S10]{isabelle-ref} advises for these examples
- that ``induction should not be applied until the goal is in the
- simplest form'' this would be a very bad idea in our setting.
-
- Simplification normalizes all arithmetic expressions involved,
- producing huge intermediate goals.  With applying induction
- afterwards, the Isar proof text would have to reflect the emerging
- configuration by appropriate sub-proofs.  This would result in badly
- structured, low-level technical reasoning, without any good idea of
- the actual point.
-
- \medskip As a general rule of good proof style, automatic methods
- such as $\idt{simp}$ or $\idt{auto}$ should normally be never used as
- initial proof methods, but only as terminal ones, solving certain
- goals completely.
-*}
-
-end
--- a/src/HOL/Isar_examples/document/proof.sty	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,254 +0,0 @@
-%       proof.sty       (Proof Figure Macros)
-%
-%       version 1.0
-%       October 13, 1990
-%       Copyright (C) 1990 Makoto Tatsuta (tatsuta@riec.tohoku.ac.jp)
-%
-% This program is free software; you can redistribute it or modify
-% it under the terms of the GNU General Public License as published by
-% the Free Software Foundation; either versions 1, or (at your option)
-% any later version.
-%
-% This program is distributed in the hope that it will be useful
-% but WITHOUT ANY WARRANTY; without even the implied warranty of
-% MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE.  See the
-% GNU General Public License for more details.
-%
-%       Usage:
-%               In \documentstyle, specify an optional style `proof', say,
-%                       \documentstyle[proof]{article}.
-%
-%       The following macros are available:
-%
-%       In all the following macros, all the arguments such as
-%       <Lowers> and <Uppers> are processed in math mode.
-%
-%       \infer<Lower><Uppers>
-%               draws an inference.
-%
-%               Use & in <Uppers> to delimit upper formulae.
-%               <Uppers> consists more than 0 formulae.
-%
-%               \infer returns \hbox{ ... } or \vbox{ ... } and
-%               sets \@LeftOffset and \@RightOffset globally.
-%
-%       \infer[<Label>]<Lower><Uppers>
-%               draws an inference labeled with <Label>.
-%
-%       \infer*<Lower><Uppers>
-%               draws a many step deduction.
-%
-%       \infer*[<Label>]<Lower><Uppers>
-%               draws a many step deduction labeled with <Label>.
-%
-%       \deduce<Lower><Uppers>
-%               draws an inference without a rule.
-%
-%       \deduce[<Proof>]<Lower><Uppers>
-%               draws a many step deduction with a proof name.
-%
-%       Example:
-%               If you want to write
-%                           B C
-%                          -----
-%                      A     D
-%                     ----------
-%                         E
-%       use
-%               \infer{E}{
-%                       A
-%                       &
-%                       \infer{D}{B & C}
-%               }
-%
-
-%       Style Parameters
-
-\newdimen\inferLineSkip         \inferLineSkip=2pt
-\newdimen\inferLabelSkip        \inferLabelSkip=5pt
-\def\inferTabSkip{\quad}
-
-%       Variables
-
-\newdimen\@LeftOffset   % global
-\newdimen\@RightOffset  % global
-\newdimen\@SavedLeftOffset      % safe from users
-
-\newdimen\UpperWidth
-\newdimen\LowerWidth
-\newdimen\LowerHeight
-\newdimen\UpperLeftOffset
-\newdimen\UpperRightOffset
-\newdimen\UpperCenter
-\newdimen\LowerCenter
-\newdimen\UpperAdjust
-\newdimen\RuleAdjust
-\newdimen\LowerAdjust
-\newdimen\RuleWidth
-\newdimen\HLabelAdjust
-\newdimen\VLabelAdjust
-\newdimen\WidthAdjust
-
-\newbox\@UpperPart
-\newbox\@LowerPart
-\newbox\@LabelPart
-\newbox\ResultBox
-
-%       Flags
-
-\newif\if@inferRule     % whether \@infer draws a rule.
-\newif\if@ReturnLeftOffset      % whether \@infer returns \@LeftOffset.
-\newif\if@MathSaved     % whether inner math mode where \infer or
-                        % \deduce appears.
-
-%       Special Fonts
-
-\def\DeduceSym{\vtop{\baselineskip4\p@ \lineskiplimit\z@
-    \vbox{\hbox{.}\hbox{.}\hbox{.}}\hbox{.}}}
-
-%       Math Save Macros
-%
-%       \@SaveMath is called in the very begining of toplevel macros
-%       which are \infer and \deduce.
-%       \@RestoreMath is called in the very last before toplevel macros end.
-%       Remark \infer and \deduce ends calling \@infer.
-
-%\def\@SaveMath{\@MathSavedfalse \ifmmode \ifinner
-%        \relax $\relax \@MathSavedtrue \fi\fi }
-%
-%\def\@RestoreMath{\if@MathSaved \relax $\relax\fi }
-
-\def\@SaveMath{\relax}
-\def\@RestoreMath{\relax}
-
-
-%       Macros
-
-\def\@ifEmpty#1#2#3{\def\@tempa{\@empty}\def\@tempb{#1}\relax
-        \ifx \@tempa \@tempb #2\else #3\fi }
-
-\def\infer{\@SaveMath \@ifnextchar *{\@inferSteps}{\@inferOneStep}}
-
-\def\@inferOneStep{\@inferRuletrue
-        \@ifnextchar [{\@infer}{\@infer[\@empty]}}
-
-\def\@inferSteps*{\@ifnextchar [{\@@inferSteps}{\@@inferSteps[\@empty]}}
-
-\def\@@inferSteps[#1]{\@deduce{#1}[\DeduceSym]}
-
-\def\deduce{\@SaveMath \@ifnextchar [{\@deduce{\@empty}}
-        {\@inferRulefalse \@infer[\@empty]}}
-
-%       \@deduce<Proof Label>[<Proof>]<Lower><Uppers>
-
-\def\@deduce#1[#2]#3#4{\@inferRulefalse
-        \@infer[\@empty]{#3}{\@SaveMath \@infer[{#1}]{#2}{#4}}}
-
-%       \@infer[<Label>]<Lower><Uppers>
-%               If \@inferRuletrue, draws a rule and <Label> is right to
-%               a rule.
-%               Otherwise, draws no rule and <Label> is right to <Lower>.
-
-\def\@infer[#1]#2#3{\relax
-% Get parameters
-        \if@ReturnLeftOffset \else \@SavedLeftOffset=\@LeftOffset \fi
-        \setbox\@LabelPart=\hbox{$#1$}\relax
-        \setbox\@LowerPart=\hbox{$#2$}\relax
-%
-        \global\@LeftOffset=0pt
-        \setbox\@UpperPart=\vbox{\tabskip=0pt \halign{\relax
-                \global\@RightOffset=0pt \@ReturnLeftOffsettrue $##$&&
-                \inferTabSkip
-                \global\@RightOffset=0pt \@ReturnLeftOffsetfalse $##$\cr
-                #3\cr}}\relax
-%                       Here is a little trick.
-%                       \@ReturnLeftOffsettrue(false) influences on \infer or
-%                       \deduce placed in ## locally
-%                       because of \@SaveMath and \@RestoreMath.
-        \UpperLeftOffset=\@LeftOffset
-        \UpperRightOffset=\@RightOffset
-% Calculate Adjustments
-        \LowerWidth=\wd\@LowerPart
-        \LowerHeight=\ht\@LowerPart
-        \LowerCenter=0.5\LowerWidth
-%
-        \UpperWidth=\wd\@UpperPart \advance\UpperWidth by -\UpperLeftOffset
-        \advance\UpperWidth by -\UpperRightOffset
-        \UpperCenter=\UpperLeftOffset
-        \advance\UpperCenter by 0.5\UpperWidth
-%
-        \ifdim \UpperWidth > \LowerWidth
-                % \UpperCenter > \LowerCenter
-        \UpperAdjust=0pt
-        \RuleAdjust=\UpperLeftOffset
-        \LowerAdjust=\UpperCenter \advance\LowerAdjust by -\LowerCenter
-        \RuleWidth=\UpperWidth
-        \global\@LeftOffset=\LowerAdjust
-%
-        \else   % \UpperWidth <= \LowerWidth
-        \ifdim \UpperCenter > \LowerCenter
-%
-        \UpperAdjust=0pt
-        \RuleAdjust=\UpperCenter \advance\RuleAdjust by -\LowerCenter
-        \LowerAdjust=\RuleAdjust
-        \RuleWidth=\LowerWidth
-        \global\@LeftOffset=\LowerAdjust
-%
-        \else   % \UpperWidth <= \LowerWidth
-                % \UpperCenter <= \LowerCenter
-%
-        \UpperAdjust=\LowerCenter \advance\UpperAdjust by -\UpperCenter
-        \RuleAdjust=0pt
-        \LowerAdjust=0pt
-        \RuleWidth=\LowerWidth
-        \global\@LeftOffset=0pt
-%
-        \fi\fi
-% Make a box
-        \if@inferRule
-%
-        \setbox\ResultBox=\vbox{
-                \moveright \UpperAdjust \box\@UpperPart
-                \nointerlineskip \kern\inferLineSkip
-                \moveright \RuleAdjust \vbox{\hrule width\RuleWidth}\relax
-                \nointerlineskip \kern\inferLineSkip
-                \moveright \LowerAdjust \box\@LowerPart }\relax
-%
-        \@ifEmpty{#1}{}{\relax
-%
-        \HLabelAdjust=\wd\ResultBox     \advance\HLabelAdjust by -\RuleAdjust
-        \advance\HLabelAdjust by -\RuleWidth
-        \WidthAdjust=\HLabelAdjust
-        \advance\WidthAdjust by -\inferLabelSkip
-        \advance\WidthAdjust by -\wd\@LabelPart
-        \ifdim \WidthAdjust < 0pt \WidthAdjust=0pt \fi
-%
-        \VLabelAdjust=\dp\@LabelPart
-        \advance\VLabelAdjust by -\ht\@LabelPart
-        \VLabelAdjust=0.5\VLabelAdjust  \advance\VLabelAdjust by \LowerHeight
-        \advance\VLabelAdjust by \inferLineSkip
-%
-        \setbox\ResultBox=\hbox{\box\ResultBox
-                \kern -\HLabelAdjust \kern\inferLabelSkip
-                \raise\VLabelAdjust \box\@LabelPart \kern\WidthAdjust}\relax
-%
-        }\relax % end @ifEmpty
-%
-        \else % \@inferRulefalse
-%
-        \setbox\ResultBox=\vbox{
-                \moveright \UpperAdjust \box\@UpperPart
-                \nointerlineskip \kern\inferLineSkip
-                \moveright \LowerAdjust \hbox{\unhbox\@LowerPart
-                        \@ifEmpty{#1}{}{\relax
-                        \kern\inferLabelSkip \unhbox\@LabelPart}}}\relax
-        \fi
-%
-        \global\@RightOffset=\wd\ResultBox
-        \global\advance\@RightOffset by -\@LeftOffset
-        \global\advance\@RightOffset by -\LowerWidth
-        \if@ReturnLeftOffset \else \global\@LeftOffset=\@SavedLeftOffset \fi
-%
-        \box\ResultBox
-        \@RestoreMath
-}
--- a/src/HOL/Isar_examples/document/root.bib	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,91 +0,0 @@
-
-@string{CUCL="Comp. Lab., Univ. Camb."}
-@string{CUP="Cambridge University Press"}
-@string{Springer="Springer-Verlag"}
-@string{TUM="TU Munich"}
-
-@Book{Concrete-Math,
-  author = 	 {R. L. Graham and D. E. Knuth and O. Patashnik},
-  title = 	 {Concrete Mathematics},
-  publisher = 	 {Addison-Wesley},
-  year = 	 1989
-}
-
-@InProceedings{Naraschewski-Wenzel:1998:HOOL,
-  author	= {Wolfgang Naraschewski and Markus Wenzel},
-  title		= {Object-Oriented Verification based on Record Subtyping in
-                  {H}igher-{O}rder {L}ogic},
-  crossref      = {tphols98}}
-
-@Article{Nipkow:1998:Winskel,
-  author = 	 {Tobias Nipkow},
-  title = 	 {Winskel is (almost) Right: Towards a Mechanized Semantics Textbook},
-  journal = 	 {Formal Aspects of Computing},
-  year = 	 1998,
-  volume =	 10,
-  pages =	 {171--186}
-}
-
-@InProceedings{Wenzel:1999:TPHOL,
-  author = 	 {Markus Wenzel},
-  title = 	 {{Isar} --- a Generic Interpretative Approach to Readable Formal Proof Documents},
-  crossref =     {tphols99}}
-
-@Book{Winskel:1993,
-  author = 	 {G. Winskel},
-  title = 	 {The Formal Semantics of Programming Languages},
-  publisher = 	 {MIT Press},
-  year = 	 1993
-}
-
-@Book{davey-priestley,
-  author	= {B. A. Davey and H. A. Priestley},
-  title		= {Introduction to Lattices and Order},
-  publisher	= CUP,
-  year		= 1990}
-
-@manual{isabelle-HOL,
-  author	= {Tobias Nipkow and Lawrence C. Paulson and Markus Wenzel},
-  title		= {{Isabelle}'s Logics: {HOL}},
-  institution	= {Institut f\"ur Informatik, Technische Universi\"at
-                  M\"unchen and Computer Laboratory, University of Cambridge}}
-
-@manual{isabelle-intro,
-  author	= {Lawrence C. Paulson},
-  title		= {Introduction to {Isabelle}},
-  institution	= CUCL}
-
-@manual{isabelle-isar-ref,
-  author	= {Markus Wenzel},
-  title		= {The {Isabelle/Isar} Reference Manual},
-  institution	= TUM}
-
-@manual{isabelle-ref,
-  author	= {Lawrence C. Paulson},
-  title		= {The {Isabelle} Reference Manual},
-  institution	= CUCL}
-
-@TechReport{paulson-mutilated-board,
-  author = 	 {Lawrence C. Paulson},
-  title = 	 {A Simple Formalization and Proof for the Mutilated Chess Board},
-  institution =  CUCL,
-  year = 	 1996,
-  number =	 394,
-  note = {\url{http://www.cl.cam.ac.uk/users/lcp/papers/Reports/mutil.pdf}}
-}
-
-@Proceedings{tphols98,
-  title		= {Theorem Proving in Higher Order Logics: {TPHOLs} '98},
-  booktitle	= {Theorem Proving in Higher Order Logics: {TPHOLs} '98},
-  editor	= {Jim Grundy and Malcom Newey},
-  series	= {LNCS},
-  volume        = 1479,
-  year		= 1998}
-
-@Proceedings{tphols99,
-  title		= {Theorem Proving in Higher Order Logics: {TPHOLs} '99},
-  booktitle	= {Theorem Proving in Higher Order Logics: {TPHOLs} '99},
-  editor	= {Bertot, Y. and Dowek, G. and Hirschowitz, A. and
-                  Paulin, C. and Thery, L.},
-  series	= {LNCS 1690},
-  year		= 1999}
--- a/src/HOL/Isar_examples/document/root.tex	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,30 +0,0 @@
-\input{style}
-
-\hyphenation{Isabelle}
-
-\begin{document}
-
-\title{Miscellaneous Isabelle/Isar examples for Higher-Order Logic}
-\author{Markus Wenzel \\ \url{http://www.in.tum.de/~wenzelm/} \\[2ex]
-  With contributions by Gertrud Bauer and Tobias Nipkow}
-\maketitle
-
-\begin{abstract}
-  Isar offers a high-level proof (and theory) language for Isabelle.
-  We give various examples of Isabelle/Isar proof developments,
-  ranging from simple demonstrations of certain language features to a
-  bit more advanced applications.  The ``real'' applications of
-  Isabelle/Isar are found elsewhere.
-\end{abstract}
-
-\tableofcontents
-
-\parindent 0pt \parskip 0.5ex
-
-\input{session}
-
-\nocite{isabelle-isar-ref,Wenzel:1999:TPHOL}
-\bibliographystyle{abbrv}
-\bibliography{root}
-
-\end{document}
--- a/src/HOL/Isar_examples/document/style.tex	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,40 +0,0 @@
-\documentclass[11pt,a4paper]{article}
-\usepackage[only,bigsqcap]{stmaryrd}
-\usepackage{ifthen,proof,amssymb,isabelle,isabellesym}
-\isabellestyle{it}
-\usepackage{pdfsetup}\urlstyle{rm}
-
-\renewcommand{\isamarkupheader}[1]{\section{#1}}
-
-\renewcommand{\isacommand}[1]
-{\ifthenelse{\equal{sorry}{#1}}{$\;$\dummyproof}
-  {\ifthenelse{\equal{oops}{#1}}{$\vdots$}{\isakeyword{#1}}}}
-
-\newcommand{\DUMMYPROOF}{{\langle\idt{proof}\rangle}}
-\newcommand{\dummyproof}{$\DUMMYPROOF$}
-
-\newcommand{\name}[1]{\textsl{#1}}
-
-\newcommand{\idt}[1]{{\mathord{\mathit{#1}}}}
-\newcommand{\var}[1]{{?\!\idt{#1}}}
-\DeclareMathSymbol{\dshsym}{\mathalpha}{letters}{"2D}
-\newcommand{\dsh}{\dshsym}
-
-\newcommand{\To}{\to}
-\newcommand{\dt}{{\mathpunct.}}
-\newcommand{\ap}{\mathbin{\!}}
-\newcommand{\lam}[1]{\mathop{\lambda} #1\dt\;}
-\newcommand{\all}[1]{\forall #1\dt\;}
-\newcommand{\ex}[1]{\exists #1\dt\;}
-\newcommand{\impl}{\to}
-\newcommand{\conj}{\land}
-\newcommand{\disj}{\lor}
-\newcommand{\Impl}{\Longrightarrow}
-
-\newcommand{\Nat}{\mathord{\mathrm{I}\mkern-3.8mu\mathrm{N}}}
-
-
-%%% Local Variables: 
-%%% mode: latex
-%%% TeX-master: "root"
-%%% End: 
--- a/src/HOL/Library/Numeral_Type.thy	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Library/Numeral_Type.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -408,7 +408,7 @@
   in bin_of n end;
 
 fun numeral_tr (*"_NumeralType"*) [Const (str, _)] =
-      mk_bintype (valOf (Int.fromString str))
+      mk_bintype (the (Int.fromString str))
   | numeral_tr (*"_NumeralType"*) ts = raise TERM ("numeral_tr", ts);
 
 in [("_NumeralType", numeral_tr)] end;
--- a/src/HOL/Library/Sum_Of_Squares/positivstellensatz_tools.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Library/Sum_Of_Squares/positivstellensatz_tools.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -97,7 +97,7 @@
 val parse_id = Scan.one Symbol.is_letter ::: Scan.many Symbol.is_letdig >> implode
 
 fun parse_varpow ctxt = parse_id -- Scan.optional (str "^" |-- nat) 1 >>
-  (fn (x, k) => (cterm_of (Context.theory_of_proof ctxt) (Free (x, @{typ real})), k)) 
+  (fn (x, k) => (cterm_of (ProofContext.theory_of ctxt) (Free (x, @{typ real})), k)) 
 
 fun parse_monomial ctxt = repeat_sep "*" (parse_varpow ctxt) >>
   List.foldl (uncurry FuncUtil.Ctermfunc.update) FuncUtil.Ctermfunc.empty
--- a/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -21,8 +21,8 @@
 val rat_2 = Rat.two;
 val rat_10 = Rat.rat_of_int 10;
 val rat_1_2 = rat_1 // rat_2;
-val max = curry Int.max;
-val min = curry Int.min;
+val max = Integer.max;
+val min = Integer.min;
 
 val denominator_rat = Rat.quotient_of_rat #> snd #> Rat.rat_of_int;
 val numerator_rat = Rat.quotient_of_rat #> fst #> Rat.rat_of_int;
@@ -520,7 +520,7 @@
 fun rat_of_quotient (a,b) = if b = 0 then rat_0 else Rat.rat_of_quotient (a,b);
 fun rat_of_string s =
  let val n = index_char s #"/" 0 in
-  if n = ~1 then s |> Int.fromString |> valOf |> Rat.rat_of_int
+  if n = ~1 then s |> Int.fromString |> the |> Rat.rat_of_int
   else
    let val SOME numer = Int.fromString(String.substring(s,0,n))
        val SOME den = Int.fromString (String.substring(s,n+1,String.size s - n - 1))
@@ -1193,7 +1193,7 @@
 fun real_nonlinear_prover proof_method ctxt =
  let
   val {add,mul,neg,pow,sub,main} =  Normalizer.semiring_normalizers_ord_wrapper ctxt
-      (valOf (NormalizerData.match ctxt @{cterm "(0::real) + 1"}))
+      (the (NormalizerData.match ctxt @{cterm "(0::real) + 1"}))
      simple_cterm_ord
   val (real_poly_add_conv,real_poly_mul_conv,real_poly_neg_conv,
        real_poly_pow_conv,real_poly_sub_conv,real_poly_conv) = (add,mul,neg,pow,sub,main)
@@ -1308,7 +1308,7 @@
 fun real_nonlinear_subst_prover prover ctxt =
  let
   val {add,mul,neg,pow,sub,main} =  Normalizer.semiring_normalizers_ord_wrapper ctxt
-      (valOf (NormalizerData.match ctxt @{cterm "(0::real) + 1"}))
+      (the (NormalizerData.match ctxt @{cterm "(0::real) + 1"}))
      simple_cterm_ord
 
   val (real_poly_add_conv,real_poly_mul_conv,real_poly_neg_conv,
--- a/src/HOL/Library/normarith.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Library/normarith.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -167,7 +167,7 @@
   (* FIXME : Should be computed statically!! *)
   val real_poly_conv = 
     Normalizer.semiring_normalize_wrapper ctxt
-     (valOf (NormalizerData.match ctxt @{cterm "(0::real) + 1"}))
+     (the (NormalizerData.match ctxt @{cterm "(0::real) + 1"}))
  in fconv_rule (arg_conv ((rewr_conv @{thm ge_iff_diff_ge_0}) then_conv arg_conv (field_comp_conv then_conv real_poly_conv)))
 end;
 
@@ -278,7 +278,7 @@
    (* FIXME: Should be computed statically!!*)
    val real_poly_conv = 
       Normalizer.semiring_normalize_wrapper ctxt
-       (valOf (NormalizerData.match ctxt @{cterm "(0::real) + 1"}))
+       (the (NormalizerData.match ctxt @{cterm "(0::real) + 1"}))
    val sources = map (Thm.dest_arg o Thm.dest_arg1 o concl) nubs
    val rawdests = fold_rev (find_normedterms o Thm.dest_arg o concl) (ges @ gts) [] 
    val _ = if not (forall fst rawdests) then error "real_vector_combo_prover: Sanity check" 
@@ -384,7 +384,7 @@
   let 
    val real_poly_neg_conv = #neg
        (Normalizer.semiring_normalizers_ord_wrapper ctxt
-        (valOf (NormalizerData.match ctxt @{cterm "(0::real) + 1"})) simple_cterm_ord)
+        (the (NormalizerData.match ctxt @{cterm "(0::real) + 1"})) simple_cterm_ord)
    val (th1,th2) = conj_pair(rawrule th)
   in th1::fconv_rule (arg_conv (arg_conv real_poly_neg_conv)) th2::acc
   end
--- a/src/HOL/Library/positivstellensatz.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Library/positivstellensatz.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -43,7 +43,7 @@
   in Tab.fold h a b end;
 
 fun choose f = case Tab.min_key f of 
-   SOME k => (k,valOf (Tab.lookup f k))
+   SOME k => (k, the (Tab.lookup f k))
  | NONE => error "FuncFun.choose : Completely empty function"
 
 fun onefunc kv = update kv empty
@@ -743,7 +743,7 @@
   fun simple_cterm_ord t u = TermOrd.term_ord (term_of t, term_of u) = LESS
   val {add,mul,neg,pow,sub,main} = 
      Normalizer.semiring_normalizers_ord_wrapper ctxt
-      (valOf (NormalizerData.match ctxt @{cterm "(0::real) + 1"})) 
+      (the (NormalizerData.match ctxt @{cterm "(0::real) + 1"})) 
      simple_cterm_ord
 in gen_real_arith ctxt
    (cterm_of_rat, field_comp_conv, field_comp_conv,field_comp_conv,
--- a/src/HOL/MetisExamples/Abstraction.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,298 +0,0 @@
-(*  Title:      HOL/MetisExamples/Abstraction.thy
-    ID:         $Id$
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-
-Testing the metis method
-*)
-
-theory Abstraction
-imports Main FuncSet
-begin
-
-(*For Christoph Benzmueller*)
-lemma "x<1 & ((op=) = (op=)) ==> ((op=) = (op=)) & (x<(2::nat))";
-  by (metis One_nat_def less_Suc0 not_less0 not_less_eq numeral_2_eq_2)
-
-(*this is a theorem, but we can't prove it unless ext is applied explicitly
-lemma "(op=) = (%x y. y=x)"
-*)
-
-consts
-  monotone :: "['a => 'a, 'a set, ('a *'a)set] => bool"
-  pset  :: "'a set => 'a set"
-  order :: "'a set => ('a * 'a) set"
-
-declare [[ atp_problem_prefix = "Abstraction__Collect_triv" ]]
-lemma (*Collect_triv:*) "a \<in> {x. P x} ==> P a"
-proof (neg_clausify)
-assume 0: "(a\<Colon>'a\<Colon>type) \<in> Collect (P\<Colon>'a\<Colon>type \<Rightarrow> bool)"
-assume 1: "\<not> (P\<Colon>'a\<Colon>type \<Rightarrow> bool) (a\<Colon>'a\<Colon>type)"
-have 2: "(P\<Colon>'a\<Colon>type \<Rightarrow> bool) (a\<Colon>'a\<Colon>type)"
-  by (metis CollectD 0)
-show "False"
-  by (metis 2 1)
-qed
-
-lemma Collect_triv: "a \<in> {x. P x} ==> P a"
-by (metis mem_Collect_eq)
-
-
-declare [[ atp_problem_prefix = "Abstraction__Collect_mp" ]]
-lemma "a \<in> {x. P x --> Q x} ==> a \<in> {x. P x} ==> a \<in> {x. Q x}"
-  by (metis CollectI Collect_imp_eq ComplD UnE mem_Collect_eq);
-  --{*34 secs*}
-
-declare [[ atp_problem_prefix = "Abstraction__Sigma_triv" ]]
-lemma "(a,b) \<in> Sigma A B ==> a \<in> A & b \<in> B a"
-proof (neg_clausify)
-assume 0: "(a\<Colon>'a\<Colon>type, b\<Colon>'b\<Colon>type) \<in> Sigma (A\<Colon>'a\<Colon>type set) (B\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>type set)"
-assume 1: "(a\<Colon>'a\<Colon>type) \<notin> (A\<Colon>'a\<Colon>type set) \<or> (b\<Colon>'b\<Colon>type) \<notin> (B\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>type set) a"
-have 2: "(a\<Colon>'a\<Colon>type) \<in> (A\<Colon>'a\<Colon>type set)"
-  by (metis SigmaD1 0)
-have 3: "(b\<Colon>'b\<Colon>type) \<in> (B\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>type set) (a\<Colon>'a\<Colon>type)"
-  by (metis SigmaD2 0)
-have 4: "(b\<Colon>'b\<Colon>type) \<notin> (B\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>type set) (a\<Colon>'a\<Colon>type)"
-  by (metis 1 2)
-show "False"
-  by (metis 3 4)
-qed
-
-lemma Sigma_triv: "(a,b) \<in> Sigma A B ==> a \<in> A & b \<in> B a"
-by (metis SigmaD1 SigmaD2)
-
-declare [[ atp_problem_prefix = "Abstraction__Sigma_Collect" ]]
-lemma "(a,b) \<in> (SIGMA x: A. {y. x = f y}) ==> a \<in> A & a = f b"
-(*???metis says this is satisfiable!
-by (metis CollectD SigmaD1 SigmaD2)
-*)
-by (meson CollectD SigmaD1 SigmaD2)
-
-
-(*single-step*)
-lemma "(a,b) \<in> (SIGMA x: A. {y. x = f y}) ==> a \<in> A & a = f b"
-by (metis SigmaD1 SigmaD2 insert_def singleton_conv2 Un_empty_right vimage_Collect_eq vimage_def vimage_singleton_eq)
-
-
-lemma "(a,b) \<in> (SIGMA x: A. {y. x = f y}) ==> a \<in> A & a = f b"
-proof (neg_clausify)
-assume 0: "(a\<Colon>'a\<Colon>type, b\<Colon>'b\<Colon>type)
-\<in> Sigma (A\<Colon>'a\<Colon>type set)
-   (COMBB Collect (COMBC (COMBB COMBB op =) (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>type)))"
-assume 1: "(a\<Colon>'a\<Colon>type) \<notin> (A\<Colon>'a\<Colon>type set) \<or> a \<noteq> (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>type) (b\<Colon>'b\<Colon>type)"
-have 2: "(a\<Colon>'a\<Colon>type) \<in> (A\<Colon>'a\<Colon>type set)"
-  by (metis 0 SigmaD1)
-have 3: "(b\<Colon>'b\<Colon>type)
-\<in> COMBB Collect (COMBC (COMBB COMBB op =) (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>type)) (a\<Colon>'a\<Colon>type)"
-  by (metis 0 SigmaD2) 
-have 4: "(b\<Colon>'b\<Colon>type) \<in> Collect (COMBB (op = (a\<Colon>'a\<Colon>type)) (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>type))"
-  by (metis 3)
-have 5: "(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>type) (b\<Colon>'b\<Colon>type) \<noteq> (a\<Colon>'a\<Colon>type)"
-  by (metis 1 2)
-have 6: "(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>type) (b\<Colon>'b\<Colon>type) = (a\<Colon>'a\<Colon>type)"
-  by (metis 4 vimage_singleton_eq insert_def singleton_conv2 Un_empty_right vimage_Collect_eq vimage_def)
-show "False"
-  by (metis 5 6)
-qed
-
-(*Alternative structured proof, untyped*)
-lemma "(a,b) \<in> (SIGMA x: A. {y. x = f y}) ==> a \<in> A & a = f b"
-proof (neg_clausify)
-assume 0: "(a, b) \<in> Sigma A (COMBB Collect (COMBC (COMBB COMBB op =) f))"
-have 1: "b \<in> Collect (COMBB (op = a) f)"
-  by (metis 0 SigmaD2)
-have 2: "f b = a"
-  by (metis 1 vimage_Collect_eq singleton_conv2 insert_def Un_empty_right vimage_singleton_eq vimage_def)
-assume 3: "a \<notin> A \<or> a \<noteq> f b"
-have 4: "a \<in> A"
-  by (metis 0 SigmaD1)
-have 5: "f b \<noteq> a"
-  by (metis 4 3)
-show "False"
-  by (metis 5 2)
-qed
-
-
-declare [[ atp_problem_prefix = "Abstraction__CLF_eq_in_pp" ]]
-lemma "(cl,f) \<in> CLF ==> CLF = (SIGMA cl: CL.{f. f \<in> pset cl}) ==> f \<in> pset cl"
-by (metis Collect_mem_eq SigmaD2)
-
-lemma "(cl,f) \<in> CLF ==> CLF = (SIGMA cl: CL.{f. f \<in> pset cl}) ==> f \<in> pset cl"
-proof (neg_clausify)
-assume 0: "(cl, f) \<in> CLF"
-assume 1: "CLF = Sigma CL (COMBB Collect (COMBB (COMBC op \<in>) pset))"
-assume 2: "f \<notin> pset cl"
-have 3: "\<And>X1 X2. X2 \<in> COMBB Collect (COMBB (COMBC op \<in>) pset) X1 \<or> (X1, X2) \<notin> CLF"
-  by (metis SigmaD2 1)
-have 4: "\<And>X1 X2. X2 \<in> pset X1 \<or> (X1, X2) \<notin> CLF"
-  by (metis 3 Collect_mem_eq)
-have 5: "(cl, f) \<notin> CLF"
-  by (metis 2 4)
-show "False"
-  by (metis 5 0)
-qed
-
-declare [[ atp_problem_prefix = "Abstraction__Sigma_Collect_Pi" ]]
-lemma
-    "(cl,f) \<in> (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl}) ==> 
-    f \<in> pset cl \<rightarrow> pset cl"
-proof (neg_clausify)
-assume 0: "f \<notin> Pi (pset cl) (COMBK (pset cl))"
-assume 1: "(cl, f)
-\<in> Sigma CL
-   (COMBB Collect
-     (COMBB (COMBC op \<in>) (COMBS (COMBB Pi pset) (COMBB COMBK pset))))"
-show "False"
-(*  by (metis 0 Collect_mem_eq SigmaD2 1) ??doesn't terminate*)
-  by (insert 0 1, simp add: COMBB_def COMBS_def COMBC_def)
-qed
-
-
-declare [[ atp_problem_prefix = "Abstraction__Sigma_Collect_Int" ]]
-lemma
-    "(cl,f) \<in> (SIGMA cl: CL. {f. f \<in> pset cl \<inter> cl}) ==>
-   f \<in> pset cl \<inter> cl"
-proof (neg_clausify)
-assume 0: "(cl, f)
-\<in> Sigma CL
-   (COMBB Collect (COMBB (COMBC op \<in>) (COMBS (COMBB op \<inter> pset) COMBI)))"
-assume 1: "f \<notin> pset cl \<inter> cl"
-have 2: "f \<in> COMBB Collect (COMBB (COMBC op \<in>) (COMBS (COMBB op \<inter> pset) COMBI)) cl" 
-  by (insert 0, simp add: COMBB_def) 
-(*  by (metis SigmaD2 0)  ??doesn't terminate*)
-have 3: "f \<in> COMBS (COMBB op \<inter> pset) COMBI cl"
-  by (metis 2 Collect_mem_eq)
-have 4: "f \<notin> cl \<inter> pset cl"
-  by (metis 1 Int_commute)
-have 5: "f \<in> cl \<inter> pset cl"
-  by (metis 3 Int_commute)
-show "False"
-  by (metis 5 4)
-qed
-
-
-declare [[ atp_problem_prefix = "Abstraction__Sigma_Collect_Pi_mono" ]]
-lemma
-    "(cl,f) \<in> (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl & monotone f (pset cl) (order cl)}) ==>
-   (f \<in> pset cl \<rightarrow> pset cl)  &  (monotone f (pset cl) (order cl))"
-by auto
-
-declare [[ atp_problem_prefix = "Abstraction__CLF_subset_Collect_Int" ]]
-lemma "(cl,f) \<in> CLF ==> 
-   CLF \<subseteq> (SIGMA cl: CL. {f. f \<in> pset cl \<inter> cl}) ==>
-   f \<in> pset cl \<inter> cl"
-by auto
-
-(*??no longer terminates, with combinators
-by (metis Collect_mem_eq Int_def SigmaD2 UnCI Un_absorb1)
-  --{*@{text Int_def} is redundant*}
-*)
-
-declare [[ atp_problem_prefix = "Abstraction__CLF_eq_Collect_Int" ]]
-lemma "(cl,f) \<in> CLF ==> 
-   CLF = (SIGMA cl: CL. {f. f \<in> pset cl \<inter> cl}) ==>
-   f \<in> pset cl \<inter> cl"
-by auto
-(*??no longer terminates, with combinators
-by (metis Collect_mem_eq Int_commute SigmaD2)
-*)
-
-declare [[ atp_problem_prefix = "Abstraction__CLF_subset_Collect_Pi" ]]
-lemma 
-   "(cl,f) \<in> CLF ==> 
-    CLF \<subseteq> (SIGMA cl': CL. {f. f \<in> pset cl' \<rightarrow> pset cl'}) ==> 
-    f \<in> pset cl \<rightarrow> pset cl"
-by fast
-(*??no longer terminates, with combinators
-by (metis Collect_mem_eq SigmaD2 subsetD)
-*)
-
-declare [[ atp_problem_prefix = "Abstraction__CLF_eq_Collect_Pi" ]]
-lemma 
-  "(cl,f) \<in> CLF ==> 
-   CLF = (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl}) ==> 
-   f \<in> pset cl \<rightarrow> pset cl"
-by auto
-(*??no longer terminates, with combinators
-by (metis Collect_mem_eq SigmaD2 contra_subsetD equalityE)
-*)
-
-declare [[ atp_problem_prefix = "Abstraction__CLF_eq_Collect_Pi_mono" ]]
-lemma 
-  "(cl,f) \<in> CLF ==> 
-   CLF = (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl & monotone f (pset cl) (order cl)}) ==>
-   (f \<in> pset cl \<rightarrow> pset cl)  &  (monotone f (pset cl) (order cl))"
-by auto
-
-declare [[ atp_problem_prefix = "Abstraction__map_eq_zipA" ]]
-lemma "map (%x. (f x, g x)) xs = zip (map f xs) (map g xs)"
-apply (induct xs)
-(*sledgehammer*)  
-apply auto
-done
-
-declare [[ atp_problem_prefix = "Abstraction__map_eq_zipB" ]]
-lemma "map (%w. (w -> w, w \<times> w)) xs = 
-       zip (map (%w. w -> w) xs) (map (%w. w \<times> w) xs)"
-apply (induct xs)
-(*sledgehammer*)  
-apply auto
-done
-
-declare [[ atp_problem_prefix = "Abstraction__image_evenA" ]]
-lemma "(%x. Suc(f x)) ` {x. even x} <= A ==> (\<forall>x. even x --> Suc(f x) \<in> A)";
-(*sledgehammer*)  
-by auto
-
-declare [[ atp_problem_prefix = "Abstraction__image_evenB" ]]
-lemma "(%x. f (f x)) ` ((%x. Suc(f x)) ` {x. even x}) <= A 
-       ==> (\<forall>x. even x --> f (f (Suc(f x))) \<in> A)";
-(*sledgehammer*)  
-by auto
-
-declare [[ atp_problem_prefix = "Abstraction__image_curry" ]]
-lemma "f \<in> (%u v. b \<times> u \<times> v) ` A ==> \<forall>u v. P (b \<times> u \<times> v) ==> P(f y)" 
-(*sledgehammer*)  
-by auto
-
-declare [[ atp_problem_prefix = "Abstraction__image_TimesA" ]]
-lemma image_TimesA: "(%(x,y). (f x, g y)) ` (A \<times> B) = (f`A) \<times> (g`B)"
-(*sledgehammer*) 
-apply (rule equalityI)
-(***Even the two inclusions are far too difficult
-using [[ atp_problem_prefix = "Abstraction__image_TimesA_simpler"]]
-***)
-apply (rule subsetI)
-apply (erule imageE)
-(*V manages from here with help: Abstraction__image_TimesA_simpler_1_b.p*)
-apply (erule ssubst)
-apply (erule SigmaE)
-(*V manages from here: Abstraction__image_TimesA_simpler_1_a.p*)
-apply (erule ssubst)
-apply (subst split_conv)
-apply (rule SigmaI) 
-apply (erule imageI) +
-txt{*subgoal 2*}
-apply (clarify );
-apply (simp add: );  
-apply (rule rev_image_eqI)  
-apply (blast intro: elim:); 
-apply (simp add: );
-done
-
-(*Given the difficulty of the previous problem, these two are probably
-impossible*)
-
-declare [[ atp_problem_prefix = "Abstraction__image_TimesB" ]]
-lemma image_TimesB:
-    "(%(x,y,z). (f x, g y, h z)) ` (A \<times> B \<times> C) = (f`A) \<times> (g`B) \<times> (h`C)" 
-(*sledgehammer*) 
-by force
-
-declare [[ atp_problem_prefix = "Abstraction__image_TimesC" ]]
-lemma image_TimesC:
-    "(%(x,y). (x \<rightarrow> x, y \<times> y)) ` (A \<times> B) = 
-     ((%x. x \<rightarrow> x) ` A) \<times> ((%y. y \<times> y) ` B)" 
-(*sledgehammer*) 
-by auto
-
-end
--- a/src/HOL/MetisExamples/BT.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,243 +0,0 @@
-(*  Title:      HOL/MetisTest/BT.thy
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-
-Testing the metis method
-*)
-
-header {* Binary trees *}
-
-theory BT
-imports Main
-begin
-
-
-datatype 'a bt =
-    Lf
-  | Br 'a  "'a bt"  "'a bt"
-
-consts
-  n_nodes   :: "'a bt => nat"
-  n_leaves  :: "'a bt => nat"
-  depth     :: "'a bt => nat"
-  reflect   :: "'a bt => 'a bt"
-  bt_map    :: "('a => 'b) => ('a bt => 'b bt)"
-  preorder  :: "'a bt => 'a list"
-  inorder   :: "'a bt => 'a list"
-  postorder :: "'a bt => 'a list"
-  appnd    :: "'a bt => 'a bt => 'a bt"
-
-primrec
-  "n_nodes Lf = 0"
-  "n_nodes (Br a t1 t2) = Suc (n_nodes t1 + n_nodes t2)"
-
-primrec
-  "n_leaves Lf = Suc 0"
-  "n_leaves (Br a t1 t2) = n_leaves t1 + n_leaves t2"
-
-primrec
-  "depth Lf = 0"
-  "depth (Br a t1 t2) = Suc (max (depth t1) (depth t2))"
-
-primrec
-  "reflect Lf = Lf"
-  "reflect (Br a t1 t2) = Br a (reflect t2) (reflect t1)"
-
-primrec
-  "bt_map f Lf = Lf"
-  "bt_map f (Br a t1 t2) = Br (f a) (bt_map f t1) (bt_map f t2)"
-
-primrec
-  "preorder Lf = []"
-  "preorder (Br a t1 t2) = [a] @ (preorder t1) @ (preorder t2)"
-
-primrec
-  "inorder Lf = []"
-  "inorder (Br a t1 t2) = (inorder t1) @ [a] @ (inorder t2)"
-
-primrec
-  "postorder Lf = []"
-  "postorder (Br a t1 t2) = (postorder t1) @ (postorder t2) @ [a]"
-
-primrec
-  "appnd Lf t = t"
-  "appnd (Br a t1 t2) t = Br a (appnd t1 t) (appnd t2 t)"
-
-
-text {* \medskip BT simplification *}
-
-declare [[ atp_problem_prefix = "BT__n_leaves_reflect" ]]
-lemma n_leaves_reflect: "n_leaves (reflect t) = n_leaves t"
-  apply (induct t)
-  apply (metis add_right_cancel n_leaves.simps(1) reflect.simps(1))
-  apply (metis add_commute n_leaves.simps(2) reflect.simps(2))
-  done
-
-declare [[ atp_problem_prefix = "BT__n_nodes_reflect" ]]
-lemma n_nodes_reflect: "n_nodes (reflect t) = n_nodes t"
-  apply (induct t)
-  apply (metis reflect.simps(1))
-  apply (metis n_nodes.simps(2) nat_add_commute reflect.simps(2))
-  done
-
-declare [[ atp_problem_prefix = "BT__depth_reflect" ]]
-lemma depth_reflect: "depth (reflect t) = depth t"
-  apply (induct t)
-  apply (metis depth.simps(1) reflect.simps(1))
-  apply (metis depth.simps(2) min_max.sup_commute reflect.simps(2))
-  done
-
-text {*
-  The famous relationship between the numbers of leaves and nodes.
-*}
-
-declare [[ atp_problem_prefix = "BT__n_leaves_nodes" ]]
-lemma n_leaves_nodes: "n_leaves t = Suc (n_nodes t)"
-  apply (induct t)
-  apply (metis n_leaves.simps(1) n_nodes.simps(1))
-  apply auto
-  done
-
-declare [[ atp_problem_prefix = "BT__reflect_reflect_ident" ]]
-lemma reflect_reflect_ident: "reflect (reflect t) = t"
-  apply (induct t)
-  apply (metis add_right_cancel reflect.simps(1));
-  apply (metis reflect.simps(2))
-  done
-
-declare [[ atp_problem_prefix = "BT__bt_map_ident" ]]
-lemma bt_map_ident: "bt_map (%x. x) = (%y. y)"
-apply (rule ext) 
-apply (induct_tac y)
-  apply (metis bt_map.simps(1))
-txt{*BUG involving flex-flex pairs*}
-(*  apply (metis bt_map.simps(2)) *)
-apply auto
-done
-
-
-declare [[ atp_problem_prefix = "BT__bt_map_appnd" ]]
-lemma bt_map_appnd: "bt_map f (appnd t u) = appnd (bt_map f t) (bt_map f u)"
-apply (induct t)
-  apply (metis appnd.simps(1) bt_map.simps(1))
-  apply (metis appnd.simps(2) bt_map.simps(2))  (*slow!!*)
-done
-
-
-declare [[ atp_problem_prefix = "BT__bt_map_compose" ]]
-lemma bt_map_compose: "bt_map (f o g) t = bt_map f (bt_map g t)"
-apply (induct t) 
-  apply (metis bt_map.simps(1))
-txt{*Metis runs forever*}
-(*  apply (metis bt_map.simps(2) o_apply)*)
-apply auto
-done
-
-
-declare [[ atp_problem_prefix = "BT__bt_map_reflect" ]]
-lemma bt_map_reflect: "bt_map f (reflect t) = reflect (bt_map f t)"
-  apply (induct t)
-  apply (metis add_right_cancel bt_map.simps(1) reflect.simps(1))
-  apply (metis add_right_cancel bt_map.simps(2) reflect.simps(2))
-  done
-
-declare [[ atp_problem_prefix = "BT__preorder_bt_map" ]]
-lemma preorder_bt_map: "preorder (bt_map f t) = map f (preorder t)"
-  apply (induct t)
-  apply (metis bt_map.simps(1) map.simps(1) preorder.simps(1))
-   apply simp
-  done
-
-declare [[ atp_problem_prefix = "BT__inorder_bt_map" ]]
-lemma inorder_bt_map: "inorder (bt_map f t) = map f (inorder t)"
-  apply (induct t)
-  apply (metis bt_map.simps(1) inorder.simps(1) map.simps(1))
-  apply simp
-  done
-
-declare [[ atp_problem_prefix = "BT__postorder_bt_map" ]]
-lemma postorder_bt_map: "postorder (bt_map f t) = map f (postorder t)"
-  apply (induct t)
-  apply (metis bt_map.simps(1) map.simps(1) postorder.simps(1))
-   apply simp
-  done
-
-declare [[ atp_problem_prefix = "BT__depth_bt_map" ]]
-lemma depth_bt_map [simp]: "depth (bt_map f t) = depth t"
-  apply (induct t)
-  apply (metis bt_map.simps(1) depth.simps(1))
-   apply simp
-  done
-
-declare [[ atp_problem_prefix = "BT__n_leaves_bt_map" ]]
-lemma n_leaves_bt_map [simp]: "n_leaves (bt_map f t) = n_leaves t"
-  apply (induct t)
-  apply (metis One_nat_def Suc_eq_plus1 bt_map.simps(1) less_add_one less_antisym linorder_neq_iff n_leaves.simps(1))
-  apply (metis bt_map.simps(2) n_leaves.simps(2))
-  done
-
-
-declare [[ atp_problem_prefix = "BT__preorder_reflect" ]]
-lemma preorder_reflect: "preorder (reflect t) = rev (postorder t)"
-  apply (induct t)
-  apply (metis postorder.simps(1) preorder.simps(1) reflect.simps(1) rev_is_Nil_conv)
-  apply (metis append_Nil Cons_eq_append_conv postorder.simps(2) preorder.simps(2) reflect.simps(2) rev.simps(2) rev_append rev_rev_ident)
-  done
-
-declare [[ atp_problem_prefix = "BT__inorder_reflect" ]]
-lemma inorder_reflect: "inorder (reflect t) = rev (inorder t)"
-  apply (induct t)
-  apply (metis inorder.simps(1) reflect.simps(1) rev.simps(1))
-  apply simp
-  done
-
-declare [[ atp_problem_prefix = "BT__postorder_reflect" ]]
-lemma postorder_reflect: "postorder (reflect t) = rev (preorder t)"
-  apply (induct t)
-  apply (metis postorder.simps(1) preorder.simps(1) reflect.simps(1) rev.simps(1))
-  apply (metis Cons_eq_appendI postorder.simps(2) preorder.simps(2) reflect.simps(2) rev.simps(2) rev_append self_append_conv2)
-  done
-
-text {*
- Analogues of the standard properties of the append function for lists.
-*}
-
-declare [[ atp_problem_prefix = "BT__appnd_assoc" ]]
-lemma appnd_assoc [simp]:
-     "appnd (appnd t1 t2) t3 = appnd t1 (appnd t2 t3)"
-  apply (induct t1)
-  apply (metis appnd.simps(1))
-  apply (metis appnd.simps(2))
-  done
-
-declare [[ atp_problem_prefix = "BT__appnd_Lf2" ]]
-lemma appnd_Lf2 [simp]: "appnd t Lf = t"
-  apply (induct t)
-  apply (metis appnd.simps(1))
-  apply (metis appnd.simps(2))
-  done
-
-declare [[ atp_problem_prefix = "BT__depth_appnd" ]]
-  declare max_add_distrib_left [simp]
-lemma depth_appnd [simp]: "depth (appnd t1 t2) = depth t1 + depth t2"
-  apply (induct t1)
-  apply (metis add_0 appnd.simps(1) depth.simps(1))
-apply (simp add: ); 
-  done
-
-declare [[ atp_problem_prefix = "BT__n_leaves_appnd" ]]
-lemma n_leaves_appnd [simp]:
-     "n_leaves (appnd t1 t2) = n_leaves t1 * n_leaves t2"
-  apply (induct t1)
-  apply (metis One_nat_def appnd.simps(1) less_irrefl less_linear n_leaves.simps(1) nat_mult_1) 
-  apply (simp add: left_distrib)
-  done
-
-declare [[ atp_problem_prefix = "BT__bt_map_appnd" ]]
-lemma (*bt_map_appnd:*)
-     "bt_map f (appnd t1 t2) = appnd (bt_map f t1) (bt_map f t2)"
-  apply (induct t1)
-  apply (metis appnd.simps(1) bt_map_appnd)
-  apply (metis bt_map_appnd)
-  done
-
-end
--- a/src/HOL/MetisExamples/BigO.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1231 +0,0 @@
-(*  Title:      HOL/MetisExamples/BigO.thy
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-
-Testing the metis method
-*)
-
-header {* Big O notation *}
-
-theory BigO
-imports "~~/src/HOL/Decision_Procs/Dense_Linear_Order" Main SetsAndFunctions 
-begin
-
-subsection {* Definitions *}
-
-definition bigo :: "('a => 'b::ordered_idom) => ('a => 'b) set"    ("(1O'(_'))") where
-  "O(f::('a => 'b)) ==   {h. EX c. ALL x. abs (h x) <= c * abs (f x)}"
-
-declare [[ atp_problem_prefix = "BigO__bigo_pos_const" ]]
-lemma bigo_pos_const: "(EX (c::'a::ordered_idom). 
-    ALL x. (abs (h x)) <= (c * (abs (f x))))
-      = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))"
-  apply auto
-  apply (case_tac "c = 0", simp)
-  apply (rule_tac x = "1" in exI, simp)
-  apply (rule_tac x = "abs c" in exI, auto)
-  apply (metis abs_ge_minus_self abs_ge_zero abs_minus_cancel abs_of_nonneg equation_minus_iff Orderings.xt1(6) abs_mult)
-  done
-
-(*** Now various verions with an increasing modulus ***)
-
-declare [[sledgehammer_modulus = 1]]
-
-lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_idom). 
-    ALL x. (abs (h x)) <= (c * (abs (f x))))
-      = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))"
-  apply auto
-  apply (case_tac "c = 0", simp)
-  apply (rule_tac x = "1" in exI, simp)
-  apply (rule_tac x = "abs c" in exI, auto)
-proof (neg_clausify)
-fix c x
-have 0: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. \<bar>X1 * X2\<bar> = \<bar>X2 * X1\<bar>"
-  by (metis abs_mult mult_commute)
-have 1: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom.
-   X1 \<le> (0\<Colon>'a\<Colon>ordered_idom) \<or> \<bar>X2\<bar> * X1 = \<bar>X2 * X1\<bar>"
-  by (metis abs_mult_pos linorder_linear)
-have 2: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom.
-   \<not> (0\<Colon>'a\<Colon>ordered_idom) < X1 * X2 \<or>
-   \<not> (0\<Colon>'a\<Colon>ordered_idom) \<le> X2 \<or> \<not> X1 \<le> (0\<Colon>'a\<Colon>ordered_idom)"
-  by (metis linorder_not_less mult_nonneg_nonpos2)
-assume 3: "\<And>x\<Colon>'b\<Colon>type.
-   \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>
-   \<le> (c\<Colon>'a\<Colon>ordered_idom) * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>"
-assume 4: "\<not> \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar>
-  \<le> \<bar>c\<Colon>'a\<Colon>ordered_idom\<bar> * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>"
-have 5: "\<not> \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar>
-  \<le> \<bar>(c\<Colon>'a\<Colon>ordered_idom) * (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>"
-  by (metis 4 abs_mult)
-have 6: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom.
-   \<not> X1 \<le> (0\<Colon>'a\<Colon>ordered_idom) \<or> X1 \<le> \<bar>X2\<bar>"
-  by (metis abs_ge_zero xt1(6))
-have 7: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom.
-   X1 \<le> \<bar>X2\<bar> \<or> (0\<Colon>'a\<Colon>ordered_idom) < X1"
-  by (metis not_leE 6)
-have 8: "(0\<Colon>'a\<Colon>ordered_idom) < \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar>"
-  by (metis 5 7)
-have 9: "\<And>X1\<Colon>'a\<Colon>ordered_idom.
-   \<not> \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar> \<le> X1 \<or>
-   (0\<Colon>'a\<Colon>ordered_idom) < X1"
-  by (metis 8 order_less_le_trans)
-have 10: "(0\<Colon>'a\<Colon>ordered_idom)
-< (c\<Colon>'a\<Colon>ordered_idom) * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar>"
-  by (metis 3 9)
-have 11: "\<not> (c\<Colon>'a\<Colon>ordered_idom) \<le> (0\<Colon>'a\<Colon>ordered_idom)"
-  by (metis abs_ge_zero 2 10)
-have 12: "\<And>X1\<Colon>'a\<Colon>ordered_idom. (c\<Colon>'a\<Colon>ordered_idom) * \<bar>X1\<bar> = \<bar>X1 * c\<bar>"
-  by (metis mult_commute 1 11)
-have 13: "\<And>X1\<Colon>'b\<Colon>type.
-   - (h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1
-   \<le> (c\<Colon>'a\<Colon>ordered_idom) * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1\<bar>"
-  by (metis 3 abs_le_D2)
-have 14: "\<And>X1\<Colon>'b\<Colon>type.
-   - (h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1
-   \<le> \<bar>(c\<Colon>'a\<Colon>ordered_idom) * (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1\<bar>"
-  by (metis 0 12 13)
-have 15: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. \<bar>X1 * \<bar>X2\<bar>\<bar> = \<bar>X1 * X2\<bar>"
-  by (metis abs_mult abs_mult_pos abs_ge_zero)
-have 16: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. X1 \<le> \<bar>X2\<bar> \<or> \<not> X1 \<le> X2"
-  by (metis xt1(6) abs_ge_self)
-have 17: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. \<not> \<bar>X1\<bar> \<le> X2 \<or> X1 \<le> \<bar>X2\<bar>"
-  by (metis 16 abs_le_D1)
-have 18: "\<And>X1\<Colon>'b\<Colon>type.
-   (h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1
-   \<le> \<bar>(c\<Colon>'a\<Colon>ordered_idom) * (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1\<bar>"
-  by (metis 17 3 15)
-show "False"
-  by (metis abs_le_iff 5 18 14)
-qed
-
-declare [[sledgehammer_modulus = 2]]
-
-lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_idom). 
-    ALL x. (abs (h x)) <= (c * (abs (f x))))
-      = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))"
-  apply auto
-  apply (case_tac "c = 0", simp)
-  apply (rule_tac x = "1" in exI, simp)
-  apply (rule_tac x = "abs c" in exI, auto);
-proof (neg_clausify)
-fix c x
-have 0: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. \<bar>X1 * X2\<bar> = \<bar>X2 * X1\<bar>"
-  by (metis abs_mult mult_commute)
-assume 1: "\<And>x\<Colon>'b\<Colon>type.
-   \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>
-   \<le> (c\<Colon>'a\<Colon>ordered_idom) * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>"
-assume 2: "\<not> \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar>
-  \<le> \<bar>c\<Colon>'a\<Colon>ordered_idom\<bar> * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>"
-have 3: "\<not> \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar>
-  \<le> \<bar>(c\<Colon>'a\<Colon>ordered_idom) * (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>"
-  by (metis 2 abs_mult)
-have 4: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom.
-   \<not> X1 \<le> (0\<Colon>'a\<Colon>ordered_idom) \<or> X1 \<le> \<bar>X2\<bar>"
-  by (metis abs_ge_zero xt1(6))
-have 5: "(0\<Colon>'a\<Colon>ordered_idom) < \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar>"
-  by (metis not_leE 4 3)
-have 6: "(0\<Colon>'a\<Colon>ordered_idom)
-< (c\<Colon>'a\<Colon>ordered_idom) * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar>"
-  by (metis 1 order_less_le_trans 5)
-have 7: "\<And>X1\<Colon>'a\<Colon>ordered_idom. (c\<Colon>'a\<Colon>ordered_idom) * \<bar>X1\<bar> = \<bar>X1 * c\<bar>"
-  by (metis abs_ge_zero linorder_not_less mult_nonneg_nonpos2 6 linorder_linear abs_mult_pos mult_commute)
-have 8: "\<And>X1\<Colon>'b\<Colon>type.
-   - (h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1
-   \<le> \<bar>(c\<Colon>'a\<Colon>ordered_idom) * (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1\<bar>"
-  by (metis 0 7 abs_le_D2 1)
-have 9: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. \<not> \<bar>X1\<bar> \<le> X2 \<or> X1 \<le> \<bar>X2\<bar>"
-  by (metis abs_ge_self xt1(6) abs_le_D1)
-show "False"
-  by (metis 8 abs_ge_zero abs_mult_pos abs_mult 1 9 3 abs_le_iff)
-qed
-
-declare [[sledgehammer_modulus = 3]]
-
-lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_idom). 
-    ALL x. (abs (h x)) <= (c * (abs (f x))))
-      = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))"
-  apply auto
-  apply (case_tac "c = 0", simp)
-  apply (rule_tac x = "1" in exI, simp)
-  apply (rule_tac x = "abs c" in exI, auto);
-proof (neg_clausify)
-fix c x
-assume 0: "\<And>x\<Colon>'b\<Colon>type.
-   \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>
-   \<le> (c\<Colon>'a\<Colon>ordered_idom) * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>"
-assume 1: "\<not> \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar>
-  \<le> \<bar>c\<Colon>'a\<Colon>ordered_idom\<bar> * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>"
-have 2: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom.
-   X1 \<le> \<bar>X2\<bar> \<or> (0\<Colon>'a\<Colon>ordered_idom) < X1"
-  by (metis abs_ge_zero xt1(6) not_leE)
-have 3: "\<not> (c\<Colon>'a\<Colon>ordered_idom) \<le> (0\<Colon>'a\<Colon>ordered_idom)"
-  by (metis abs_ge_zero mult_nonneg_nonpos2 linorder_not_less order_less_le_trans 1 abs_mult 2 0)
-have 4: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. \<bar>X1 * \<bar>X2\<bar>\<bar> = \<bar>X1 * X2\<bar>"
-  by (metis abs_ge_zero abs_mult_pos abs_mult)
-have 5: "\<And>X1\<Colon>'b\<Colon>type.
-   (h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1
-   \<le> \<bar>(c\<Colon>'a\<Colon>ordered_idom) * (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1\<bar>"
-  by (metis 4 0 xt1(6) abs_ge_self abs_le_D1)
-show "False"
-  by (metis abs_mult mult_commute 3 abs_mult_pos linorder_linear 0 abs_le_D2 5 1 abs_le_iff)
-qed
-
-
-declare [[sledgehammer_modulus = 1]]
-
-lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_idom). 
-    ALL x. (abs (h x)) <= (c * (abs (f x))))
-      = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))"
-  apply auto
-  apply (case_tac "c = 0", simp)
-  apply (rule_tac x = "1" in exI, simp)
-  apply (rule_tac x = "abs c" in exI, auto);
-proof (neg_clausify)
-fix c x  (*sort/type constraint inserted by hand!*)
-have 0: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2. \<bar>X1 * \<bar>X2\<bar>\<bar> = \<bar>X1 * X2\<bar>"
-  by (metis abs_ge_zero abs_mult_pos abs_mult)
-assume 1: "\<And>A. \<bar>h A\<bar> \<le> c * \<bar>f A\<bar>"
-have 2: "\<And>X1 X2. \<not> \<bar>X1\<bar> \<le> X2 \<or> (0\<Colon>'a) \<le> X2"
-  by (metis abs_ge_zero order_trans)
-have 3: "\<And>X1. (0\<Colon>'a) \<le> c * \<bar>f X1\<bar>"
-  by (metis 1 2)
-have 4: "\<And>X1. c * \<bar>f X1\<bar> = \<bar>c * f X1\<bar>"
-  by (metis 0 abs_of_nonneg 3)
-have 5: "\<And>X1. - h X1 \<le> c * \<bar>f X1\<bar>"
-  by (metis 1 abs_le_D2)
-have 6: "\<And>X1. - h X1 \<le> \<bar>c * f X1\<bar>"
-  by (metis 4 5)
-have 7: "\<And>X1. h X1 \<le> c * \<bar>f X1\<bar>"
-  by (metis 1 abs_le_D1)
-have 8: "\<And>X1. h X1 \<le> \<bar>c * f X1\<bar>"
-  by (metis 4 7)
-assume 9: "\<not> \<bar>h x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>"
-have 10: "\<not> \<bar>h x\<bar> \<le> \<bar>c * f x\<bar>"
-  by (metis abs_mult 9)
-show "False"
-  by (metis 6 8 10 abs_leI)
-qed
-
-
-declare [[sledgehammer_sorts = true]]
-
-lemma bigo_alt_def: "O(f) = 
-    {h. EX c. (0 < c & (ALL x. abs (h x) <= c * abs (f x)))}"
-by (auto simp add: bigo_def bigo_pos_const)
-
-declare [[ atp_problem_prefix = "BigO__bigo_elt_subset" ]]
-lemma bigo_elt_subset [intro]: "f : O(g) ==> O(f) <= O(g)"
-  apply (auto simp add: bigo_alt_def)
-  apply (rule_tac x = "ca * c" in exI)
-  apply (rule conjI)
-  apply (rule mult_pos_pos)
-  apply (assumption)+ 
-(*sledgehammer*);
-  apply (rule allI)
-  apply (drule_tac x = "xa" in spec)+
-  apply (subgoal_tac "ca * abs(f xa) <= ca * (c * abs(g xa))");
-  apply (erule order_trans)
-  apply (simp add: mult_ac)
-  apply (rule mult_left_mono, assumption)
-  apply (rule order_less_imp_le, assumption);
-done
-
-
-declare [[ atp_problem_prefix = "BigO__bigo_refl" ]]
-lemma bigo_refl [intro]: "f : O(f)"
-  apply(auto simp add: bigo_def)
-proof (neg_clausify)
-fix x
-assume 0: "\<And>xa. \<not> \<bar>f (x xa)\<bar> \<le> xa * \<bar>f (x xa)\<bar>"
-have 1: "\<And>X2. X2 \<le> (1\<Colon>'b) * X2 \<or> \<not> (1\<Colon>'b) \<le> (1\<Colon>'b)"
-  by (metis mult_le_cancel_right1 order_eq_iff)
-have 2: "\<And>X2. X2 \<le> (1\<Colon>'b) * X2"
-  by (metis order_eq_iff 1)
-show "False"
-  by (metis 0 2)
-qed
-
-declare [[ atp_problem_prefix = "BigO__bigo_zero" ]]
-lemma bigo_zero: "0 : O(g)"
-  apply (auto simp add: bigo_def func_zero)
-proof (neg_clausify)
-fix x
-assume 0: "\<And>xa. \<not> (0\<Colon>'b) \<le> xa * \<bar>g (x xa)\<bar>"
-have 1: "\<not> (0\<Colon>'b) \<le> (0\<Colon>'b)"
-  by (metis 0 mult_eq_0_iff)
-show "False"
-  by (metis 1 linorder_neq_iff linorder_antisym_conv1)
-qed
-
-lemma bigo_zero2: "O(%x.0) = {%x.0}"
-  apply (auto simp add: bigo_def) 
-  apply (rule ext)
-  apply auto
-done
-
-lemma bigo_plus_self_subset [intro]: 
-  "O(f) \<oplus> O(f) <= O(f)"
-  apply (auto simp add: bigo_alt_def set_plus_def)
-  apply (rule_tac x = "c + ca" in exI)
-  apply auto
-  apply (simp add: ring_distribs func_plus)
-  apply (blast intro:order_trans abs_triangle_ineq add_mono elim:) 
-done
-
-lemma bigo_plus_idemp [simp]: "O(f) \<oplus> O(f) = O(f)"
-  apply (rule equalityI)
-  apply (rule bigo_plus_self_subset)
-  apply (rule set_zero_plus2) 
-  apply (rule bigo_zero)
-done
-
-lemma bigo_plus_subset [intro]: "O(f + g) <= O(f) \<oplus> O(g)"
-  apply (rule subsetI)
-  apply (auto simp add: bigo_def bigo_pos_const func_plus set_plus_def)
-  apply (subst bigo_pos_const [symmetric])+
-  apply (rule_tac x = 
-    "%n. if abs (g n) <= (abs (f n)) then x n else 0" in exI)
-  apply (rule conjI)
-  apply (rule_tac x = "c + c" in exI)
-  apply (clarsimp)
-  apply (auto)
-  apply (subgoal_tac "c * abs (f xa + g xa) <= (c + c) * abs (f xa)")
-  apply (erule_tac x = xa in allE)
-  apply (erule order_trans)
-  apply (simp)
-  apply (subgoal_tac "c * abs (f xa + g xa) <= c * (abs (f xa) + abs (g xa))")
-  apply (erule order_trans)
-  apply (simp add: ring_distribs)
-  apply (rule mult_left_mono)
-  apply assumption
-  apply (simp add: order_less_le)
-  apply (rule mult_left_mono)
-  apply (simp add: abs_triangle_ineq)
-  apply (simp add: order_less_le)
-  apply (rule mult_nonneg_nonneg)
-  apply (rule add_nonneg_nonneg)
-  apply auto
-  apply (rule_tac x = "%n. if (abs (f n)) <  abs (g n) then x n else 0" 
-     in exI)
-  apply (rule conjI)
-  apply (rule_tac x = "c + c" in exI)
-  apply auto
-  apply (subgoal_tac "c * abs (f xa + g xa) <= (c + c) * abs (g xa)")
-  apply (erule_tac x = xa in allE)
-  apply (erule order_trans)
-  apply (simp)
-  apply (subgoal_tac "c * abs (f xa + g xa) <= c * (abs (f xa) + abs (g xa))")
-  apply (erule order_trans)
-  apply (simp add: ring_distribs)
-  apply (rule mult_left_mono)
-  apply (simp add: order_less_le)
-  apply (simp add: order_less_le)
-  apply (rule mult_left_mono)
-  apply (rule abs_triangle_ineq)
-  apply (simp add: order_less_le)
-apply (metis abs_not_less_zero double_less_0_iff less_not_permute linorder_not_less mult_less_0_iff)
-  apply (rule ext)
-  apply (auto simp add: if_splits linorder_not_le)
-done
-
-lemma bigo_plus_subset2 [intro]: "A <= O(f) ==> B <= O(f) ==> A \<oplus> B <= O(f)"
-  apply (subgoal_tac "A \<oplus> B <= O(f) \<oplus> O(f)")
-  apply (erule order_trans)
-  apply simp
-  apply (auto del: subsetI simp del: bigo_plus_idemp)
-done
-
-declare [[ atp_problem_prefix = "BigO__bigo_plus_eq" ]]
-lemma bigo_plus_eq: "ALL x. 0 <= f x ==> ALL x. 0 <= g x ==> 
-  O(f + g) = O(f) \<oplus> O(g)"
-  apply (rule equalityI)
-  apply (rule bigo_plus_subset)
-  apply (simp add: bigo_alt_def set_plus_def func_plus)
-  apply clarify 
-(*sledgehammer*); 
-  apply (rule_tac x = "max c ca" in exI)
-  apply (rule conjI)
-   apply (metis Orderings.less_max_iff_disj)
-  apply clarify
-  apply (drule_tac x = "xa" in spec)+
-  apply (subgoal_tac "0 <= f xa + g xa")
-  apply (simp add: ring_distribs)
-  apply (subgoal_tac "abs(a xa + b xa) <= abs(a xa) + abs(b xa)")
-  apply (subgoal_tac "abs(a xa) + abs(b xa) <= 
-      max c ca * f xa + max c ca * g xa")
-  apply (blast intro: order_trans)
-  defer 1
-  apply (rule abs_triangle_ineq)
-  apply (metis add_nonneg_nonneg)
-  apply (rule add_mono)
-using [[ atp_problem_prefix = "BigO__bigo_plus_eq_simpler" ]] 
-(*Found by SPASS; SLOW*)
-apply (metis le_maxI2 linorder_linear linorder_not_le min_max.sup_absorb1 mult_le_cancel_right order_trans)
-apply (metis le_maxI2 linorder_not_le mult_le_cancel_right order_trans)
-done
-
-declare [[ atp_problem_prefix = "BigO__bigo_bounded_alt" ]]
-lemma bigo_bounded_alt: "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> 
-    f : O(g)" 
-  apply (auto simp add: bigo_def)
-(*Version 1: one-shot proof*)
-  apply (metis OrderedGroup.abs_le_D1 linorder_class.not_less  order_less_le  Orderings.xt1(12)  Ring_and_Field.abs_mult)
-  done
-
-lemma (*bigo_bounded_alt:*) "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> 
-    f : O(g)" 
-  apply (auto simp add: bigo_def)
-(*Version 2: single-step proof*)
-proof (neg_clausify)
-fix x
-assume 0: "\<And>x. f x \<le> c * g x"
-assume 1: "\<And>xa. \<not> f (x xa) \<le> xa * \<bar>g (x xa)\<bar>"
-have 2: "\<And>X3. c * g X3 = f X3 \<or> \<not> c * g X3 \<le> f X3"
-  by (metis 0 order_antisym_conv)
-have 3: "\<And>X3. \<not> f (x \<bar>X3\<bar>) \<le> \<bar>X3 * g (x \<bar>X3\<bar>)\<bar>"
-  by (metis 1 abs_mult)
-have 4: "\<And>X1 X3\<Colon>'b\<Colon>ordered_idom. X3 \<le> X1 \<or> X1 \<le> \<bar>X3\<bar>"
-  by (metis linorder_linear abs_le_D1)
-have 5: "\<And>X3::'b. \<bar>X3\<bar> * \<bar>X3\<bar> = X3 * X3"
-  by (metis abs_mult_self)
-have 6: "\<And>X3. \<not> X3 * X3 < (0\<Colon>'b\<Colon>ordered_idom)"
-  by (metis not_square_less_zero)
-have 7: "\<And>X1 X3::'b. \<bar>X1\<bar> * \<bar>X3\<bar> = \<bar>X3 * X1\<bar>"
-  by (metis abs_mult mult_commute)
-have 8: "\<And>X3::'b. X3 * X3 = \<bar>X3 * X3\<bar>"
-  by (metis abs_mult 5)
-have 9: "\<And>X3. X3 * g (x \<bar>X3\<bar>) \<le> f (x \<bar>X3\<bar>)"
-  by (metis 3 4)
-have 10: "c * g (x \<bar>c\<bar>) = f (x \<bar>c\<bar>)"
-  by (metis 2 9)
-have 11: "\<And>X3::'b. \<bar>X3\<bar> * \<bar>\<bar>X3\<bar>\<bar> = \<bar>X3\<bar> * \<bar>X3\<bar>"
-  by (metis abs_idempotent abs_mult 8)
-have 12: "\<And>X3::'b. \<bar>X3 * \<bar>X3\<bar>\<bar> = \<bar>X3\<bar> * \<bar>X3\<bar>"
-  by (metis mult_commute 7 11)
-have 13: "\<And>X3::'b. \<bar>X3 * \<bar>X3\<bar>\<bar> = X3 * X3"
-  by (metis 8 7 12)
-have 14: "\<And>X3. X3 \<le> \<bar>X3\<bar> \<or> X3 < (0\<Colon>'b)"
-  by (metis abs_ge_self abs_le_D1 abs_if)
-have 15: "\<And>X3. X3 \<le> \<bar>X3\<bar> \<or> \<bar>X3\<bar> < (0\<Colon>'b)"
-  by (metis abs_ge_self abs_le_D1 abs_if)
-have 16: "\<And>X3. X3 * X3 < (0\<Colon>'b) \<or> X3 * \<bar>X3\<bar> \<le> X3 * X3"
-  by (metis 15 13)
-have 17: "\<And>X3::'b. X3 * \<bar>X3\<bar> \<le> X3 * X3"
-  by (metis 16 6)
-have 18: "\<And>X3. X3 \<le> \<bar>X3\<bar> \<or> \<not> X3 < (0\<Colon>'b)"
-  by (metis mult_le_cancel_left 17)
-have 19: "\<And>X3::'b. X3 \<le> \<bar>X3\<bar>"
-  by (metis 18 14)
-have 20: "\<not> f (x \<bar>c\<bar>) \<le> \<bar>f (x \<bar>c\<bar>)\<bar>"
-  by (metis 3 10)
-show "False"
-  by (metis 20 19)
-qed
-
-
-text{*So here is the easier (and more natural) problem using transitivity*}
-declare [[ atp_problem_prefix = "BigO__bigo_bounded_alt_trans" ]]
-lemma "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)" 
-  apply (auto simp add: bigo_def)
-  (*Version 1: one-shot proof*) 
-  apply (metis Orderings.leD Orderings.leI abs_ge_self abs_le_D1 abs_mult abs_of_nonneg order_le_less)
-  done
-
-text{*So here is the easier (and more natural) problem using transitivity*}
-declare [[ atp_problem_prefix = "BigO__bigo_bounded_alt_trans" ]]
-lemma "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)" 
-  apply (auto simp add: bigo_def)
-(*Version 2: single-step proof*)
-proof (neg_clausify)
-fix x
-assume 0: "\<And>A\<Colon>'a\<Colon>type.
-   (f\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) A
-   \<le> (c\<Colon>'b\<Colon>ordered_idom) * (g\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) A"
-assume 1: "\<And>A\<Colon>'b\<Colon>ordered_idom.
-   \<not> (f\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) ((x\<Colon>'b\<Colon>ordered_idom \<Rightarrow> 'a\<Colon>type) A)
-     \<le> A * \<bar>(g\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) (x A)\<bar>"
-have 2: "\<And>X2\<Colon>'a\<Colon>type.
-   \<not> (c\<Colon>'b\<Colon>ordered_idom) * (g\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) X2
-     < (f\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) X2"
-  by (metis 0 linorder_not_le)
-have 3: "\<And>X2\<Colon>'b\<Colon>ordered_idom.
-   \<not> (f\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) ((x\<Colon>'b\<Colon>ordered_idom \<Rightarrow> 'a\<Colon>type) \<bar>X2\<bar>)
-     \<le> \<bar>X2 * (g\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) (x \<bar>X2\<bar>)\<bar>"
-  by (metis abs_mult 1)
-have 4: "\<And>X2\<Colon>'b\<Colon>ordered_idom.
-   \<bar>X2 * (g\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) ((x\<Colon>'b\<Colon>ordered_idom \<Rightarrow> 'a\<Colon>type) \<bar>X2\<bar>)\<bar>
-   < (f\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) (x \<bar>X2\<bar>)"
-  by (metis 3 linorder_not_less)
-have 5: "\<And>X2\<Colon>'b\<Colon>ordered_idom.
-   X2 * (g\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) ((x\<Colon>'b\<Colon>ordered_idom \<Rightarrow> 'a\<Colon>type) \<bar>X2\<bar>)
-   < (f\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) (x \<bar>X2\<bar>)"
-  by (metis abs_less_iff 4)
-show "False"
-  by (metis 2 5)
-qed
-
-
-lemma bigo_bounded: "ALL x. 0 <= f x ==> ALL x. f x <= g x ==> 
-    f : O(g)" 
-  apply (erule bigo_bounded_alt [of f 1 g])
-  apply simp
-done
-
-declare [[ atp_problem_prefix = "BigO__bigo_bounded2" ]]
-lemma bigo_bounded2: "ALL x. lb x <= f x ==> ALL x. f x <= lb x + g x ==>
-    f : lb +o O(g)"
-  apply (rule set_minus_imp_plus)
-  apply (rule bigo_bounded)
-  apply (auto simp add: diff_minus fun_Compl_def func_plus)
-  prefer 2
-  apply (drule_tac x = x in spec)+ 
-  apply arith (*not clear that it's provable otherwise*) 
-proof (neg_clausify)
-fix x
-assume 0: "\<And>y. lb y \<le> f y"
-assume 1: "\<not> (0\<Colon>'b) \<le> f x + - lb x"
-have 2: "\<And>X3. (0\<Colon>'b) + X3 = X3"
-  by (metis diff_eq_eq right_minus_eq)
-have 3: "\<not> (0\<Colon>'b) \<le> f x - lb x"
-  by (metis 1 diff_minus)
-have 4: "\<not> (0\<Colon>'b) + lb x \<le> f x"
-  by (metis 3 le_diff_eq)
-show "False"
-  by (metis 4 2 0)
-qed
-
-declare [[ atp_problem_prefix = "BigO__bigo_abs" ]]
-lemma bigo_abs: "(%x. abs(f x)) =o O(f)" 
-  apply (unfold bigo_def)
-  apply auto
-proof (neg_clausify)
-fix x
-assume 0: "\<And>xa. \<not> \<bar>f (x xa)\<bar> \<le> xa * \<bar>f (x xa)\<bar>"
-have 1: "\<And>X2. X2 \<le> (1\<Colon>'b) * X2 \<or> \<not> (1\<Colon>'b) \<le> (1\<Colon>'b)"
-  by (metis mult_le_cancel_right1 order_eq_iff)
-have 2: "\<And>X2. X2 \<le> (1\<Colon>'b) * X2"
-  by (metis order_eq_iff 1)
-show "False"
-  by (metis 0 2)
-qed
-
-declare [[ atp_problem_prefix = "BigO__bigo_abs2" ]]
-lemma bigo_abs2: "f =o O(%x. abs(f x))"
-  apply (unfold bigo_def)
-  apply auto
-proof (neg_clausify)
-fix x
-assume 0: "\<And>xa. \<not> \<bar>f (x xa)\<bar> \<le> xa * \<bar>f (x xa)\<bar>"
-have 1: "\<And>X2. X2 \<le> (1\<Colon>'b) * X2 \<or> \<not> (1\<Colon>'b) \<le> (1\<Colon>'b)"
-  by (metis mult_le_cancel_right1 order_eq_iff)
-have 2: "\<And>X2. X2 \<le> (1\<Colon>'b) * X2"
-  by (metis order_eq_iff 1)
-show "False"
-  by (metis 0 2)
-qed
- 
-lemma bigo_abs3: "O(f) = O(%x. abs(f x))"
-  apply (rule equalityI)
-  apply (rule bigo_elt_subset)
-  apply (rule bigo_abs2)
-  apply (rule bigo_elt_subset)
-  apply (rule bigo_abs)
-done
-
-lemma bigo_abs4: "f =o g +o O(h) ==> 
-    (%x. abs (f x)) =o (%x. abs (g x)) +o O(h)"
-  apply (drule set_plus_imp_minus)
-  apply (rule set_minus_imp_plus)
-  apply (subst fun_diff_def)
-proof -
-  assume a: "f - g : O(h)"
-  have "(%x. abs (f x) - abs (g x)) =o O(%x. abs(abs (f x) - abs (g x)))"
-    by (rule bigo_abs2)
-  also have "... <= O(%x. abs (f x - g x))"
-    apply (rule bigo_elt_subset)
-    apply (rule bigo_bounded)
-    apply force
-    apply (rule allI)
-    apply (rule abs_triangle_ineq3)
-    done
-  also have "... <= O(f - g)"
-    apply (rule bigo_elt_subset)
-    apply (subst fun_diff_def)
-    apply (rule bigo_abs)
-    done
-  also have "... <= O(h)"
-    using a by (rule bigo_elt_subset)
-  finally show "(%x. abs (f x) - abs (g x)) : O(h)".
-qed
-
-lemma bigo_abs5: "f =o O(g) ==> (%x. abs(f x)) =o O(g)" 
-by (unfold bigo_def, auto)
-
-lemma bigo_elt_subset2 [intro]: "f : g +o O(h) ==> O(f) <= O(g) \<oplus> O(h)"
-proof -
-  assume "f : g +o O(h)"
-  also have "... <= O(g) \<oplus> O(h)"
-    by (auto del: subsetI)
-  also have "... = O(%x. abs(g x)) \<oplus> O(%x. abs(h x))"
-    apply (subst bigo_abs3 [symmetric])+
-    apply (rule refl)
-    done
-  also have "... = O((%x. abs(g x)) + (%x. abs(h x)))"
-    by (rule bigo_plus_eq [symmetric], auto)
-  finally have "f : ...".
-  then have "O(f) <= ..."
-    by (elim bigo_elt_subset)
-  also have "... = O(%x. abs(g x)) \<oplus> O(%x. abs(h x))"
-    by (rule bigo_plus_eq, auto)
-  finally show ?thesis
-    by (simp add: bigo_abs3 [symmetric])
-qed
-
-declare [[ atp_problem_prefix = "BigO__bigo_mult" ]]
-lemma bigo_mult [intro]: "O(f)\<otimes>O(g) <= O(f * g)"
-  apply (rule subsetI)
-  apply (subst bigo_def)
-  apply (auto simp del: abs_mult mult_ac
-              simp add: bigo_alt_def set_times_def func_times)
-(*sledgehammer*); 
-  apply (rule_tac x = "c * ca" in exI)
-  apply(rule allI)
-  apply(erule_tac x = x in allE)+
-  apply(subgoal_tac "c * ca * abs(f x * g x) = 
-      (c * abs(f x)) * (ca * abs(g x))")
-using [[ atp_problem_prefix = "BigO__bigo_mult_simpler" ]]
-prefer 2 
-apply (metis mult_assoc mult_left_commute
-  OrderedGroup.abs_of_pos OrderedGroup.mult_left_commute
-  Ring_and_Field.abs_mult Ring_and_Field.mult_pos_pos)
-  apply (erule ssubst) 
-  apply (subst abs_mult)
-(*not qute BigO__bigo_mult_simpler_1 (a hard problem!) as abs_mult has
-  just been done*)
-proof (neg_clausify)
-fix a c b ca x
-assume 0: "(0\<Colon>'b\<Colon>ordered_idom) < (c\<Colon>'b\<Colon>ordered_idom)"
-assume 1: "\<bar>(a\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) (x\<Colon>'a)\<bar>
-\<le> (c\<Colon>'b\<Colon>ordered_idom) * \<bar>(f\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>"
-assume 2: "\<bar>(b\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) (x\<Colon>'a)\<bar>
-\<le> (ca\<Colon>'b\<Colon>ordered_idom) * \<bar>(g\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>"
-assume 3: "\<not> \<bar>(a\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) (x\<Colon>'a)\<bar> *
-  \<bar>(b\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>
-  \<le> (c\<Colon>'b\<Colon>ordered_idom) * \<bar>(f\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar> *
-    ((ca\<Colon>'b\<Colon>ordered_idom) * \<bar>(g\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>)"
-have 4: "\<bar>c\<Colon>'b\<Colon>ordered_idom\<bar> = c"
-  by (metis OrderedGroup.abs_of_pos 0)
-have 5: "\<And>X1\<Colon>'b\<Colon>ordered_idom. (c\<Colon>'b\<Colon>ordered_idom) * \<bar>X1\<bar> = \<bar>c * X1\<bar>"
-  by (metis Ring_and_Field.abs_mult 4)
-have 6: "(0\<Colon>'b\<Colon>ordered_idom) = (1\<Colon>'b\<Colon>ordered_idom) \<or>
-(0\<Colon>'b\<Colon>ordered_idom) < (1\<Colon>'b\<Colon>ordered_idom)"
-  by (metis OrderedGroup.abs_not_less_zero Ring_and_Field.abs_one Ring_and_Field.linorder_neqE_ordered_idom)
-have 7: "(0\<Colon>'b\<Colon>ordered_idom) < (1\<Colon>'b\<Colon>ordered_idom)"
-  by (metis 6 Ring_and_Field.one_neq_zero)
-have 8: "\<bar>1\<Colon>'b\<Colon>ordered_idom\<bar> = (1\<Colon>'b\<Colon>ordered_idom)"
-  by (metis OrderedGroup.abs_of_pos 7)
-have 9: "\<And>X1\<Colon>'b\<Colon>ordered_idom. (0\<Colon>'b\<Colon>ordered_idom) \<le> (c\<Colon>'b\<Colon>ordered_idom) * \<bar>X1\<bar>"
-  by (metis OrderedGroup.abs_ge_zero 5)
-have 10: "\<And>X1\<Colon>'b\<Colon>ordered_idom. X1 * (1\<Colon>'b\<Colon>ordered_idom) = X1"
-  by (metis Ring_and_Field.mult_cancel_right2 mult_commute)
-have 11: "\<And>X1\<Colon>'b\<Colon>ordered_idom. \<bar>\<bar>X1\<bar>\<bar> = \<bar>X1\<bar> * \<bar>1\<Colon>'b\<Colon>ordered_idom\<bar>"
-  by (metis Ring_and_Field.abs_mult OrderedGroup.abs_idempotent 10)
-have 12: "\<And>X1\<Colon>'b\<Colon>ordered_idom. \<bar>\<bar>X1\<bar>\<bar> = \<bar>X1\<bar>"
-  by (metis 11 8 10)
-have 13: "\<And>X1\<Colon>'b\<Colon>ordered_idom. (0\<Colon>'b\<Colon>ordered_idom) \<le> \<bar>X1\<bar>"
-  by (metis OrderedGroup.abs_ge_zero 12)
-have 14: "\<not> (0\<Colon>'b\<Colon>ordered_idom)
-  \<le> (c\<Colon>'b\<Colon>ordered_idom) * \<bar>(f\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) (x\<Colon>'a)\<bar> \<or>
-\<not> (0\<Colon>'b\<Colon>ordered_idom) \<le> \<bar>(b\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar> \<or>
-\<not> \<bar>b x\<bar> \<le> (ca\<Colon>'b\<Colon>ordered_idom) * \<bar>(g\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar> \<or>
-\<not> \<bar>(a\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar> \<le> c * \<bar>f x\<bar>"
-  by (metis 3 Ring_and_Field.mult_mono)
-have 15: "\<not> (0\<Colon>'b\<Colon>ordered_idom) \<le> \<bar>(b\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) (x\<Colon>'a)\<bar> \<or>
-\<not> \<bar>b x\<bar> \<le> (ca\<Colon>'b\<Colon>ordered_idom) * \<bar>(g\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar> \<or>
-\<not> \<bar>(a\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>
-  \<le> (c\<Colon>'b\<Colon>ordered_idom) * \<bar>(f\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>"
-  by (metis 14 9)
-have 16: "\<not> \<bar>(b\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) (x\<Colon>'a)\<bar>
-  \<le> (ca\<Colon>'b\<Colon>ordered_idom) * \<bar>(g\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar> \<or>
-\<not> \<bar>(a\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>
-  \<le> (c\<Colon>'b\<Colon>ordered_idom) * \<bar>(f\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>"
-  by (metis 15 13)
-have 17: "\<not> \<bar>(a\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) (x\<Colon>'a)\<bar>
-  \<le> (c\<Colon>'b\<Colon>ordered_idom) * \<bar>(f\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>"
-  by (metis 16 2)
-show 18: "False"
-  by (metis 17 1)
-qed
-
-
-declare [[ atp_problem_prefix = "BigO__bigo_mult2" ]]
-lemma bigo_mult2 [intro]: "f *o O(g) <= O(f * g)"
-  apply (auto simp add: bigo_def elt_set_times_def func_times abs_mult)
-(*sledgehammer*); 
-  apply (rule_tac x = c in exI)
-  apply clarify
-  apply (drule_tac x = x in spec)
-using [[ atp_problem_prefix = "BigO__bigo_mult2_simpler" ]]
-(*sledgehammer [no luck]*); 
-  apply (subgoal_tac "abs(f x) * abs(b x) <= abs(f x) * (c * abs(g x))")
-  apply (simp add: mult_ac)
-  apply (rule mult_left_mono, assumption)
-  apply (rule abs_ge_zero)
-done
-
-declare [[ atp_problem_prefix = "BigO__bigo_mult3" ]]
-lemma bigo_mult3: "f : O(h) ==> g : O(j) ==> f * g : O(h * j)"
-by (metis bigo_mult set_times_intro subset_iff)
-
-declare [[ atp_problem_prefix = "BigO__bigo_mult4" ]]
-lemma bigo_mult4 [intro]:"f : k +o O(h) ==> g * f : (g * k) +o O(g * h)"
-by (metis bigo_mult2 set_plus_mono_b set_times_intro2 set_times_plus_distrib)
-
-
-lemma bigo_mult5: "ALL x. f x ~= 0 ==>
-    O(f * g) <= (f::'a => ('b::ordered_field)) *o O(g)"
-proof -
-  assume "ALL x. f x ~= 0"
-  show "O(f * g) <= f *o O(g)"
-  proof
-    fix h
-    assume "h : O(f * g)"
-    then have "(%x. 1 / (f x)) * h : (%x. 1 / f x) *o O(f * g)"
-      by auto
-    also have "... <= O((%x. 1 / f x) * (f * g))"
-      by (rule bigo_mult2)
-    also have "(%x. 1 / f x) * (f * g) = g"
-      apply (simp add: func_times) 
-      apply (rule ext)
-      apply (simp add: prems nonzero_divide_eq_eq mult_ac)
-      done
-    finally have "(%x. (1::'b) / f x) * h : O(g)".
-    then have "f * ((%x. (1::'b) / f x) * h) : f *o O(g)"
-      by auto
-    also have "f * ((%x. (1::'b) / f x) * h) = h"
-      apply (simp add: func_times) 
-      apply (rule ext)
-      apply (simp add: prems nonzero_divide_eq_eq mult_ac)
-      done
-    finally show "h : f *o O(g)".
-  qed
-qed
-
-declare [[ atp_problem_prefix = "BigO__bigo_mult6" ]]
-lemma bigo_mult6: "ALL x. f x ~= 0 ==>
-    O(f * g) = (f::'a => ('b::ordered_field)) *o O(g)"
-by (metis bigo_mult2 bigo_mult5 order_antisym)
-
-(*proof requires relaxing relevance: 2007-01-25*)
-declare [[ atp_problem_prefix = "BigO__bigo_mult7" ]]
-  declare bigo_mult6 [simp]
-lemma bigo_mult7: "ALL x. f x ~= 0 ==>
-    O(f * g) <= O(f::'a => ('b::ordered_field)) \<otimes> O(g)"
-(*sledgehammer*)
-  apply (subst bigo_mult6)
-  apply assumption
-  apply (rule set_times_mono3) 
-  apply (rule bigo_refl)
-done
-  declare bigo_mult6 [simp del]
-
-declare [[ atp_problem_prefix = "BigO__bigo_mult8" ]]
-  declare bigo_mult7[intro!]
-lemma bigo_mult8: "ALL x. f x ~= 0 ==>
-    O(f * g) = O(f::'a => ('b::ordered_field)) \<otimes> O(g)"
-by (metis bigo_mult bigo_mult7 order_antisym_conv)
-
-lemma bigo_minus [intro]: "f : O(g) ==> - f : O(g)"
-  by (auto simp add: bigo_def fun_Compl_def)
-
-lemma bigo_minus2: "f : g +o O(h) ==> -f : -g +o O(h)"
-  apply (rule set_minus_imp_plus)
-  apply (drule set_plus_imp_minus)
-  apply (drule bigo_minus)
-  apply (simp add: diff_minus)
-done
-
-lemma bigo_minus3: "O(-f) = O(f)"
-  by (auto simp add: bigo_def fun_Compl_def abs_minus_cancel)
-
-lemma bigo_plus_absorb_lemma1: "f : O(g) ==> f +o O(g) <= O(g)"
-proof -
-  assume a: "f : O(g)"
-  show "f +o O(g) <= O(g)"
-  proof -
-    have "f : O(f)" by auto
-    then have "f +o O(g) <= O(f) \<oplus> O(g)"
-      by (auto del: subsetI)
-    also have "... <= O(g) \<oplus> O(g)"
-    proof -
-      from a have "O(f) <= O(g)" by (auto del: subsetI)
-      thus ?thesis by (auto del: subsetI)
-    qed
-    also have "... <= O(g)" by (simp add: bigo_plus_idemp)
-    finally show ?thesis .
-  qed
-qed
-
-lemma bigo_plus_absorb_lemma2: "f : O(g) ==> O(g) <= f +o O(g)"
-proof -
-  assume a: "f : O(g)"
-  show "O(g) <= f +o O(g)"
-  proof -
-    from a have "-f : O(g)" by auto
-    then have "-f +o O(g) <= O(g)" by (elim bigo_plus_absorb_lemma1)
-    then have "f +o (-f +o O(g)) <= f +o O(g)" by auto
-    also have "f +o (-f +o O(g)) = O(g)"
-      by (simp add: set_plus_rearranges)
-    finally show ?thesis .
-  qed
-qed
-
-declare [[ atp_problem_prefix = "BigO__bigo_plus_absorb" ]]
-lemma bigo_plus_absorb [simp]: "f : O(g) ==> f +o O(g) = O(g)"
-by (metis bigo_plus_absorb_lemma1 bigo_plus_absorb_lemma2 order_eq_iff);
-
-lemma bigo_plus_absorb2 [intro]: "f : O(g) ==> A <= O(g) ==> f +o A <= O(g)"
-  apply (subgoal_tac "f +o A <= f +o O(g)")
-  apply force+
-done
-
-lemma bigo_add_commute_imp: "f : g +o O(h) ==> g : f +o O(h)"
-  apply (subst set_minus_plus [symmetric])
-  apply (subgoal_tac "g - f = - (f - g)")
-  apply (erule ssubst)
-  apply (rule bigo_minus)
-  apply (subst set_minus_plus)
-  apply assumption
-  apply  (simp add: diff_minus add_ac)
-done
-
-lemma bigo_add_commute: "(f : g +o O(h)) = (g : f +o O(h))"
-  apply (rule iffI)
-  apply (erule bigo_add_commute_imp)+
-done
-
-lemma bigo_const1: "(%x. c) : O(%x. 1)"
-by (auto simp add: bigo_def mult_ac)
-
-declare [[ atp_problem_prefix = "BigO__bigo_const2" ]]
-lemma (*bigo_const2 [intro]:*) "O(%x. c) <= O(%x. 1)"
-by (metis bigo_const1 bigo_elt_subset);
-
-lemma bigo_const2 [intro]: "O(%x. c::'b::ordered_idom) <= O(%x. 1)";
-(*??FAILS because the two occurrences of COMBK have different polymorphic types
-proof (neg_clausify)
-assume 0: "\<not> O(COMBK (c\<Colon>'b\<Colon>ordered_idom)) \<subseteq> O(COMBK (1\<Colon>'b\<Colon>ordered_idom))"
-have 1: "COMBK (c\<Colon>'b\<Colon>ordered_idom) \<notin> O(COMBK (1\<Colon>'b\<Colon>ordered_idom))"
-apply (rule notI) 
-apply (rule 0 [THEN notE]) 
-apply (rule bigo_elt_subset) 
-apply assumption; 
-sorry
-  by (metis 0 bigo_elt_subset)  loops??
-show "False"
-  by (metis 1 bigo_const1)
-qed
-*)
-  apply (rule bigo_elt_subset)
-  apply (rule bigo_const1)
-done
-
-declare [[ atp_problem_prefix = "BigO__bigo_const3" ]]
-lemma bigo_const3: "(c::'a::ordered_field) ~= 0 ==> (%x. 1) : O(%x. c)"
-apply (simp add: bigo_def)
-proof (neg_clausify)
-assume 0: "(c\<Colon>'a\<Colon>ordered_field) \<noteq> (0\<Colon>'a\<Colon>ordered_field)"
-assume 1: "\<And>A\<Colon>'a\<Colon>ordered_field. \<not> (1\<Colon>'a\<Colon>ordered_field) \<le> A * \<bar>c\<Colon>'a\<Colon>ordered_field\<bar>"
-have 2: "(0\<Colon>'a\<Colon>ordered_field) = \<bar>c\<Colon>'a\<Colon>ordered_field\<bar> \<or>
-\<not> (1\<Colon>'a\<Colon>ordered_field) \<le> (1\<Colon>'a\<Colon>ordered_field)"
-  by (metis 1 field_inverse)
-have 3: "\<bar>c\<Colon>'a\<Colon>ordered_field\<bar> = (0\<Colon>'a\<Colon>ordered_field)"
-  by (metis linorder_neq_iff linorder_antisym_conv1 2)
-have 4: "(0\<Colon>'a\<Colon>ordered_field) = (c\<Colon>'a\<Colon>ordered_field)"
-  by (metis 3 abs_eq_0)
-show "False"
-  by (metis 0 4)
-qed
-
-lemma bigo_const4: "(c::'a::ordered_field) ~= 0 ==> O(%x. 1) <= O(%x. c)"
-by (rule bigo_elt_subset, rule bigo_const3, assumption)
-
-lemma bigo_const [simp]: "(c::'a::ordered_field) ~= 0 ==> 
-    O(%x. c) = O(%x. 1)"
-by (rule equalityI, rule bigo_const2, rule bigo_const4, assumption)
-
-declare [[ atp_problem_prefix = "BigO__bigo_const_mult1" ]]
-lemma bigo_const_mult1: "(%x. c * f x) : O(f)"
-  apply (simp add: bigo_def abs_mult)
-proof (neg_clausify)
-fix x
-assume 0: "\<And>xa\<Colon>'b\<Colon>ordered_idom.
-   \<not> \<bar>c\<Colon>'b\<Colon>ordered_idom\<bar> *
-     \<bar>(f\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) ((x\<Colon>'b\<Colon>ordered_idom \<Rightarrow> 'a\<Colon>type) xa)\<bar>
-     \<le> xa * \<bar>f (x xa)\<bar>"
-show "False"
-  by (metis linorder_neq_iff linorder_antisym_conv1 0)
-qed
-
-lemma bigo_const_mult2: "O(%x. c * f x) <= O(f)"
-by (rule bigo_elt_subset, rule bigo_const_mult1)
-
-declare [[ atp_problem_prefix = "BigO__bigo_const_mult3" ]]
-lemma bigo_const_mult3: "(c::'a::ordered_field) ~= 0 ==> f : O(%x. c * f x)"
-  apply (simp add: bigo_def)
-(*sledgehammer [no luck]*); 
-  apply (rule_tac x = "abs(inverse c)" in exI)
-  apply (simp only: abs_mult [symmetric] mult_assoc [symmetric])
-apply (subst left_inverse) 
-apply (auto ); 
-done
-
-lemma bigo_const_mult4: "(c::'a::ordered_field) ~= 0 ==> 
-    O(f) <= O(%x. c * f x)"
-by (rule bigo_elt_subset, rule bigo_const_mult3, assumption)
-
-lemma bigo_const_mult [simp]: "(c::'a::ordered_field) ~= 0 ==> 
-    O(%x. c * f x) = O(f)"
-by (rule equalityI, rule bigo_const_mult2, erule bigo_const_mult4)
-
-declare [[ atp_problem_prefix = "BigO__bigo_const_mult5" ]]
-lemma bigo_const_mult5 [simp]: "(c::'a::ordered_field) ~= 0 ==> 
-    (%x. c) *o O(f) = O(f)"
-  apply (auto del: subsetI)
-  apply (rule order_trans)
-  apply (rule bigo_mult2)
-  apply (simp add: func_times)
-  apply (auto intro!: subsetI simp add: bigo_def elt_set_times_def func_times)
-  apply (rule_tac x = "%y. inverse c * x y" in exI)
-  apply (rename_tac g d) 
-  apply safe
-  apply (rule_tac [2] ext) 
-   prefer 2 
-   apply simp
-  apply (simp add: mult_assoc [symmetric] abs_mult)
-  (*couldn't get this proof without the step above; SLOW*)
-  apply (metis mult_assoc abs_ge_zero mult_left_mono)
-done
-
-
-declare [[ atp_problem_prefix = "BigO__bigo_const_mult6" ]]
-lemma bigo_const_mult6 [intro]: "(%x. c) *o O(f) <= O(f)"
-  apply (auto intro!: subsetI
-    simp add: bigo_def elt_set_times_def func_times
-    simp del: abs_mult mult_ac)
-(*sledgehammer*); 
-  apply (rule_tac x = "ca * (abs c)" in exI)
-  apply (rule allI)
-  apply (subgoal_tac "ca * abs(c) * abs(f x) = abs(c) * (ca * abs(f x))")
-  apply (erule ssubst)
-  apply (subst abs_mult)
-  apply (rule mult_left_mono)
-  apply (erule spec)
-  apply simp
-  apply(simp add: mult_ac)
-done
-
-lemma bigo_const_mult7 [intro]: "f =o O(g) ==> (%x. c * f x) =o O(g)"
-proof -
-  assume "f =o O(g)"
-  then have "(%x. c) * f =o (%x. c) *o O(g)"
-    by auto
-  also have "(%x. c) * f = (%x. c * f x)"
-    by (simp add: func_times)
-  also have "(%x. c) *o O(g) <= O(g)"
-    by (auto del: subsetI)
-  finally show ?thesis .
-qed
-
-lemma bigo_compose1: "f =o O(g) ==> (%x. f(k x)) =o O(%x. g(k x))"
-by (unfold bigo_def, auto)
-
-lemma bigo_compose2: "f =o g +o O(h) ==> (%x. f(k x)) =o (%x. g(k x)) +o 
-    O(%x. h(k x))"
-  apply (simp only: set_minus_plus [symmetric] diff_minus fun_Compl_def
-      func_plus)
-  apply (erule bigo_compose1)
-done
-
-subsection {* Setsum *}
-
-lemma bigo_setsum_main: "ALL x. ALL y : A x. 0 <= h x y ==> 
-    EX c. ALL x. ALL y : A x. abs(f x y) <= c * (h x y) ==>
-      (%x. SUM y : A x. f x y) =o O(%x. SUM y : A x. h x y)"  
-  apply (auto simp add: bigo_def)
-  apply (rule_tac x = "abs c" in exI)
-  apply (subst abs_of_nonneg) back back
-  apply (rule setsum_nonneg)
-  apply force
-  apply (subst setsum_right_distrib)
-  apply (rule allI)
-  apply (rule order_trans)
-  apply (rule setsum_abs)
-  apply (rule setsum_mono)
-apply (blast intro: order_trans mult_right_mono abs_ge_self) 
-done
-
-declare [[ atp_problem_prefix = "BigO__bigo_setsum1" ]]
-lemma bigo_setsum1: "ALL x y. 0 <= h x y ==> 
-    EX c. ALL x y. abs(f x y) <= c * (h x y) ==>
-      (%x. SUM y : A x. f x y) =o O(%x. SUM y : A x. h x y)"
-  apply (rule bigo_setsum_main)
-(*sledgehammer*); 
-  apply force
-  apply clarsimp
-  apply (rule_tac x = c in exI)
-  apply force
-done
-
-lemma bigo_setsum2: "ALL y. 0 <= h y ==> 
-    EX c. ALL y. abs(f y) <= c * (h y) ==>
-      (%x. SUM y : A x. f y) =o O(%x. SUM y : A x. h y)"
-by (rule bigo_setsum1, auto)  
-
-declare [[ atp_problem_prefix = "BigO__bigo_setsum3" ]]
-lemma bigo_setsum3: "f =o O(h) ==>
-    (%x. SUM y : A x. (l x y) * f(k x y)) =o
-      O(%x. SUM y : A x. abs(l x y * h(k x y)))"
-  apply (rule bigo_setsum1)
-  apply (rule allI)+
-  apply (rule abs_ge_zero)
-  apply (unfold bigo_def)
-  apply (auto simp add: abs_mult);
-(*sledgehammer*); 
-  apply (rule_tac x = c in exI)
-  apply (rule allI)+
-  apply (subst mult_left_commute)
-  apply (rule mult_left_mono)
-  apply (erule spec)
-  apply (rule abs_ge_zero)
-done
-
-lemma bigo_setsum4: "f =o g +o O(h) ==>
-    (%x. SUM y : A x. l x y * f(k x y)) =o
-      (%x. SUM y : A x. l x y * g(k x y)) +o
-        O(%x. SUM y : A x. abs(l x y * h(k x y)))"
-  apply (rule set_minus_imp_plus)
-  apply (subst fun_diff_def)
-  apply (subst setsum_subtractf [symmetric])
-  apply (subst right_diff_distrib [symmetric])
-  apply (rule bigo_setsum3)
-  apply (subst fun_diff_def [symmetric])
-  apply (erule set_plus_imp_minus)
-done
-
-declare [[ atp_problem_prefix = "BigO__bigo_setsum5" ]]
-lemma bigo_setsum5: "f =o O(h) ==> ALL x y. 0 <= l x y ==> 
-    ALL x. 0 <= h x ==>
-      (%x. SUM y : A x. (l x y) * f(k x y)) =o
-        O(%x. SUM y : A x. (l x y) * h(k x y))" 
-  apply (subgoal_tac "(%x. SUM y : A x. (l x y) * h(k x y)) = 
-      (%x. SUM y : A x. abs((l x y) * h(k x y)))")
-  apply (erule ssubst)
-  apply (erule bigo_setsum3)
-  apply (rule ext)
-  apply (rule setsum_cong2)
-  apply (thin_tac "f \<in> O(h)") 
-apply (metis abs_of_nonneg zero_le_mult_iff)
-done
-
-lemma bigo_setsum6: "f =o g +o O(h) ==> ALL x y. 0 <= l x y ==>
-    ALL x. 0 <= h x ==>
-      (%x. SUM y : A x. (l x y) * f(k x y)) =o
-        (%x. SUM y : A x. (l x y) * g(k x y)) +o
-          O(%x. SUM y : A x. (l x y) * h(k x y))" 
-  apply (rule set_minus_imp_plus)
-  apply (subst fun_diff_def)
-  apply (subst setsum_subtractf [symmetric])
-  apply (subst right_diff_distrib [symmetric])
-  apply (rule bigo_setsum5)
-  apply (subst fun_diff_def [symmetric])
-  apply (drule set_plus_imp_minus)
-  apply auto
-done
-
-subsection {* Misc useful stuff *}
-
-lemma bigo_useful_intro: "A <= O(f) ==> B <= O(f) ==>
-  A \<oplus> B <= O(f)"
-  apply (subst bigo_plus_idemp [symmetric])
-  apply (rule set_plus_mono2)
-  apply assumption+
-done
-
-lemma bigo_useful_add: "f =o O(h) ==> g =o O(h) ==> f + g =o O(h)"
-  apply (subst bigo_plus_idemp [symmetric])
-  apply (rule set_plus_intro)
-  apply assumption+
-done
-  
-lemma bigo_useful_const_mult: "(c::'a::ordered_field) ~= 0 ==> 
-    (%x. c) * f =o O(h) ==> f =o O(h)"
-  apply (rule subsetD)
-  apply (subgoal_tac "(%x. 1 / c) *o O(h) <= O(h)")
-  apply assumption
-  apply (rule bigo_const_mult6)
-  apply (subgoal_tac "f = (%x. 1 / c) * ((%x. c) * f)")
-  apply (erule ssubst)
-  apply (erule set_times_intro2)
-  apply (simp add: func_times) 
-done
-
-declare [[ atp_problem_prefix = "BigO__bigo_fix" ]]
-lemma bigo_fix: "(%x. f ((x::nat) + 1)) =o O(%x. h(x + 1)) ==> f 0 = 0 ==>
-    f =o O(h)"
-  apply (simp add: bigo_alt_def)
-(*sledgehammer*); 
-  apply clarify
-  apply (rule_tac x = c in exI)
-  apply safe
-  apply (case_tac "x = 0")
-apply (metis OrderedGroup.abs_ge_zero  OrderedGroup.abs_zero  order_less_le  Ring_and_Field.split_mult_pos_le) 
-  apply (subgoal_tac "x = Suc (x - 1)")
-  apply metis
-  apply simp
-  done
-
-
-lemma bigo_fix2: 
-    "(%x. f ((x::nat) + 1)) =o (%x. g(x + 1)) +o O(%x. h(x + 1)) ==> 
-       f 0 = g 0 ==> f =o g +o O(h)"
-  apply (rule set_minus_imp_plus)
-  apply (rule bigo_fix)
-  apply (subst fun_diff_def)
-  apply (subst fun_diff_def [symmetric])
-  apply (rule set_plus_imp_minus)
-  apply simp
-  apply (simp add: fun_diff_def)
-done
-
-subsection {* Less than or equal to *}
-
-constdefs 
-  lesso :: "('a => 'b::ordered_idom) => ('a => 'b) => ('a => 'b)"
-      (infixl "<o" 70)
-  "f <o g == (%x. max (f x - g x) 0)"
-
-lemma bigo_lesseq1: "f =o O(h) ==> ALL x. abs (g x) <= abs (f x) ==>
-    g =o O(h)"
-  apply (unfold bigo_def)
-  apply clarsimp
-apply (blast intro: order_trans) 
-done
-
-lemma bigo_lesseq2: "f =o O(h) ==> ALL x. abs (g x) <= f x ==>
-      g =o O(h)"
-  apply (erule bigo_lesseq1)
-apply (blast intro: abs_ge_self order_trans) 
-done
-
-lemma bigo_lesseq3: "f =o O(h) ==> ALL x. 0 <= g x ==> ALL x. g x <= f x ==>
-      g =o O(h)"
-  apply (erule bigo_lesseq2)
-  apply (rule allI)
-  apply (subst abs_of_nonneg)
-  apply (erule spec)+
-done
-
-lemma bigo_lesseq4: "f =o O(h) ==>
-    ALL x. 0 <= g x ==> ALL x. g x <= abs (f x) ==>
-      g =o O(h)"
-  apply (erule bigo_lesseq1)
-  apply (rule allI)
-  apply (subst abs_of_nonneg)
-  apply (erule spec)+
-done
-
-declare [[ atp_problem_prefix = "BigO__bigo_lesso1" ]]
-lemma bigo_lesso1: "ALL x. f x <= g x ==> f <o g =o O(h)"
-  apply (unfold lesso_def)
-  apply (subgoal_tac "(%x. max (f x - g x) 0) = 0")
-(*??Translation of TSTP raised an exception: Type unification failed: Variable ?'X2.0::type not of sort ord*)
-apply (metis bigo_zero)
-  apply (unfold func_zero)
-  apply (rule ext)
-  apply (simp split: split_max)
-done
-
-
-declare [[ atp_problem_prefix = "BigO__bigo_lesso2" ]]
-lemma bigo_lesso2: "f =o g +o O(h) ==>
-    ALL x. 0 <= k x ==> ALL x. k x <= f x ==>
-      k <o g =o O(h)"
-  apply (unfold lesso_def)
-  apply (rule bigo_lesseq4)
-  apply (erule set_plus_imp_minus)
-  apply (rule allI)
-  apply (rule le_maxI2)
-  apply (rule allI)
-  apply (subst fun_diff_def)
-apply (erule thin_rl)
-(*sledgehammer*);  
-  apply (case_tac "0 <= k x - g x")
-  prefer 2 (*re-order subgoals because I don't know what to put after a structured proof*)
-   apply (metis abs_ge_zero abs_minus_commute linorder_linear min_max.sup_absorb1 min_max.sup_commute)
-proof (neg_clausify)
-fix x
-assume 0: "\<And>A. k A \<le> f A"
-have 1: "\<And>(X1\<Colon>'b\<Colon>ordered_idom) X2. \<not> max X1 X2 < X1"
-  by (metis linorder_not_less le_maxI1)  (*sort inserted by hand*)
-assume 2: "(0\<Colon>'b) \<le> k x - g x"
-have 3: "\<not> k x - g x < (0\<Colon>'b)"
-  by (metis 2 linorder_not_less)
-have 4: "\<And>X1 X2. min X1 (k X2) \<le> f X2"
-  by (metis min_max.inf_le2 min_max.le_inf_iff min_max.le_iff_inf 0)
-have 5: "\<bar>g x - f x\<bar> = f x - g x"
-  by (metis abs_minus_commute combine_common_factor mult_zero_right minus_add_cancel minus_zero abs_if diff_less_eq min_max.inf_commute 4 linorder_not_le min_max.le_iff_inf 3 diff_less_0_iff_less linorder_not_less)
-have 6: "max (0\<Colon>'b) (k x - g x) = k x - g x"
-  by (metis min_max.le_iff_sup 2)
-assume 7: "\<not> max (k x - g x) (0\<Colon>'b) \<le> \<bar>f x - g x\<bar>"
-have 8: "\<not> k x - g x \<le> f x - g x"
-  by (metis 5 abs_minus_commute 7 min_max.sup_commute 6)
-show "False"
-  by (metis min_max.sup_commute min_max.inf_commute min_max.sup_inf_absorb min_max.le_iff_inf 0 max_diff_distrib_left 1 linorder_not_le 8)
-qed
-
-declare [[ atp_problem_prefix = "BigO__bigo_lesso3" ]]
-lemma bigo_lesso3: "f =o g +o O(h) ==>
-    ALL x. 0 <= k x ==> ALL x. g x <= k x ==>
-      f <o k =o O(h)"
-  apply (unfold lesso_def)
-  apply (rule bigo_lesseq4)
-  apply (erule set_plus_imp_minus)
-  apply (rule allI)
-  apply (rule le_maxI2)
-  apply (rule allI)
-  apply (subst fun_diff_def)
-apply (erule thin_rl) 
-(*sledgehammer*); 
-  apply (case_tac "0 <= f x - k x")
-  apply (simp)
-  apply (subst abs_of_nonneg)
-  apply (drule_tac x = x in spec) back
-using [[ atp_problem_prefix = "BigO__bigo_lesso3_simpler" ]]
-apply (metis diff_less_0_iff_less linorder_not_le not_leE uminus_add_conv_diff xt1(12) xt1(6))
-apply (metis add_minus_cancel diff_le_eq le_diff_eq uminus_add_conv_diff)
-apply (metis abs_ge_zero linorder_linear min_max.sup_absorb1 min_max.sup_commute)
-done
-
-lemma bigo_lesso4: "f <o g =o O(k::'a=>'b::ordered_field) ==>
-    g =o h +o O(k) ==> f <o h =o O(k)"
-  apply (unfold lesso_def)
-  apply (drule set_plus_imp_minus)
-  apply (drule bigo_abs5) back
-  apply (simp add: fun_diff_def)
-  apply (drule bigo_useful_add)
-  apply assumption
-  apply (erule bigo_lesseq2) back
-  apply (rule allI)
-  apply (auto simp add: func_plus fun_diff_def algebra_simps
-    split: split_max abs_split)
-done
-
-declare [[ atp_problem_prefix = "BigO__bigo_lesso5" ]]
-lemma bigo_lesso5: "f <o g =o O(h) ==>
-    EX C. ALL x. f x <= g x + C * abs(h x)"
-  apply (simp only: lesso_def bigo_alt_def)
-  apply clarsimp
-  apply (metis abs_if abs_mult add_commute diff_le_eq less_not_permute)  
-done
-
-end
--- a/src/HOL/MetisExamples/Message.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,812 +0,0 @@
-(*  Title:      HOL/MetisTest/Message.thy
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-
-Testing the metis method.
-*)
-
-theory Message imports Main begin
-
-(*Needed occasionally with spy_analz_tac, e.g. in analz_insert_Key_newK*)
-lemma strange_Un_eq [simp]: "A \<union> (B \<union> A) = B \<union> A"
-by blast
-
-types 
-  key = nat
-
-consts
-  all_symmetric :: bool        --{*true if all keys are symmetric*}
-  invKey        :: "key=>key"  --{*inverse of a symmetric key*}
-
-specification (invKey)
-  invKey [simp]: "invKey (invKey K) = K"
-  invKey_symmetric: "all_symmetric --> invKey = id"
-    by (rule exI [of _ id], auto)
-
-
-text{*The inverse of a symmetric key is itself; that of a public key
-      is the private key and vice versa*}
-
-constdefs
-  symKeys :: "key set"
-  "symKeys == {K. invKey K = K}"
-
-datatype  --{*We allow any number of friendly agents*}
-  agent = Server | Friend nat | Spy
-
-datatype
-     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*}
-
-
-text{*Concrete syntax: messages appear as {|A,B,NA|}, etc...*}
-syntax
-  "@MTuple"      :: "['a, args] => 'a * 'b"       ("(2{|_,/ _|})")
-
-syntax (xsymbols)
-  "@MTuple"      :: "['a, args] => 'a * 'b"       ("(2\<lbrace>_,/ _\<rbrace>)")
-
-translations
-  "{|x, y, z|}"   == "{|x, {|y, z|}|}"
-  "{|x, y|}"      == "MPair x y"
-
-
-constdefs
-  HPair :: "[msg,msg] => msg"                       ("(4Hash[_] /_)" [0, 1000])
-    --{*Message Y paired with a MAC computed with the help of X*}
-    "Hash[X] Y == {| Hash{|X,Y|}, Y|}"
-
-  keysFor :: "msg set => key set"
-    --{*Keys useful to decrypt elements of a message set*}
-  "keysFor H == invKey ` {K. \<exists>X. Crypt K X \<in> H}"
-
-
-subsubsection{*Inductive Definition of All Parts" of a Message*}
-
-inductive_set
-  parts :: "msg set => msg set"
-  for H :: "msg set"
-  where
-    Inj [intro]:               "X \<in> H ==> X \<in> parts H"
-  | Fst:         "{|X,Y|}   \<in> parts H ==> X \<in> parts H"
-  | Snd:         "{|X,Y|}   \<in> parts H ==> Y \<in> parts H"
-  | Body:        "Crypt K X \<in> parts H ==> X \<in> parts H"
-
-
-declare [[ atp_problem_prefix = "Message__parts_mono" ]]
-lemma parts_mono: "G \<subseteq> H ==> parts(G) \<subseteq> parts(H)"
-apply auto
-apply (erule parts.induct) 
-apply (metis Inj set_mp)
-apply (metis Fst)
-apply (metis Snd)
-apply (metis Body)
-done
-
-
-text{*Equations hold because constructors are injective.*}
-lemma Friend_image_eq [simp]: "(Friend x \<in> Friend`A) = (x:A)"
-by auto
-
-lemma Key_image_eq [simp]: "(Key x \<in> Key`A) = (x\<in>A)"
-by auto
-
-lemma Nonce_Key_image_eq [simp]: "(Nonce x \<notin> Key`A)"
-by auto
-
-
-subsubsection{*Inverse of keys *}
-
-declare [[ atp_problem_prefix = "Message__invKey_eq" ]]
-lemma invKey_eq [simp]: "(invKey K = invKey K') = (K=K')"
-by (metis invKey)
-
-
-subsection{*keysFor operator*}
-
-lemma keysFor_empty [simp]: "keysFor {} = {}"
-by (unfold keysFor_def, blast)
-
-lemma keysFor_Un [simp]: "keysFor (H \<union> H') = keysFor H \<union> keysFor H'"
-by (unfold keysFor_def, blast)
-
-lemma keysFor_UN [simp]: "keysFor (\<Union>i\<in>A. H i) = (\<Union>i\<in>A. keysFor (H i))"
-by (unfold keysFor_def, blast)
-
-text{*Monotonicity*}
-lemma keysFor_mono: "G \<subseteq> H ==> keysFor(G) \<subseteq> keysFor(H)"
-by (unfold keysFor_def, blast)
-
-lemma keysFor_insert_Agent [simp]: "keysFor (insert (Agent A) H) = keysFor H"
-by (unfold keysFor_def, auto)
-
-lemma keysFor_insert_Nonce [simp]: "keysFor (insert (Nonce N) H) = keysFor H"
-by (unfold keysFor_def, auto)
-
-lemma keysFor_insert_Number [simp]: "keysFor (insert (Number N) H) = keysFor H"
-by (unfold keysFor_def, auto)
-
-lemma keysFor_insert_Key [simp]: "keysFor (insert (Key K) H) = keysFor H"
-by (unfold keysFor_def, auto)
-
-lemma keysFor_insert_Hash [simp]: "keysFor (insert (Hash X) H) = keysFor H"
-by (unfold keysFor_def, auto)
-
-lemma keysFor_insert_MPair [simp]: "keysFor (insert {|X,Y|} H) = keysFor H"
-by (unfold keysFor_def, auto)
-
-lemma keysFor_insert_Crypt [simp]: 
-    "keysFor (insert (Crypt K X) H) = insert (invKey K) (keysFor H)"
-by (unfold keysFor_def, auto)
-
-lemma keysFor_image_Key [simp]: "keysFor (Key`E) = {}"
-by (unfold keysFor_def, auto)
-
-lemma Crypt_imp_invKey_keysFor: "Crypt K X \<in> H ==> invKey K \<in> keysFor H"
-by (unfold keysFor_def, blast)
-
-
-subsection{*Inductive relation "parts"*}
-
-lemma MPair_parts:
-     "[| {|X,Y|} \<in> parts H;        
-         [| X \<in> parts H; Y \<in> parts H |] ==> P |] ==> P"
-by (blast dest: parts.Fst parts.Snd) 
-
-    declare MPair_parts [elim!]  parts.Body [dest!]
-text{*NB These two rules are UNSAFE in the formal sense, as they discard the
-     compound message.  They work well on THIS FILE.  
-  @{text MPair_parts} is left as SAFE because it speeds up proofs.
-  The Crypt rule is normally kept UNSAFE to avoid breaking up certificates.*}
-
-lemma parts_increasing: "H \<subseteq> parts(H)"
-by blast
-
-lemmas parts_insertI = subset_insertI [THEN parts_mono, THEN subsetD, standard]
-
-lemma parts_empty [simp]: "parts{} = {}"
-apply safe
-apply (erule parts.induct)
-apply blast+
-done
-
-lemma parts_emptyE [elim!]: "X\<in> parts{} ==> P"
-by simp
-
-text{*WARNING: loops if H = {Y}, therefore must not be repeated!*}
-lemma parts_singleton: "X\<in> parts H ==> \<exists>Y\<in>H. X\<in> parts {Y}"
-apply (erule parts.induct)
-apply fast+
-done
-
-
-subsubsection{*Unions *}
-
-lemma parts_Un_subset1: "parts(G) \<union> parts(H) \<subseteq> parts(G \<union> H)"
-by (intro Un_least parts_mono Un_upper1 Un_upper2)
-
-lemma parts_Un_subset2: "parts(G \<union> H) \<subseteq> parts(G) \<union> parts(H)"
-apply (rule subsetI)
-apply (erule parts.induct, blast+)
-done
-
-lemma parts_Un [simp]: "parts(G \<union> H) = parts(G) \<union> parts(H)"
-by (intro equalityI parts_Un_subset1 parts_Un_subset2)
-
-lemma parts_insert: "parts (insert X H) = parts {X} \<union> parts H"
-apply (subst insert_is_Un [of _ H])
-apply (simp only: parts_Un)
-done
-
-declare [[ atp_problem_prefix = "Message__parts_insert_two" ]]
-lemma parts_insert2:
-     "parts (insert X (insert Y H)) = parts {X} \<union> parts {Y} \<union> parts H"
-by (metis Un_commute Un_empty_left Un_empty_right Un_insert_left Un_insert_right parts_Un)
-
-
-lemma parts_UN_subset1: "(\<Union>x\<in>A. parts(H x)) \<subseteq> parts(\<Union>x\<in>A. H x)"
-by (intro UN_least parts_mono UN_upper)
-
-lemma parts_UN_subset2: "parts(\<Union>x\<in>A. H x) \<subseteq> (\<Union>x\<in>A. parts(H x))"
-apply (rule subsetI)
-apply (erule parts.induct, blast+)
-done
-
-lemma parts_UN [simp]: "parts(\<Union>x\<in>A. H x) = (\<Union>x\<in>A. parts(H x))"
-by (intro equalityI parts_UN_subset1 parts_UN_subset2)
-
-text{*Added to simplify arguments to parts, analz and synth.
-  NOTE: the UN versions are no longer used!*}
-
-
-text{*This allows @{text blast} to simplify occurrences of 
-  @{term "parts(G\<union>H)"} in the assumption.*}
-lemmas in_parts_UnE = parts_Un [THEN equalityD1, THEN subsetD, THEN UnE] 
-declare in_parts_UnE [elim!]
-
-lemma parts_insert_subset: "insert X (parts H) \<subseteq> parts(insert X H)"
-by (blast intro: parts_mono [THEN [2] rev_subsetD])
-
-subsubsection{*Idempotence and transitivity *}
-
-lemma parts_partsD [dest!]: "X\<in> parts (parts H) ==> X\<in> parts H"
-by (erule parts.induct, blast+)
-
-lemma parts_idem [simp]: "parts (parts H) = parts H"
-by blast
-
-declare [[ atp_problem_prefix = "Message__parts_subset_iff" ]]
-lemma parts_subset_iff [simp]: "(parts G \<subseteq> parts H) = (G \<subseteq> parts H)"
-apply (rule iffI) 
-apply (metis Un_absorb1 Un_subset_iff parts_Un parts_increasing)
-apply (metis parts_idem parts_mono)
-done
-
-lemma parts_trans: "[| X\<in> parts G;  G \<subseteq> parts H |] ==> X\<in> parts H"
-by (blast dest: parts_mono); 
-
-
-declare [[ atp_problem_prefix = "Message__parts_cut" ]]
-lemma parts_cut: "[|Y\<in> parts(insert X G);  X\<in> parts H|] ==> Y\<in> parts(G \<union> H)"
-by (metis Un_subset_iff insert_subset parts_increasing parts_trans) 
-
-
-
-subsubsection{*Rewrite rules for pulling out atomic messages *}
-
-lemmas parts_insert_eq_I = equalityI [OF subsetI parts_insert_subset]
-
-
-lemma parts_insert_Agent [simp]:
-     "parts (insert (Agent agt) H) = insert (Agent agt) (parts H)"
-apply (rule parts_insert_eq_I) 
-apply (erule parts.induct, auto) 
-done
-
-lemma parts_insert_Nonce [simp]:
-     "parts (insert (Nonce N) H) = insert (Nonce N) (parts H)"
-apply (rule parts_insert_eq_I) 
-apply (erule parts.induct, auto) 
-done
-
-lemma parts_insert_Number [simp]:
-     "parts (insert (Number N) H) = insert (Number N) (parts H)"
-apply (rule parts_insert_eq_I) 
-apply (erule parts.induct, auto) 
-done
-
-lemma parts_insert_Key [simp]:
-     "parts (insert (Key K) H) = insert (Key K) (parts H)"
-apply (rule parts_insert_eq_I) 
-apply (erule parts.induct, auto) 
-done
-
-lemma parts_insert_Hash [simp]:
-     "parts (insert (Hash X) H) = insert (Hash X) (parts H)"
-apply (rule parts_insert_eq_I) 
-apply (erule parts.induct, auto) 
-done
-
-lemma parts_insert_Crypt [simp]:
-     "parts (insert (Crypt K X) H) =  
-          insert (Crypt K X) (parts (insert X H))"
-apply (rule equalityI)
-apply (rule subsetI)
-apply (erule parts.induct, auto)
-apply (blast intro: parts.Body)
-done
-
-lemma parts_insert_MPair [simp]:
-     "parts (insert {|X,Y|} H) =  
-          insert {|X,Y|} (parts (insert X (insert Y H)))"
-apply (rule equalityI)
-apply (rule subsetI)
-apply (erule parts.induct, auto)
-apply (blast intro: parts.Fst parts.Snd)+
-done
-
-lemma parts_image_Key [simp]: "parts (Key`N) = Key`N"
-apply auto
-apply (erule parts.induct, auto)
-done
-
-
-declare [[ atp_problem_prefix = "Message__msg_Nonce_supply" ]]
-lemma msg_Nonce_supply: "\<exists>N. \<forall>n. N\<le>n --> Nonce n \<notin> parts {msg}"
-apply (induct_tac "msg") 
-apply (simp_all add: parts_insert2)
-apply (metis Suc_n_not_le_n)
-apply (metis le_trans linorder_linear)
-done
-
-subsection{*Inductive relation "analz"*}
-
-text{*Inductive definition of "analz" -- what can be broken down from a set of
-    messages, including keys.  A form of downward closure.  Pairs can
-    be taken apart; messages decrypted with known keys.  *}
-
-inductive_set
-  analz :: "msg set => msg set"
-  for H :: "msg set"
-  where
-    Inj [intro,simp] :    "X \<in> H ==> X \<in> analz H"
-  | Fst:     "{|X,Y|} \<in> analz H ==> X \<in> analz H"
-  | Snd:     "{|X,Y|} \<in> analz H ==> Y \<in> analz H"
-  | Decrypt [dest]: 
-             "[|Crypt K X \<in> analz H; Key(invKey K): analz H|] ==> X \<in> analz H"
-
-
-text{*Monotonicity; Lemma 1 of Lowe's paper*}
-lemma analz_mono: "G\<subseteq>H ==> analz(G) \<subseteq> analz(H)"
-apply auto
-apply (erule analz.induct) 
-apply (auto dest: analz.Fst analz.Snd) 
-done
-
-text{*Making it safe speeds up proofs*}
-lemma MPair_analz [elim!]:
-     "[| {|X,Y|} \<in> analz H;        
-             [| X \<in> analz H; Y \<in> analz H |] ==> P   
-          |] ==> P"
-by (blast dest: analz.Fst analz.Snd)
-
-lemma analz_increasing: "H \<subseteq> analz(H)"
-by blast
-
-lemma analz_subset_parts: "analz H \<subseteq> parts H"
-apply (rule subsetI)
-apply (erule analz.induct, blast+)
-done
-
-lemmas analz_into_parts = analz_subset_parts [THEN subsetD, standard]
-
-lemmas not_parts_not_analz = analz_subset_parts [THEN contra_subsetD, standard]
-
-
-declare [[ atp_problem_prefix = "Message__parts_analz" ]]
-lemma parts_analz [simp]: "parts (analz H) = parts H"
-apply (rule equalityI)
-apply (metis analz_subset_parts parts_subset_iff)
-apply (metis analz_increasing parts_mono)
-done
-
-
-lemma analz_parts [simp]: "analz (parts H) = parts H"
-apply auto
-apply (erule analz.induct, auto)
-done
-
-lemmas analz_insertI = subset_insertI [THEN analz_mono, THEN [2] rev_subsetD, standard]
-
-subsubsection{*General equational properties *}
-
-lemma analz_empty [simp]: "analz{} = {}"
-apply safe
-apply (erule analz.induct, blast+)
-done
-
-text{*Converse fails: we can analz more from the union than from the 
-  separate parts, as a key in one might decrypt a message in the other*}
-lemma analz_Un: "analz(G) \<union> analz(H) \<subseteq> analz(G \<union> H)"
-by (intro Un_least analz_mono Un_upper1 Un_upper2)
-
-lemma analz_insert: "insert X (analz H) \<subseteq> analz(insert X H)"
-by (blast intro: analz_mono [THEN [2] rev_subsetD])
-
-subsubsection{*Rewrite rules for pulling out atomic messages *}
-
-lemmas analz_insert_eq_I = equalityI [OF subsetI analz_insert]
-
-lemma analz_insert_Agent [simp]:
-     "analz (insert (Agent agt) H) = insert (Agent agt) (analz H)"
-apply (rule analz_insert_eq_I) 
-apply (erule analz.induct, auto) 
-done
-
-lemma analz_insert_Nonce [simp]:
-     "analz (insert (Nonce N) H) = insert (Nonce N) (analz H)"
-apply (rule analz_insert_eq_I) 
-apply (erule analz.induct, auto) 
-done
-
-lemma analz_insert_Number [simp]:
-     "analz (insert (Number N) H) = insert (Number N) (analz H)"
-apply (rule analz_insert_eq_I) 
-apply (erule analz.induct, auto) 
-done
-
-lemma analz_insert_Hash [simp]:
-     "analz (insert (Hash X) H) = insert (Hash X) (analz H)"
-apply (rule analz_insert_eq_I) 
-apply (erule analz.induct, auto) 
-done
-
-text{*Can only pull out Keys if they are not needed to decrypt the rest*}
-lemma analz_insert_Key [simp]: 
-    "K \<notin> keysFor (analz H) ==>   
-          analz (insert (Key K) H) = insert (Key K) (analz H)"
-apply (unfold keysFor_def)
-apply (rule analz_insert_eq_I) 
-apply (erule analz.induct, auto) 
-done
-
-lemma analz_insert_MPair [simp]:
-     "analz (insert {|X,Y|} H) =  
-          insert {|X,Y|} (analz (insert X (insert Y H)))"
-apply (rule equalityI)
-apply (rule subsetI)
-apply (erule analz.induct, auto)
-apply (erule analz.induct)
-apply (blast intro: analz.Fst analz.Snd)+
-done
-
-text{*Can pull out enCrypted message if the Key is not known*}
-lemma analz_insert_Crypt:
-     "Key (invKey K) \<notin> analz H 
-      ==> analz (insert (Crypt K X) H) = insert (Crypt K X) (analz H)"
-apply (rule analz_insert_eq_I) 
-apply (erule analz.induct, auto) 
-
-done
-
-lemma lemma1: "Key (invKey K) \<in> analz H ==>   
-               analz (insert (Crypt K X) H) \<subseteq>  
-               insert (Crypt K X) (analz (insert X H))" 
-apply (rule subsetI)
-apply (erule_tac x = x in analz.induct, auto)
-done
-
-lemma lemma2: "Key (invKey K) \<in> analz H ==>   
-               insert (Crypt K X) (analz (insert X H)) \<subseteq>  
-               analz (insert (Crypt K X) H)"
-apply auto
-apply (erule_tac x = x in analz.induct, auto)
-apply (blast intro: analz_insertI analz.Decrypt)
-done
-
-lemma analz_insert_Decrypt:
-     "Key (invKey K) \<in> analz H ==>   
-               analz (insert (Crypt K X) H) =  
-               insert (Crypt K X) (analz (insert X H))"
-by (intro equalityI lemma1 lemma2)
-
-text{*Case analysis: either the message is secure, or it is not! Effective,
-but can cause subgoals to blow up! Use with @{text "split_if"}; apparently
-@{text "split_tac"} does not cope with patterns such as @{term"analz (insert
-(Crypt K X) H)"} *} 
-lemma analz_Crypt_if [simp]:
-     "analz (insert (Crypt K X) H) =                 
-          (if (Key (invKey K) \<in> analz H)                 
-           then insert (Crypt K X) (analz (insert X H))  
-           else insert (Crypt K X) (analz H))"
-by (simp add: analz_insert_Crypt analz_insert_Decrypt)
-
-
-text{*This rule supposes "for the sake of argument" that we have the key.*}
-lemma analz_insert_Crypt_subset:
-     "analz (insert (Crypt K X) H) \<subseteq>   
-           insert (Crypt K X) (analz (insert X H))"
-apply (rule subsetI)
-apply (erule analz.induct, auto)
-done
-
-
-lemma analz_image_Key [simp]: "analz (Key`N) = Key`N"
-apply auto
-apply (erule analz.induct, auto)
-done
-
-
-subsubsection{*Idempotence and transitivity *}
-
-lemma analz_analzD [dest!]: "X\<in> analz (analz H) ==> X\<in> analz H"
-by (erule analz.induct, blast+)
-
-lemma analz_idem [simp]: "analz (analz H) = analz H"
-by blast
-
-lemma analz_subset_iff [simp]: "(analz G \<subseteq> analz H) = (G \<subseteq> analz H)"
-apply (rule iffI)
-apply (iprover intro: subset_trans analz_increasing)  
-apply (frule analz_mono, simp) 
-done
-
-lemma analz_trans: "[| X\<in> analz G;  G \<subseteq> analz H |] ==> X\<in> analz H"
-by (drule analz_mono, blast)
-
-
-declare [[ atp_problem_prefix = "Message__analz_cut" ]]
-    declare analz_trans[intro]
-lemma analz_cut: "[| Y\<in> analz (insert X H);  X\<in> analz H |] ==> Y\<in> analz H"
-(*TOO SLOW
-by (metis analz_idem analz_increasing analz_mono insert_absorb insert_mono insert_subset) --{*317s*}
-??*)
-by (erule analz_trans, blast)
-
-
-text{*This rewrite rule helps in the simplification of messages that involve
-  the forwarding of unknown components (X).  Without it, removing occurrences
-  of X can be very complicated. *}
-lemma analz_insert_eq: "X\<in> analz H ==> analz (insert X H) = analz H"
-by (blast intro: analz_cut analz_insertI)
-
-
-text{*A congruence rule for "analz" *}
-
-declare [[ atp_problem_prefix = "Message__analz_subset_cong" ]]
-lemma analz_subset_cong:
-     "[| analz G \<subseteq> analz G'; analz H \<subseteq> analz H' |] 
-      ==> analz (G \<union> H) \<subseteq> analz (G' \<union> H')"
-apply simp
-apply (metis Un_absorb2 Un_commute Un_subset_iff Un_upper1 Un_upper2 analz_mono)
-done
-
-
-lemma analz_cong:
-     "[| analz G = analz G'; analz H = analz H'  
-               |] ==> analz (G \<union> H) = analz (G' \<union> H')"
-by (intro equalityI analz_subset_cong, simp_all) 
-
-lemma analz_insert_cong:
-     "analz H = analz H' ==> analz(insert X H) = analz(insert X H')"
-by (force simp only: insert_def intro!: analz_cong)
-
-text{*If there are no pairs or encryptions then analz does nothing*}
-lemma analz_trivial:
-     "[| \<forall>X Y. {|X,Y|} \<notin> H;  \<forall>X K. Crypt K X \<notin> H |] ==> analz H = H"
-apply safe
-apply (erule analz.induct, blast+)
-done
-
-text{*These two are obsolete (with a single Spy) but cost little to prove...*}
-lemma analz_UN_analz_lemma:
-     "X\<in> analz (\<Union>i\<in>A. analz (H i)) ==> X\<in> analz (\<Union>i\<in>A. H i)"
-apply (erule analz.induct)
-apply (blast intro: analz_mono [THEN [2] rev_subsetD])+
-done
-
-lemma analz_UN_analz [simp]: "analz (\<Union>i\<in>A. analz (H i)) = analz (\<Union>i\<in>A. H i)"
-by (blast intro: analz_UN_analz_lemma analz_mono [THEN [2] rev_subsetD])
-
-
-subsection{*Inductive relation "synth"*}
-
-text{*Inductive definition of "synth" -- what can be built up from a set of
-    messages.  A form of upward closure.  Pairs can be built, messages
-    encrypted with known keys.  Agent names are public domain.
-    Numbers can be guessed, but Nonces cannot be.  *}
-
-inductive_set
-  synth :: "msg set => msg set"
-  for H :: "msg set"
-  where
-    Inj    [intro]:   "X \<in> H ==> X \<in> synth H"
-  | Agent  [intro]:   "Agent agt \<in> synth H"
-  | Number [intro]:   "Number n  \<in> synth H"
-  | Hash   [intro]:   "X \<in> synth H ==> Hash X \<in> synth H"
-  | MPair  [intro]:   "[|X \<in> synth H;  Y \<in> synth H|] ==> {|X,Y|} \<in> synth H"
-  | Crypt  [intro]:   "[|X \<in> synth H;  Key(K) \<in> H|] ==> Crypt K X \<in> synth H"
-
-text{*Monotonicity*}
-lemma synth_mono: "G\<subseteq>H ==> synth(G) \<subseteq> synth(H)"
-  by (auto, erule synth.induct, auto)  
-
-text{*NO @{text Agent_synth}, as any Agent name can be synthesized.  
-  The same holds for @{term Number}*}
-inductive_cases Nonce_synth [elim!]: "Nonce n \<in> synth H"
-inductive_cases Key_synth   [elim!]: "Key K \<in> synth H"
-inductive_cases Hash_synth  [elim!]: "Hash X \<in> synth H"
-inductive_cases MPair_synth [elim!]: "{|X,Y|} \<in> synth H"
-inductive_cases Crypt_synth [elim!]: "Crypt K X \<in> synth H"
-
-
-lemma synth_increasing: "H \<subseteq> synth(H)"
-by blast
-
-subsubsection{*Unions *}
-
-text{*Converse fails: we can synth more from the union than from the 
-  separate parts, building a compound message using elements of each.*}
-lemma synth_Un: "synth(G) \<union> synth(H) \<subseteq> synth(G \<union> H)"
-by (intro Un_least synth_mono Un_upper1 Un_upper2)
-
-
-declare [[ atp_problem_prefix = "Message__synth_insert" ]]
- 
-lemma synth_insert: "insert X (synth H) \<subseteq> synth(insert X H)"
-by (metis insert_iff insert_subset subset_insertI synth.Inj synth_mono)
-
-subsubsection{*Idempotence and transitivity *}
-
-lemma synth_synthD [dest!]: "X\<in> synth (synth H) ==> X\<in> synth H"
-by (erule synth.induct, blast+)
-
-lemma synth_idem: "synth (synth H) = synth H"
-by blast
-
-lemma synth_subset_iff [simp]: "(synth G \<subseteq> synth H) = (G \<subseteq> synth H)"
-apply (rule iffI)
-apply (iprover intro: subset_trans synth_increasing)  
-apply (frule synth_mono, simp add: synth_idem) 
-done
-
-lemma synth_trans: "[| X\<in> synth G;  G \<subseteq> synth H |] ==> X\<in> synth H"
-by (drule synth_mono, blast)
-
-declare [[ atp_problem_prefix = "Message__synth_cut" ]]
-lemma synth_cut: "[| Y\<in> synth (insert X H);  X\<in> synth H |] ==> Y\<in> synth H"
-(*TOO SLOW
-by (metis insert_absorb insert_mono insert_subset synth_idem synth_increasing synth_mono)
-*)
-by (erule synth_trans, blast)
-
-
-lemma Agent_synth [simp]: "Agent A \<in> synth H"
-by blast
-
-lemma Number_synth [simp]: "Number n \<in> synth H"
-by blast
-
-lemma Nonce_synth_eq [simp]: "(Nonce N \<in> synth H) = (Nonce N \<in> H)"
-by blast
-
-lemma Key_synth_eq [simp]: "(Key K \<in> synth H) = (Key K \<in> H)"
-by blast
-
-lemma Crypt_synth_eq [simp]:
-     "Key K \<notin> H ==> (Crypt K X \<in> synth H) = (Crypt K X \<in> H)"
-by blast
-
-
-lemma keysFor_synth [simp]: 
-    "keysFor (synth H) = keysFor H \<union> invKey`{K. Key K \<in> H}"
-by (unfold keysFor_def, blast)
-
-
-subsubsection{*Combinations of parts, analz and synth *}
-
-declare [[ atp_problem_prefix = "Message__parts_synth" ]]
-lemma parts_synth [simp]: "parts (synth H) = parts H \<union> synth H"
-apply (rule equalityI)
-apply (rule subsetI)
-apply (erule parts.induct)
-apply (metis UnCI)
-apply (metis MPair_synth UnCI UnE insert_absorb insert_subset parts.Fst parts_increasing)
-apply (metis MPair_synth UnCI UnE insert_absorb insert_subset parts.Snd parts_increasing)
-apply (metis Body Crypt_synth UnCI UnE insert_absorb insert_subset parts_increasing)
-apply (metis Un_subset_iff parts_increasing parts_mono synth_increasing)
-done
-
-
-
-
-declare [[ atp_problem_prefix = "Message__analz_analz_Un" ]]
-lemma analz_analz_Un [simp]: "analz (analz G \<union> H) = analz (G \<union> H)"
-apply (rule equalityI);
-apply (metis analz_idem analz_subset_cong order_eq_refl)
-apply (metis analz_increasing analz_subset_cong order_eq_refl)
-done
-
-declare [[ atp_problem_prefix = "Message__analz_synth_Un" ]]
-    declare analz_mono [intro] analz.Fst [intro] analz.Snd [intro] Un_least [intro]
-lemma analz_synth_Un [simp]: "analz (synth G \<union> H) = analz (G \<union> H) \<union> synth G"
-apply (rule equalityI)
-apply (rule subsetI)
-apply (erule analz.induct)
-apply (metis UnCI UnE Un_commute analz.Inj)
-apply (metis MPair_synth UnCI UnE Un_commute Un_upper1 analz.Fst analz_increasing analz_mono insert_absorb insert_subset)
-apply (metis MPair_synth UnCI UnE Un_commute Un_upper1 analz.Snd analz_increasing analz_mono insert_absorb insert_subset)
-apply (blast intro: analz.Decrypt)
-apply blast
-done
-
-
-declare [[ atp_problem_prefix = "Message__analz_synth" ]]
-lemma analz_synth [simp]: "analz (synth H) = analz H \<union> synth H"
-proof (neg_clausify)
-assume 0: "analz (synth H) \<noteq> analz H \<union> synth H"
-have 1: "\<And>X1 X3. sup (analz (sup X3 X1)) (synth X3) = analz (sup (synth X3) X1)"
-  by (metis analz_synth_Un)
-have 2: "sup (analz H) (synth H) \<noteq> analz (synth H)"
-  by (metis 0)
-have 3: "\<And>X1 X3. sup (synth X3) (analz (sup X3 X1)) = analz (sup (synth X3) X1)"
-  by (metis 1 Un_commute)
-have 4: "\<And>X3. sup (synth X3) (analz X3) = analz (sup (synth X3) {})"
-  by (metis 3 Un_empty_right)
-have 5: "\<And>X3. sup (synth X3) (analz X3) = analz (synth X3)"
-  by (metis 4 Un_empty_right)
-have 6: "\<And>X3. sup (analz X3) (synth X3) = analz (synth X3)"
-  by (metis 5 Un_commute)
-show "False"
-  by (metis 2 6)
-qed
-
-
-subsubsection{*For reasoning about the Fake rule in traces *}
-
-declare [[ atp_problem_prefix = "Message__parts_insert_subset_Un" ]]
-lemma parts_insert_subset_Un: "X\<in> G ==> parts(insert X H) \<subseteq> parts G \<union> parts H"
-proof (neg_clausify)
-assume 0: "X \<in> G"
-assume 1: "\<not> parts (insert X H) \<subseteq> parts G \<union> parts H"
-have 2: "\<not> parts (insert X H) \<subseteq> parts (G \<union> H)"
-  by (metis 1 parts_Un)
-have 3: "\<not> insert X H \<subseteq> G \<union> H"
-  by (metis 2 parts_mono)
-have 4: "X \<notin> G \<union> H \<or> \<not> H \<subseteq> G \<union> H"
-  by (metis 3 insert_subset)
-have 5: "X \<notin> G \<union> H"
-  by (metis 4 Un_upper2)
-have 6: "X \<notin> G"
-  by (metis 5 UnCI)
-show "False"
-  by (metis 6 0)
-qed
-
-declare [[ atp_problem_prefix = "Message__Fake_parts_insert" ]]
-lemma Fake_parts_insert:
-     "X \<in> synth (analz H) ==>  
-      parts (insert X H) \<subseteq> synth (analz H) \<union> parts H"
-proof (neg_clausify)
-assume 0: "X \<in> synth (analz H)"
-assume 1: "\<not> parts (insert X H) \<subseteq> synth (analz H) \<union> parts H"
-have 2: "\<And>X3. parts X3 \<union> synth (analz X3) = parts (synth (analz X3))"
-  by (metis parts_synth parts_analz)
-have 3: "\<And>X3. analz X3 \<union> synth (analz X3) = analz (synth (analz X3))"
-  by (metis analz_synth analz_idem)
-have 4: "\<And>X3. analz X3 \<subseteq> analz (synth X3)"
-  by (metis Un_upper1 analz_synth)
-have 5: "\<not> parts (insert X H) \<subseteq> parts H \<union> synth (analz H)"
-  by (metis 1 Un_commute)
-have 6: "\<not> parts (insert X H) \<subseteq> parts (synth (analz H))"
-  by (metis 5 2)
-have 7: "\<not> insert X H \<subseteq> synth (analz H)"
-  by (metis 6 parts_mono)
-have 8: "X \<notin> synth (analz H) \<or> \<not> H \<subseteq> synth (analz H)"
-  by (metis 7 insert_subset)
-have 9: "\<not> H \<subseteq> synth (analz H)"
-  by (metis 8 0)
-have 10: "\<And>X3. X3 \<subseteq> analz (synth X3)"
-  by (metis analz_subset_iff 4)
-have 11: "\<And>X3. X3 \<subseteq> analz (synth (analz X3))"
-  by (metis analz_subset_iff 10)
-have 12: "\<And>X3. analz (synth (analz X3)) = synth (analz X3) \<or>
-     \<not> analz X3 \<subseteq> synth (analz X3)"
-  by (metis Un_absorb1 3)
-have 13: "\<And>X3. analz (synth (analz X3)) = synth (analz X3)"
-  by (metis 12 synth_increasing)
-have 14: "\<And>X3. X3 \<subseteq> synth (analz X3)"
-  by (metis 11 13)
-show "False"
-  by (metis 9 14)
-qed
-
-lemma Fake_parts_insert_in_Un:
-     "[|Z \<in> parts (insert X H);  X: synth (analz H)|] 
-      ==> Z \<in>  synth (analz H) \<union> parts H";
-by (blast dest: Fake_parts_insert  [THEN subsetD, dest])
-
-declare [[ atp_problem_prefix = "Message__Fake_analz_insert" ]]
-    declare analz_mono [intro] synth_mono [intro] 
-lemma Fake_analz_insert:
-     "X\<in> synth (analz G) ==>  
-      analz (insert X H) \<subseteq> synth (analz G) \<union> analz (G \<union> H)"
-by (metis Un_commute Un_insert_left Un_insert_right Un_upper1 analz_analz_Un analz_mono analz_synth_Un equalityE insert_absorb order_le_less xt1(12))
-
-declare [[ atp_problem_prefix = "Message__Fake_analz_insert_simpler" ]]
-(*simpler problems?  BUT METIS CAN'T PROVE
-lemma Fake_analz_insert_simpler:
-     "X\<in> synth (analz G) ==>  
-      analz (insert X H) \<subseteq> synth (analz G) \<union> analz (G \<union> H)"
-apply (rule subsetI)
-apply (subgoal_tac "x \<in> analz (synth (analz G) \<union> H) ")
-apply (metis Un_commute analz_analz_Un analz_synth_Un)
-apply (metis Un_commute Un_upper1 Un_upper2 analz_cut analz_increasing analz_mono insert_absorb insert_mono insert_subset)
-done
-*)
-
-end
--- a/src/HOL/MetisExamples/ROOT.ML	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,9 +0,0 @@
-(*  Title:      HOL/MetisExamples/ROOT.ML
-    ID:         $Id$
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-
-Testing the metis method.
-*)
-
-use_thys ["set", "BigO", "Abstraction", "BT", "Message", "Tarski", "TransClosure"];
-
--- a/src/HOL/MetisExamples/Tarski.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1110 +0,0 @@
-(*  Title:      HOL/MetisTest/Tarski.thy
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-
-Testing the metis method.
-*)
-
-header {* The Full Theorem of Tarski *}
-
-theory Tarski
-imports Main FuncSet
-begin
-
-(*Many of these higher-order problems appear to be impossible using the
-current linkup. They often seem to need either higher-order unification
-or explicit reasoning about connectives such as conjunction. The numerous
-set comprehensions are to blame.*)
-
-
-record 'a potype =
-  pset  :: "'a set"
-  order :: "('a * 'a) set"
-
-constdefs
-  monotone :: "['a => 'a, 'a set, ('a *'a)set] => bool"
-  "monotone f A r == \<forall>x\<in>A. \<forall>y\<in>A. (x, y): r --> ((f x), (f y)) : r"
-
-  least :: "['a => bool, 'a potype] => 'a"
-  "least P po == @ x. x: pset po & P x &
-                       (\<forall>y \<in> pset po. P y --> (x,y): order po)"
-
-  greatest :: "['a => bool, 'a potype] => 'a"
-  "greatest P po == @ x. x: pset po & P x &
-                          (\<forall>y \<in> pset po. P y --> (y,x): order po)"
-
-  lub  :: "['a set, 'a potype] => 'a"
-  "lub S po == least (%x. \<forall>y\<in>S. (y,x): order po) po"
-
-  glb  :: "['a set, 'a potype] => 'a"
-  "glb S po == greatest (%x. \<forall>y\<in>S. (x,y): order po) po"
-
-  isLub :: "['a set, 'a potype, 'a] => bool"
-  "isLub S po == %L. (L: pset po & (\<forall>y\<in>S. (y,L): order po) &
-                   (\<forall>z\<in>pset po. (\<forall>y\<in>S. (y,z): order po) --> (L,z): order po))"
-
-  isGlb :: "['a set, 'a potype, 'a] => bool"
-  "isGlb S po == %G. (G: pset po & (\<forall>y\<in>S. (G,y): order po) &
-                 (\<forall>z \<in> pset po. (\<forall>y\<in>S. (z,y): order po) --> (z,G): order po))"
-
-  "fix"    :: "[('a => 'a), 'a set] => 'a set"
-  "fix f A  == {x. x: A & f x = x}"
-
-  interval :: "[('a*'a) set,'a, 'a ] => 'a set"
-  "interval r a b == {x. (a,x): r & (x,b): r}"
-
-constdefs
-  Bot :: "'a potype => 'a"
-  "Bot po == least (%x. True) po"
-
-  Top :: "'a potype => 'a"
-  "Top po == greatest (%x. True) po"
-
-  PartialOrder :: "('a potype) set"
-  "PartialOrder == {P. refl_on (pset P) (order P) & antisym (order P) &
-                       trans (order P)}"
-
-  CompleteLattice :: "('a potype) set"
-  "CompleteLattice == {cl. cl: PartialOrder &
-                        (\<forall>S. S \<subseteq> pset cl --> (\<exists>L. isLub S cl L)) &
-                        (\<forall>S. S \<subseteq> pset cl --> (\<exists>G. isGlb S cl G))}"
-
-  induced :: "['a set, ('a * 'a) set] => ('a *'a)set"
-  "induced A r == {(a,b). a : A & b: A & (a,b): r}"
-
-constdefs
-  sublattice :: "('a potype * 'a set)set"
-  "sublattice ==
-      SIGMA cl: CompleteLattice.
-          {S. S \<subseteq> pset cl &
-           (| pset = S, order = induced S (order cl) |): CompleteLattice }"
-
-syntax
-  "@SL"  :: "['a set, 'a potype] => bool" ("_ <<= _" [51,50]50)
-
-translations
-  "S <<= cl" == "S : sublattice `` {cl}"
-
-constdefs
-  dual :: "'a potype => 'a potype"
-  "dual po == (| pset = pset po, order = converse (order po) |)"
-
-locale PO =
-  fixes cl :: "'a potype"
-    and A  :: "'a set"
-    and r  :: "('a * 'a) set"
-  assumes cl_po:  "cl : PartialOrder"
-  defines A_def: "A == pset cl"
-     and  r_def: "r == order cl"
-
-locale CL = PO +
-  assumes cl_co:  "cl : CompleteLattice"
-
-definition CLF_set :: "('a potype * ('a => 'a)) set" where
-  "CLF_set = (SIGMA cl: CompleteLattice.
-            {f. f: pset cl -> pset cl & monotone f (pset cl) (order cl)})"
-
-locale CLF = CL +
-  fixes f :: "'a => 'a"
-    and P :: "'a set"
-  assumes f_cl:  "(cl,f) : CLF_set" (*was the equivalent "f : CLF``{cl}"*)
-  defines P_def: "P == fix f A"
-
-
-locale Tarski = CLF +
-  fixes Y     :: "'a set"
-    and intY1 :: "'a set"
-    and v     :: "'a"
-  assumes
-    Y_ss: "Y \<subseteq> P"
-  defines
-    intY1_def: "intY1 == interval r (lub Y cl) (Top cl)"
-    and v_def: "v == glb {x. ((%x: intY1. f x) x, x): induced intY1 r &
-                             x: intY1}
-                      (| pset=intY1, order=induced intY1 r|)"
-
-
-subsection {* Partial Order *}
-
-lemma (in PO) PO_imp_refl_on: "refl_on A r"
-apply (insert cl_po)
-apply (simp add: PartialOrder_def A_def r_def)
-done
-
-lemma (in PO) PO_imp_sym: "antisym r"
-apply (insert cl_po)
-apply (simp add: PartialOrder_def r_def)
-done
-
-lemma (in PO) PO_imp_trans: "trans r"
-apply (insert cl_po)
-apply (simp add: PartialOrder_def r_def)
-done
-
-lemma (in PO) reflE: "x \<in> A ==> (x, x) \<in> r"
-apply (insert cl_po)
-apply (simp add: PartialOrder_def refl_on_def A_def r_def)
-done
-
-lemma (in PO) antisymE: "[| (a, b) \<in> r; (b, a) \<in> r |] ==> a = b"
-apply (insert cl_po)
-apply (simp add: PartialOrder_def antisym_def r_def)
-done
-
-lemma (in PO) transE: "[| (a, b) \<in> r; (b, c) \<in> r|] ==> (a,c) \<in> r"
-apply (insert cl_po)
-apply (simp add: PartialOrder_def r_def)
-apply (unfold trans_def, fast)
-done
-
-lemma (in PO) monotoneE:
-     "[| monotone f A r;  x \<in> A; y \<in> A; (x, y) \<in> r |] ==> (f x, f y) \<in> r"
-by (simp add: monotone_def)
-
-lemma (in PO) po_subset_po:
-     "S \<subseteq> A ==> (| pset = S, order = induced S r |) \<in> PartialOrder"
-apply (simp (no_asm) add: PartialOrder_def)
-apply auto
--- {* refl *}
-apply (simp add: refl_on_def induced_def)
-apply (blast intro: reflE)
--- {* antisym *}
-apply (simp add: antisym_def induced_def)
-apply (blast intro: antisymE)
--- {* trans *}
-apply (simp add: trans_def induced_def)
-apply (blast intro: transE)
-done
-
-lemma (in PO) indE: "[| (x, y) \<in> induced S r; S \<subseteq> A |] ==> (x, y) \<in> r"
-by (simp add: add: induced_def)
-
-lemma (in PO) indI: "[| (x, y) \<in> r; x \<in> S; y \<in> S |] ==> (x, y) \<in> induced S r"
-by (simp add: add: induced_def)
-
-lemma (in CL) CL_imp_ex_isLub: "S \<subseteq> A ==> \<exists>L. isLub S cl L"
-apply (insert cl_co)
-apply (simp add: CompleteLattice_def A_def)
-done
-
-declare (in CL) cl_co [simp]
-
-lemma isLub_lub: "(\<exists>L. isLub S cl L) = isLub S cl (lub S cl)"
-by (simp add: lub_def least_def isLub_def some_eq_ex [symmetric])
-
-lemma isGlb_glb: "(\<exists>G. isGlb S cl G) = isGlb S cl (glb S cl)"
-by (simp add: glb_def greatest_def isGlb_def some_eq_ex [symmetric])
-
-lemma isGlb_dual_isLub: "isGlb S cl = isLub S (dual cl)"
-by (simp add: isLub_def isGlb_def dual_def converse_def)
-
-lemma isLub_dual_isGlb: "isLub S cl = isGlb S (dual cl)"
-by (simp add: isLub_def isGlb_def dual_def converse_def)
-
-lemma (in PO) dualPO: "dual cl \<in> PartialOrder"
-apply (insert cl_po)
-apply (simp add: PartialOrder_def dual_def refl_on_converse
-                 trans_converse antisym_converse)
-done
-
-lemma Rdual:
-     "\<forall>S. (S \<subseteq> A -->( \<exists>L. isLub S (| pset = A, order = r|) L))
-      ==> \<forall>S. (S \<subseteq> A --> (\<exists>G. isGlb S (| pset = A, order = r|) G))"
-apply safe
-apply (rule_tac x = "lub {y. y \<in> A & (\<forall>k \<in> S. (y, k) \<in> r)}
-                      (|pset = A, order = r|) " in exI)
-apply (drule_tac x = "{y. y \<in> A & (\<forall>k \<in> S. (y,k) \<in> r) }" in spec)
-apply (drule mp, fast)
-apply (simp add: isLub_lub isGlb_def)
-apply (simp add: isLub_def, blast)
-done
-
-lemma lub_dual_glb: "lub S cl = glb S (dual cl)"
-by (simp add: lub_def glb_def least_def greatest_def dual_def converse_def)
-
-lemma glb_dual_lub: "glb S cl = lub S (dual cl)"
-by (simp add: lub_def glb_def least_def greatest_def dual_def converse_def)
-
-lemma CL_subset_PO: "CompleteLattice \<subseteq> PartialOrder"
-by (simp add: PartialOrder_def CompleteLattice_def, fast)
-
-lemmas CL_imp_PO = CL_subset_PO [THEN subsetD]
-
-declare PO.PO_imp_refl_on  [OF PO.intro [OF CL_imp_PO], simp]
-declare PO.PO_imp_sym   [OF PO.intro [OF CL_imp_PO], simp]
-declare PO.PO_imp_trans [OF PO.intro [OF CL_imp_PO], simp]
-
-lemma (in CL) CO_refl_on: "refl_on A r"
-by (rule PO_imp_refl_on)
-
-lemma (in CL) CO_antisym: "antisym r"
-by (rule PO_imp_sym)
-
-lemma (in CL) CO_trans: "trans r"
-by (rule PO_imp_trans)
-
-lemma CompleteLatticeI:
-     "[| po \<in> PartialOrder; (\<forall>S. S \<subseteq> pset po --> (\<exists>L. isLub S po L));
-         (\<forall>S. S \<subseteq> pset po --> (\<exists>G. isGlb S po G))|]
-      ==> po \<in> CompleteLattice"
-apply (unfold CompleteLattice_def, blast)
-done
-
-lemma (in CL) CL_dualCL: "dual cl \<in> CompleteLattice"
-apply (insert cl_co)
-apply (simp add: CompleteLattice_def dual_def)
-apply (fold dual_def)
-apply (simp add: isLub_dual_isGlb [symmetric] isGlb_dual_isLub [symmetric]
-                 dualPO)
-done
-
-lemma (in PO) dualA_iff: "pset (dual cl) = pset cl"
-by (simp add: dual_def)
-
-lemma (in PO) dualr_iff: "((x, y) \<in> (order(dual cl))) = ((y, x) \<in> order cl)"
-by (simp add: dual_def)
-
-lemma (in PO) monotone_dual:
-     "monotone f (pset cl) (order cl) 
-     ==> monotone f (pset (dual cl)) (order(dual cl))"
-by (simp add: monotone_def dualA_iff dualr_iff)
-
-lemma (in PO) interval_dual:
-     "[| x \<in> A; y \<in> A|] ==> interval r x y = interval (order(dual cl)) y x"
-apply (simp add: interval_def dualr_iff)
-apply (fold r_def, fast)
-done
-
-lemma (in PO) interval_not_empty:
-     "[| trans r; interval r a b \<noteq> {} |] ==> (a, b) \<in> r"
-apply (simp add: interval_def)
-apply (unfold trans_def, blast)
-done
-
-lemma (in PO) interval_imp_mem: "x \<in> interval r a b ==> (a, x) \<in> r"
-by (simp add: interval_def)
-
-lemma (in PO) left_in_interval:
-     "[| a \<in> A; b \<in> A; interval r a b \<noteq> {} |] ==> a \<in> interval r a b"
-apply (simp (no_asm_simp) add: interval_def)
-apply (simp add: PO_imp_trans interval_not_empty)
-apply (simp add: reflE)
-done
-
-lemma (in PO) right_in_interval:
-     "[| a \<in> A; b \<in> A; interval r a b \<noteq> {} |] ==> b \<in> interval r a b"
-apply (simp (no_asm_simp) add: interval_def)
-apply (simp add: PO_imp_trans interval_not_empty)
-apply (simp add: reflE)
-done
-
-
-subsection {* sublattice *}
-
-lemma (in PO) sublattice_imp_CL:
-     "S <<= cl  ==> (| pset = S, order = induced S r |) \<in> CompleteLattice"
-by (simp add: sublattice_def CompleteLattice_def A_def r_def)
-
-lemma (in CL) sublatticeI:
-     "[| S \<subseteq> A; (| pset = S, order = induced S r |) \<in> CompleteLattice |]
-      ==> S <<= cl"
-by (simp add: sublattice_def A_def r_def)
-
-
-subsection {* lub *}
-
-lemma (in CL) lub_unique: "[| S \<subseteq> A; isLub S cl x; isLub S cl L|] ==> x = L"
-apply (rule antisymE)
-apply (auto simp add: isLub_def r_def)
-done
-
-lemma (in CL) lub_upper: "[|S \<subseteq> A; x \<in> S|] ==> (x, lub S cl) \<in> r"
-apply (rule CL_imp_ex_isLub [THEN exE], assumption)
-apply (unfold lub_def least_def)
-apply (rule some_equality [THEN ssubst])
-  apply (simp add: isLub_def)
- apply (simp add: lub_unique A_def isLub_def)
-apply (simp add: isLub_def r_def)
-done
-
-lemma (in CL) lub_least:
-     "[| S \<subseteq> A; L \<in> A; \<forall>x \<in> S. (x,L) \<in> r |] ==> (lub S cl, L) \<in> r"
-apply (rule CL_imp_ex_isLub [THEN exE], assumption)
-apply (unfold lub_def least_def)
-apply (rule_tac s=x in some_equality [THEN ssubst])
-  apply (simp add: isLub_def)
- apply (simp add: lub_unique A_def isLub_def)
-apply (simp add: isLub_def r_def A_def)
-done
-
-lemma (in CL) lub_in_lattice: "S \<subseteq> A ==> lub S cl \<in> A"
-apply (rule CL_imp_ex_isLub [THEN exE], assumption)
-apply (unfold lub_def least_def)
-apply (subst some_equality)
-apply (simp add: isLub_def)
-prefer 2 apply (simp add: isLub_def A_def)
-apply (simp add: lub_unique A_def isLub_def)
-done
-
-lemma (in CL) lubI:
-     "[| S \<subseteq> A; L \<in> A; \<forall>x \<in> S. (x,L) \<in> r;
-         \<forall>z \<in> A. (\<forall>y \<in> S. (y,z) \<in> r) --> (L,z) \<in> r |] ==> L = lub S cl"
-apply (rule lub_unique, assumption)
-apply (simp add: isLub_def A_def r_def)
-apply (unfold isLub_def)
-apply (rule conjI)
-apply (fold A_def r_def)
-apply (rule lub_in_lattice, assumption)
-apply (simp add: lub_upper lub_least)
-done
-
-lemma (in CL) lubIa: "[| S \<subseteq> A; isLub S cl L |] ==> L = lub S cl"
-by (simp add: lubI isLub_def A_def r_def)
-
-lemma (in CL) isLub_in_lattice: "isLub S cl L ==> L \<in> A"
-by (simp add: isLub_def  A_def)
-
-lemma (in CL) isLub_upper: "[|isLub S cl L; y \<in> S|] ==> (y, L) \<in> r"
-by (simp add: isLub_def r_def)
-
-lemma (in CL) isLub_least:
-     "[| isLub S cl L; z \<in> A; \<forall>y \<in> S. (y, z) \<in> r|] ==> (L, z) \<in> r"
-by (simp add: isLub_def A_def r_def)
-
-lemma (in CL) isLubI:
-     "[| L \<in> A; \<forall>y \<in> S. (y, L) \<in> r;
-         (\<forall>z \<in> A. (\<forall>y \<in> S. (y, z):r) --> (L, z) \<in> r)|] ==> isLub S cl L"
-by (simp add: isLub_def A_def r_def)
-
-
-
-subsection {* glb *}
-
-lemma (in CL) glb_in_lattice: "S \<subseteq> A ==> glb S cl \<in> A"
-apply (subst glb_dual_lub)
-apply (simp add: A_def)
-apply (rule dualA_iff [THEN subst])
-apply (rule CL.lub_in_lattice)
-apply (rule CL.intro)
-apply (rule PO.intro)
-apply (rule dualPO)
-apply (rule CL_axioms.intro)
-apply (rule CL_dualCL)
-apply (simp add: dualA_iff)
-done
-
-lemma (in CL) glb_lower: "[|S \<subseteq> A; x \<in> S|] ==> (glb S cl, x) \<in> r"
-apply (subst glb_dual_lub)
-apply (simp add: r_def)
-apply (rule dualr_iff [THEN subst])
-apply (rule CL.lub_upper)
-apply (rule CL.intro)
-apply (rule PO.intro)
-apply (rule dualPO)
-apply (rule CL_axioms.intro)
-apply (rule CL_dualCL)
-apply (simp add: dualA_iff A_def, assumption)
-done
-
-text {*
-  Reduce the sublattice property by using substructural properties;
-  abandoned see @{text "Tarski_4.ML"}.
-*}
-
-declare (in CLF) f_cl [simp]
-
-(*never proved, 2007-01-22: Tarski__CLF_unnamed_lemma
-  NOT PROVABLE because of the conjunction used in the definition: we don't
-  allow reasoning with rules like conjE, which is essential here.*)
-declare [[ atp_problem_prefix = "Tarski__CLF_unnamed_lemma" ]]
-lemma (in CLF) [simp]:
-    "f: pset cl -> pset cl & monotone f (pset cl) (order cl)" 
-apply (insert f_cl)
-apply (unfold CLF_set_def)
-apply (erule SigmaE2) 
-apply (erule CollectE) 
-apply assumption
-done
-
-lemma (in CLF) f_in_funcset: "f \<in> A -> A"
-by (simp add: A_def)
-
-lemma (in CLF) monotone_f: "monotone f A r"
-by (simp add: A_def r_def)
-
-(*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__CLF_CLF_dual" ]]
-declare (in CLF) CLF_set_def [simp] CL_dualCL [simp] monotone_dual [simp] dualA_iff [simp]
-
-lemma (in CLF) CLF_dual: "(dual cl, f) \<in> CLF_set" 
-apply (simp del: dualA_iff)
-apply (simp)
-done
-
-declare (in CLF) CLF_set_def[simp del] CL_dualCL[simp del] monotone_dual[simp del]
-          dualA_iff[simp del]
-
-
-subsection {* fixed points *}
-
-lemma fix_subset: "fix f A \<subseteq> A"
-by (simp add: fix_def, fast)
-
-lemma fix_imp_eq: "x \<in> fix f A ==> f x = x"
-by (simp add: fix_def)
-
-lemma fixf_subset:
-     "[| A \<subseteq> B; x \<in> fix (%y: A. f y) A |] ==> x \<in> fix f B"
-by (simp add: fix_def, auto)
-
-
-subsection {* lemmas for Tarski, lub *}
-
-(*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__CLF_lubH_le_flubH" ]]
-  declare CL.lub_least[intro] CLF.f_in_funcset[intro] funcset_mem[intro] CL.lub_in_lattice[intro] PO.transE[intro] PO.monotoneE[intro] CLF.monotone_f[intro] CL.lub_upper[intro] 
-lemma (in CLF) lubH_le_flubH:
-     "H = {x. (x, f x) \<in> r & x \<in> A} ==> (lub H cl, f (lub H cl)) \<in> r"
-apply (rule lub_least, fast)
-apply (rule f_in_funcset [THEN funcset_mem])
-apply (rule lub_in_lattice, fast)
--- {* @{text "\<forall>x:H. (x, f (lub H r)) \<in> r"} *}
-apply (rule ballI)
-(*never proved, 2007-01-22*)
-using [[ atp_problem_prefix = "Tarski__CLF_lubH_le_flubH_simpler" ]]
-apply (rule transE)
--- {* instantiates @{text "(x, ?z) \<in> order cl to (x, f x)"}, *}
--- {* because of the def of @{text H} *}
-apply fast
--- {* so it remains to show @{text "(f x, f (lub H cl)) \<in> r"} *}
-apply (rule_tac f = "f" in monotoneE)
-apply (rule monotone_f, fast)
-apply (rule lub_in_lattice, fast)
-apply (rule lub_upper, fast)
-apply assumption
-done
-  declare CL.lub_least[rule del] CLF.f_in_funcset[rule del] 
-          funcset_mem[rule del] CL.lub_in_lattice[rule del] 
-          PO.transE[rule del] PO.monotoneE[rule del] 
-          CLF.monotone_f[rule del] CL.lub_upper[rule del] 
-
-(*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__CLF_flubH_le_lubH" ]]
-  declare CLF.f_in_funcset[intro] funcset_mem[intro] CL.lub_in_lattice[intro]
-       PO.monotoneE[intro] CLF.monotone_f[intro] CL.lub_upper[intro] 
-       CLF.lubH_le_flubH[simp]
-lemma (in CLF) flubH_le_lubH:
-     "[|  H = {x. (x, f x) \<in> r & x \<in> A} |] ==> (f (lub H cl), lub H cl) \<in> r"
-apply (rule lub_upper, fast)
-apply (rule_tac t = "H" in ssubst, assumption)
-apply (rule CollectI)
-apply (rule conjI)
-using [[ atp_problem_prefix = "Tarski__CLF_flubH_le_lubH_simpler" ]]
-(*??no longer terminates, with combinators
-apply (metis CO_refl_on lubH_le_flubH monotone_def monotone_f reflD1 reflD2) 
-*)
-apply (metis CO_refl_on lubH_le_flubH monotoneE [OF monotone_f] refl_onD1 refl_onD2)
-apply (metis CO_refl_on lubH_le_flubH refl_onD2)
-done
-  declare CLF.f_in_funcset[rule del] funcset_mem[rule del] 
-          CL.lub_in_lattice[rule del] PO.monotoneE[rule del] 
-          CLF.monotone_f[rule del] CL.lub_upper[rule del] 
-          CLF.lubH_le_flubH[simp del]
-
-
-(*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__CLF_lubH_is_fixp" ]]
-(*Single-step version fails. The conjecture clauses refer to local abstraction
-functions (Frees), which prevents expand_defs_tac from removing those 
-"definitions" at the end of the proof. *)
-lemma (in CLF) lubH_is_fixp:
-     "H = {x. (x, f x) \<in> r & x \<in> A} ==> lub H cl \<in> fix f A"
-apply (simp add: fix_def)
-apply (rule conjI)
-proof (neg_clausify)
-assume 0: "H =
-Collect
- (COMBS (COMBB op \<and> (COMBC (COMBB op \<in> (COMBS Pair f)) r)) (COMBC op \<in> A))"
-assume 1: "lub (Collect
-      (COMBS (COMBB op \<and> (COMBC (COMBB op \<in> (COMBS Pair f)) r))
-        (COMBC op \<in> A)))
- cl
-\<notin> A"
-have 2: "lub H cl \<notin> A"
-  by (metis 1 0)
-have 3: "(lub H cl, f (lub H cl)) \<in> r"
-  by (metis lubH_le_flubH 0)
-have 4: "(f (lub H cl), lub H cl) \<in> r"
-  by (metis flubH_le_lubH 0)
-have 5: "lub H cl = f (lub H cl) \<or> (lub H cl, f (lub H cl)) \<notin> r"
-  by (metis antisymE 4)
-have 6: "lub H cl = f (lub H cl)"
-  by (metis 5 3)
-have 7: "(lub H cl, lub H cl) \<in> r"
-  by (metis 6 4)
-have 8: "\<And>X1. lub H cl \<in> X1 \<or> \<not> refl_on X1 r"
-  by (metis 7 refl_onD2)
-have 9: "\<not> refl_on A r"
-  by (metis 8 2)
-show "False"
-  by (metis CO_refl_on 9);
-next --{*apparently the way to insert a second structured proof*}
-  show "H = {x. (x, f x) \<in> r \<and> x \<in> A} \<Longrightarrow>
-  f (lub {x. (x, f x) \<in> r \<and> x \<in> A} cl) = lub {x. (x, f x) \<in> r \<and> x \<in> A} cl"
-  proof (neg_clausify)
-  assume 0: "H =
-  Collect
-   (COMBS (COMBB op \<and> (COMBC (COMBB op \<in> (COMBS Pair f)) r)) (COMBC op \<in> A))"
-  assume 1: "f (lub (Collect
-           (COMBS (COMBB op \<and> (COMBC (COMBB op \<in> (COMBS Pair f)) r))
-             (COMBC op \<in> A)))
-      cl) \<noteq>
-  lub (Collect
-        (COMBS (COMBB op \<and> (COMBC (COMBB op \<in> (COMBS Pair f)) r))
-          (COMBC op \<in> A)))
-   cl"
-  have 2: "f (lub H cl) \<noteq>
-  lub (Collect
-        (COMBS (COMBB op \<and> (COMBC (COMBB op \<in> (COMBS Pair f)) r))
-          (COMBC op \<in> A)))
-   cl"
-    by (metis 1 0)
-  have 3: "f (lub H cl) \<noteq> lub H cl"
-    by (metis 2 0)
-  have 4: "(lub H cl, f (lub H cl)) \<in> r"
-    by (metis lubH_le_flubH 0)
-  have 5: "(f (lub H cl), lub H cl) \<in> r"
-    by (metis flubH_le_lubH 0)
-  have 6: "lub H cl = f (lub H cl) \<or> (lub H cl, f (lub H cl)) \<notin> r"
-    by (metis antisymE 5)
-  have 7: "lub H cl = f (lub H cl)"
-    by (metis 6 4)
-  show "False"
-    by (metis 3 7)
-  qed
-qed
-
-lemma (in CLF) (*lubH_is_fixp:*)
-     "H = {x. (x, f x) \<in> r & x \<in> A} ==> lub H cl \<in> fix f A"
-apply (simp add: fix_def)
-apply (rule conjI)
-using [[ atp_problem_prefix = "Tarski__CLF_lubH_is_fixp_simpler" ]]
-apply (metis CO_refl_on lubH_le_flubH refl_onD1)
-apply (metis antisymE flubH_le_lubH lubH_le_flubH)
-done
-
-lemma (in CLF) fix_in_H:
-     "[| H = {x. (x, f x) \<in> r & x \<in> A};  x \<in> P |] ==> x \<in> H"
-by (simp add: P_def fix_imp_eq [of _ f A] reflE CO_refl_on
-                    fix_subset [of f A, THEN subsetD])
-
-
-lemma (in CLF) fixf_le_lubH:
-     "H = {x. (x, f x) \<in> r & x \<in> A} ==> \<forall>x \<in> fix f A. (x, lub H cl) \<in> r"
-apply (rule ballI)
-apply (rule lub_upper, fast)
-apply (rule fix_in_H)
-apply (simp_all add: P_def)
-done
-
-declare [[ atp_problem_prefix = "Tarski__CLF_lubH_least_fixf" ]]
-lemma (in CLF) lubH_least_fixf:
-     "H = {x. (x, f x) \<in> r & x \<in> A}
-      ==> \<forall>L. (\<forall>y \<in> fix f A. (y,L) \<in> r) --> (lub H cl, L) \<in> r"
-apply (metis P_def lubH_is_fixp)
-done
-
-subsection {* Tarski fixpoint theorem 1, first part *}
-declare [[ atp_problem_prefix = "Tarski__CLF_T_thm_1_lub" ]]
-  declare CL.lubI[intro] fix_subset[intro] CL.lub_in_lattice[intro] 
-          CLF.fixf_le_lubH[simp] CLF.lubH_least_fixf[simp]
-lemma (in CLF) T_thm_1_lub: "lub P cl = lub {x. (x, f x) \<in> r & x \<in> A} cl"
-(*sledgehammer;*)
-apply (rule sym)
-apply (simp add: P_def)
-apply (rule lubI)
-using [[ atp_problem_prefix = "Tarski__CLF_T_thm_1_lub_simpler" ]]
-apply (metis P_def fix_subset) 
-apply (metis Collect_conj_eq Collect_mem_eq Int_commute Int_lower1 lub_in_lattice vimage_def)
-(*??no longer terminates, with combinators
-apply (metis P_def fix_def fixf_le_lubH)
-apply (metis P_def fix_def lubH_least_fixf)
-*)
-apply (simp add: fixf_le_lubH)
-apply (simp add: lubH_least_fixf)
-done
-  declare CL.lubI[rule del] fix_subset[rule del] CL.lub_in_lattice[rule del] 
-          CLF.fixf_le_lubH[simp del] CLF.lubH_least_fixf[simp del]
-
-
-(*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__CLF_glbH_is_fixp" ]]
-  declare glb_dual_lub[simp] PO.dualA_iff[intro] CLF.lubH_is_fixp[intro] 
-          PO.dualPO[intro] CL.CL_dualCL[intro] PO.dualr_iff[simp]
-lemma (in CLF) glbH_is_fixp: "H = {x. (f x, x) \<in> r & x \<in> A} ==> glb H cl \<in> P"
-  -- {* Tarski for glb *}
-(*sledgehammer;*)
-apply (simp add: glb_dual_lub P_def A_def r_def)
-apply (rule dualA_iff [THEN subst])
-apply (rule CLF.lubH_is_fixp)
-apply (rule CLF.intro)
-apply (rule CL.intro)
-apply (rule PO.intro)
-apply (rule dualPO)
-apply (rule CL_axioms.intro)
-apply (rule CL_dualCL)
-apply (rule CLF_axioms.intro)
-apply (rule CLF_dual)
-apply (simp add: dualr_iff dualA_iff)
-done
-  declare glb_dual_lub[simp del] PO.dualA_iff[rule del] CLF.lubH_is_fixp[rule del] 
-          PO.dualPO[rule del] CL.CL_dualCL[rule del] PO.dualr_iff[simp del]
-
-
-(*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__T_thm_1_glb" ]]  (*ALL THEOREMS*)
-lemma (in CLF) T_thm_1_glb: "glb P cl = glb {x. (f x, x) \<in> r & x \<in> A} cl"
-(*sledgehammer;*)
-apply (simp add: glb_dual_lub P_def A_def r_def)
-apply (rule dualA_iff [THEN subst])
-(*never proved, 2007-01-22*)
-using [[ atp_problem_prefix = "Tarski__T_thm_1_glb_simpler" ]]  (*ALL THEOREMS*)
-(*sledgehammer;*)
-apply (simp add: CLF.T_thm_1_lub [of _ f, OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro,
-  OF dualPO CL_dualCL] dualPO CL_dualCL CLF_dual dualr_iff)
-done
-
-subsection {* interval *}
-
-
-declare [[ atp_problem_prefix = "Tarski__rel_imp_elem" ]]
-  declare (in CLF) CO_refl_on[simp] refl_on_def [simp]
-lemma (in CLF) rel_imp_elem: "(x, y) \<in> r ==> x \<in> A"
-by (metis CO_refl_on refl_onD1)
-  declare (in CLF) CO_refl_on[simp del]  refl_on_def [simp del]
-
-declare [[ atp_problem_prefix = "Tarski__interval_subset" ]]
-  declare (in CLF) rel_imp_elem[intro] 
-  declare interval_def [simp]
-lemma (in CLF) interval_subset: "[| a \<in> A; b \<in> A |] ==> interval r a b \<subseteq> A"
-by (metis CO_refl_on interval_imp_mem refl_onD refl_onD2 rel_imp_elem subset_eq)
-  declare (in CLF) rel_imp_elem[rule del] 
-  declare interval_def [simp del]
-
-
-lemma (in CLF) intervalI:
-     "[| (a, x) \<in> r; (x, b) \<in> r |] ==> x \<in> interval r a b"
-by (simp add: interval_def)
-
-lemma (in CLF) interval_lemma1:
-     "[| S \<subseteq> interval r a b; x \<in> S |] ==> (a, x) \<in> r"
-by (unfold interval_def, fast)
-
-lemma (in CLF) interval_lemma2:
-     "[| S \<subseteq> interval r a b; x \<in> S |] ==> (x, b) \<in> r"
-by (unfold interval_def, fast)
-
-lemma (in CLF) a_less_lub:
-     "[| S \<subseteq> A; S \<noteq> {};
-         \<forall>x \<in> S. (a,x) \<in> r; \<forall>y \<in> S. (y, L) \<in> r |] ==> (a,L) \<in> r"
-by (blast intro: transE)
-
-lemma (in CLF) glb_less_b:
-     "[| S \<subseteq> A; S \<noteq> {};
-         \<forall>x \<in> S. (x,b) \<in> r; \<forall>y \<in> S. (G, y) \<in> r |] ==> (G,b) \<in> r"
-by (blast intro: transE)
-
-lemma (in CLF) S_intv_cl:
-     "[| a \<in> A; b \<in> A; S \<subseteq> interval r a b |]==> S \<subseteq> A"
-by (simp add: subset_trans [OF _ interval_subset])
-
-declare [[ atp_problem_prefix = "Tarski__L_in_interval" ]]  (*ALL THEOREMS*)
-lemma (in CLF) L_in_interval:
-     "[| a \<in> A; b \<in> A; S \<subseteq> interval r a b;
-         S \<noteq> {}; isLub S cl L; interval r a b \<noteq> {} |] ==> L \<in> interval r a b" 
-(*WON'T TERMINATE
-apply (metis CO_trans intervalI interval_lemma1 interval_lemma2 isLub_least isLub_upper subset_empty subset_iff trans_def)
-*)
-apply (rule intervalI)
-apply (rule a_less_lub)
-prefer 2 apply assumption
-apply (simp add: S_intv_cl)
-apply (rule ballI)
-apply (simp add: interval_lemma1)
-apply (simp add: isLub_upper)
--- {* @{text "(L, b) \<in> r"} *}
-apply (simp add: isLub_least interval_lemma2)
-done
-
-(*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__G_in_interval" ]]  (*ALL THEOREMS*)
-lemma (in CLF) G_in_interval:
-     "[| a \<in> A; b \<in> A; interval r a b \<noteq> {}; S \<subseteq> interval r a b; isGlb S cl G;
-         S \<noteq> {} |] ==> G \<in> interval r a b"
-apply (simp add: interval_dual)
-apply (simp add: CLF.L_in_interval [of _ f, OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro]
-                 dualA_iff A_def dualPO CL_dualCL CLF_dual isGlb_dual_isLub)
-done
-
-declare [[ atp_problem_prefix = "Tarski__intervalPO" ]]  (*ALL THEOREMS*)
-lemma (in CLF) intervalPO:
-     "[| a \<in> A; b \<in> A; interval r a b \<noteq> {} |]
-      ==> (| pset = interval r a b, order = induced (interval r a b) r |)
-          \<in> PartialOrder"
-proof (neg_clausify)
-assume 0: "a \<in> A"
-assume 1: "b \<in> A"
-assume 2: "\<lparr>pset = interval r a b, order = induced (interval r a b) r\<rparr> \<notin> PartialOrder"
-have 3: "\<not> interval r a b \<subseteq> A"
-  by (metis 2 po_subset_po)
-have 4: "b \<notin> A \<or> a \<notin> A"
-  by (metis 3 interval_subset)
-have 5: "a \<notin> A"
-  by (metis 4 1)
-show "False"
-  by (metis 5 0)
-qed
-
-lemma (in CLF) intv_CL_lub:
- "[| a \<in> A; b \<in> A; interval r a b \<noteq> {} |]
-  ==> \<forall>S. S \<subseteq> interval r a b -->
-          (\<exists>L. isLub S (| pset = interval r a b,
-                          order = induced (interval r a b) r |)  L)"
-apply (intro strip)
-apply (frule S_intv_cl [THEN CL_imp_ex_isLub])
-prefer 2 apply assumption
-apply assumption
-apply (erule exE)
--- {* define the lub for the interval as *}
-apply (rule_tac x = "if S = {} then a else L" in exI)
-apply (simp (no_asm_simp) add: isLub_def split del: split_if)
-apply (intro impI conjI)
--- {* @{text "(if S = {} then a else L) \<in> interval r a b"} *}
-apply (simp add: CL_imp_PO L_in_interval)
-apply (simp add: left_in_interval)
--- {* lub prop 1 *}
-apply (case_tac "S = {}")
--- {* @{text "S = {}, y \<in> S = False => everything"} *}
-apply fast
--- {* @{text "S \<noteq> {}"} *}
-apply simp
--- {* @{text "\<forall>y:S. (y, L) \<in> induced (interval r a b) r"} *}
-apply (rule ballI)
-apply (simp add: induced_def  L_in_interval)
-apply (rule conjI)
-apply (rule subsetD)
-apply (simp add: S_intv_cl, assumption)
-apply (simp add: isLub_upper)
--- {* @{text "\<forall>z:interval r a b. (\<forall>y:S. (y, z) \<in> induced (interval r a b) r \<longrightarrow> (if S = {} then a else L, z) \<in> induced (interval r a b) r"} *}
-apply (rule ballI)
-apply (rule impI)
-apply (case_tac "S = {}")
--- {* @{text "S = {}"} *}
-apply simp
-apply (simp add: induced_def  interval_def)
-apply (rule conjI)
-apply (rule reflE, assumption)
-apply (rule interval_not_empty)
-apply (rule CO_trans)
-apply (simp add: interval_def)
--- {* @{text "S \<noteq> {}"} *}
-apply simp
-apply (simp add: induced_def  L_in_interval)
-apply (rule isLub_least, assumption)
-apply (rule subsetD)
-prefer 2 apply assumption
-apply (simp add: S_intv_cl, fast)
-done
-
-lemmas (in CLF) intv_CL_glb = intv_CL_lub [THEN Rdual]
-
-(*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__interval_is_sublattice" ]]  (*ALL THEOREMS*)
-lemma (in CLF) interval_is_sublattice:
-     "[| a \<in> A; b \<in> A; interval r a b \<noteq> {} |]
-        ==> interval r a b <<= cl"
-(*sledgehammer *)
-apply (rule sublatticeI)
-apply (simp add: interval_subset)
-(*never proved, 2007-01-22*)
-using [[ atp_problem_prefix = "Tarski__interval_is_sublattice_simpler" ]]
-(*sledgehammer *)
-apply (rule CompleteLatticeI)
-apply (simp add: intervalPO)
- apply (simp add: intv_CL_lub)
-apply (simp add: intv_CL_glb)
-done
-
-lemmas (in CLF) interv_is_compl_latt =
-    interval_is_sublattice [THEN sublattice_imp_CL]
-
-
-subsection {* Top and Bottom *}
-lemma (in CLF) Top_dual_Bot: "Top cl = Bot (dual cl)"
-by (simp add: Top_def Bot_def least_def greatest_def dualA_iff dualr_iff)
-
-lemma (in CLF) Bot_dual_Top: "Bot cl = Top (dual cl)"
-by (simp add: Top_def Bot_def least_def greatest_def dualA_iff dualr_iff)
-
-declare [[ atp_problem_prefix = "Tarski__Bot_in_lattice" ]]  (*ALL THEOREMS*)
-lemma (in CLF) Bot_in_lattice: "Bot cl \<in> A"
-(*sledgehammer; *)
-apply (simp add: Bot_def least_def)
-apply (rule_tac a="glb A cl" in someI2)
-apply (simp_all add: glb_in_lattice glb_lower 
-                     r_def [symmetric] A_def [symmetric])
-done
-
-(*first proved 2007-01-25 after relaxing relevance*)
-declare [[ atp_problem_prefix = "Tarski__Top_in_lattice" ]]  (*ALL THEOREMS*)
-lemma (in CLF) Top_in_lattice: "Top cl \<in> A"
-(*sledgehammer;*)
-apply (simp add: Top_dual_Bot A_def)
-(*first proved 2007-01-25 after relaxing relevance*)
-using [[ atp_problem_prefix = "Tarski__Top_in_lattice_simpler" ]]  (*ALL THEOREMS*)
-(*sledgehammer*)
-apply (rule dualA_iff [THEN subst])
-apply (blast intro!: CLF.Bot_in_lattice [OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro] dualPO CL_dualCL CLF_dual)
-done
-
-lemma (in CLF) Top_prop: "x \<in> A ==> (x, Top cl) \<in> r"
-apply (simp add: Top_def greatest_def)
-apply (rule_tac a="lub A cl" in someI2)
-apply (rule someI2)
-apply (simp_all add: lub_in_lattice lub_upper 
-                     r_def [symmetric] A_def [symmetric])
-done
-
-(*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__Bot_prop" ]]  (*ALL THEOREMS*) 
-lemma (in CLF) Bot_prop: "x \<in> A ==> (Bot cl, x) \<in> r"
-(*sledgehammer*) 
-apply (simp add: Bot_dual_Top r_def)
-apply (rule dualr_iff [THEN subst])
-apply (simp add: CLF.Top_prop [of _ f, OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro]
-                 dualA_iff A_def dualPO CL_dualCL CLF_dual)
-done
-
-declare [[ atp_problem_prefix = "Tarski__Bot_in_lattice" ]]  (*ALL THEOREMS*)
-lemma (in CLF) Top_intv_not_empty: "x \<in> A  ==> interval r x (Top cl) \<noteq> {}" 
-apply (metis Top_in_lattice Top_prop empty_iff intervalI reflE)
-done
-
-declare [[ atp_problem_prefix = "Tarski__Bot_intv_not_empty" ]]  (*ALL THEOREMS*)
-lemma (in CLF) Bot_intv_not_empty: "x \<in> A ==> interval r (Bot cl) x \<noteq> {}" 
-apply (metis Bot_prop ex_in_conv intervalI reflE rel_imp_elem)
-done
-
-
-subsection {* fixed points form a partial order *}
-
-lemma (in CLF) fixf_po: "(| pset = P, order = induced P r|) \<in> PartialOrder"
-by (simp add: P_def fix_subset po_subset_po)
-
-(*first proved 2007-01-25 after relaxing relevance*)
-declare [[ atp_problem_prefix = "Tarski__Y_subset_A" ]]
-  declare (in Tarski) P_def[simp] Y_ss [simp]
-  declare fix_subset [intro] subset_trans [intro]
-lemma (in Tarski) Y_subset_A: "Y \<subseteq> A"
-(*sledgehammer*) 
-apply (rule subset_trans [OF _ fix_subset])
-apply (rule Y_ss [simplified P_def])
-done
-  declare (in Tarski) P_def[simp del] Y_ss [simp del]
-  declare fix_subset [rule del] subset_trans [rule del]
-
-
-lemma (in Tarski) lubY_in_A: "lub Y cl \<in> A"
-  by (rule Y_subset_A [THEN lub_in_lattice])
-
-(*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__lubY_le_flubY" ]]  (*ALL THEOREMS*)
-lemma (in Tarski) lubY_le_flubY: "(lub Y cl, f (lub Y cl)) \<in> r"
-(*sledgehammer*) 
-apply (rule lub_least)
-apply (rule Y_subset_A)
-apply (rule f_in_funcset [THEN funcset_mem])
-apply (rule lubY_in_A)
--- {* @{text "Y \<subseteq> P ==> f x = x"} *}
-apply (rule ballI)
-using [[ atp_problem_prefix = "Tarski__lubY_le_flubY_simpler" ]]  (*ALL THEOREMS*)
-(*sledgehammer *)
-apply (rule_tac t = "x" in fix_imp_eq [THEN subst])
-apply (erule Y_ss [simplified P_def, THEN subsetD])
--- {* @{text "reduce (f x, f (lub Y cl)) \<in> r to (x, lub Y cl) \<in> r"} by monotonicity *}
-using [[ atp_problem_prefix = "Tarski__lubY_le_flubY_simplest" ]]  (*ALL THEOREMS*)
-(*sledgehammer*)
-apply (rule_tac f = "f" in monotoneE)
-apply (rule monotone_f)
-apply (simp add: Y_subset_A [THEN subsetD])
-apply (rule lubY_in_A)
-apply (simp add: lub_upper Y_subset_A)
-done
-
-(*first proved 2007-01-25 after relaxing relevance*)
-declare [[ atp_problem_prefix = "Tarski__intY1_subset" ]]  (*ALL THEOREMS*)
-lemma (in Tarski) intY1_subset: "intY1 \<subseteq> A"
-(*sledgehammer*) 
-apply (unfold intY1_def)
-apply (rule interval_subset)
-apply (rule lubY_in_A)
-apply (rule Top_in_lattice)
-done
-
-lemmas (in Tarski) intY1_elem = intY1_subset [THEN subsetD]
-
-(*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__intY1_f_closed" ]]  (*ALL THEOREMS*)
-lemma (in Tarski) intY1_f_closed: "x \<in> intY1 \<Longrightarrow> f x \<in> intY1"
-(*sledgehammer*) 
-apply (simp add: intY1_def  interval_def)
-apply (rule conjI)
-apply (rule transE)
-apply (rule lubY_le_flubY)
--- {* @{text "(f (lub Y cl), f x) \<in> r"} *}
-using [[ atp_problem_prefix = "Tarski__intY1_f_closed_simpler" ]]  (*ALL THEOREMS*)
-(*sledgehammer [has been proved before now...]*)
-apply (rule_tac f=f in monotoneE)
-apply (rule monotone_f)
-apply (rule lubY_in_A)
-apply (simp add: intY1_def interval_def  intY1_elem)
-apply (simp add: intY1_def  interval_def)
--- {* @{text "(f x, Top cl) \<in> r"} *} 
-apply (rule Top_prop)
-apply (rule f_in_funcset [THEN funcset_mem])
-apply (simp add: intY1_def interval_def  intY1_elem)
-done
-
-declare [[ atp_problem_prefix = "Tarski__intY1_func" ]]  (*ALL THEOREMS*)
-lemma (in Tarski) intY1_func: "(%x: intY1. f x) \<in> intY1 -> intY1"
-apply (rule restrict_in_funcset)
-apply (metis intY1_f_closed restrict_in_funcset)
-done
-
-declare [[ atp_problem_prefix = "Tarski__intY1_mono" ]]  (*ALL THEOREMS*)
-lemma (in Tarski) intY1_mono:
-     "monotone (%x: intY1. f x) intY1 (induced intY1 r)"
-(*sledgehammer *)
-apply (auto simp add: monotone_def induced_def intY1_f_closed)
-apply (blast intro: intY1_elem monotone_f [THEN monotoneE])
-done
-
-(*proof requires relaxing relevance: 2007-01-25*)
-declare [[ atp_problem_prefix = "Tarski__intY1_is_cl" ]]  (*ALL THEOREMS*)
-lemma (in Tarski) intY1_is_cl:
-    "(| pset = intY1, order = induced intY1 r |) \<in> CompleteLattice"
-(*sledgehammer*) 
-apply (unfold intY1_def)
-apply (rule interv_is_compl_latt)
-apply (rule lubY_in_A)
-apply (rule Top_in_lattice)
-apply (rule Top_intv_not_empty)
-apply (rule lubY_in_A)
-done
-
-(*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__v_in_P" ]]  (*ALL THEOREMS*)
-lemma (in Tarski) v_in_P: "v \<in> P"
-(*sledgehammer*) 
-apply (unfold P_def)
-apply (rule_tac A = "intY1" in fixf_subset)
-apply (rule intY1_subset)
-apply (simp add: CLF.glbH_is_fixp [OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro, OF _ intY1_is_cl, simplified]
-                 v_def CL_imp_PO intY1_is_cl CLF_set_def intY1_func intY1_mono)
-done
-
-declare [[ atp_problem_prefix = "Tarski__z_in_interval" ]]  (*ALL THEOREMS*)
-lemma (in Tarski) z_in_interval:
-     "[| z \<in> P; \<forall>y\<in>Y. (y, z) \<in> induced P r |] ==> z \<in> intY1"
-(*sledgehammer *)
-apply (unfold intY1_def P_def)
-apply (rule intervalI)
-prefer 2
- apply (erule fix_subset [THEN subsetD, THEN Top_prop])
-apply (rule lub_least)
-apply (rule Y_subset_A)
-apply (fast elim!: fix_subset [THEN subsetD])
-apply (simp add: induced_def)
-done
-
-declare [[ atp_problem_prefix = "Tarski__fz_in_int_rel" ]]  (*ALL THEOREMS*)
-lemma (in Tarski) f'z_in_int_rel: "[| z \<in> P; \<forall>y\<in>Y. (y, z) \<in> induced P r |]
-      ==> ((%x: intY1. f x) z, z) \<in> induced intY1 r" 
-apply (metis P_def acc_def fix_imp_eq fix_subset indI reflE restrict_apply subset_eq z_in_interval)
-done
-
-(*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__tarski_full_lemma" ]]  (*ALL THEOREMS*)
-lemma (in Tarski) tarski_full_lemma:
-     "\<exists>L. isLub Y (| pset = P, order = induced P r |) L"
-apply (rule_tac x = "v" in exI)
-apply (simp add: isLub_def)
--- {* @{text "v \<in> P"} *}
-apply (simp add: v_in_P)
-apply (rule conjI)
-(*sledgehammer*) 
--- {* @{text v} is lub *}
--- {* @{text "1. \<forall>y:Y. (y, v) \<in> induced P r"} *}
-apply (rule ballI)
-apply (simp add: induced_def subsetD v_in_P)
-apply (rule conjI)
-apply (erule Y_ss [THEN subsetD])
-apply (rule_tac b = "lub Y cl" in transE)
-apply (rule lub_upper)
-apply (rule Y_subset_A, assumption)
-apply (rule_tac b = "Top cl" in interval_imp_mem)
-apply (simp add: v_def)
-apply (fold intY1_def)
-apply (rule CL.glb_in_lattice [OF CL.intro, OF PO.intro CL_axioms.intro, OF _ intY1_is_cl, simplified])
- apply (simp add: CL_imp_PO intY1_is_cl, force)
--- {* @{text v} is LEAST ub *}
-apply clarify
-apply (rule indI)
-  prefer 3 apply assumption
- prefer 2 apply (simp add: v_in_P)
-apply (unfold v_def)
-(*never proved, 2007-01-22*)
-using [[ atp_problem_prefix = "Tarski__tarski_full_lemma_simpler" ]]
-(*sledgehammer*) 
-apply (rule indE)
-apply (rule_tac [2] intY1_subset)
-(*never proved, 2007-01-22*)
-using [[ atp_problem_prefix = "Tarski__tarski_full_lemma_simplest" ]]
-(*sledgehammer*) 
-apply (rule CL.glb_lower [OF CL.intro, OF PO.intro CL_axioms.intro, OF _ intY1_is_cl, simplified])
-  apply (simp add: CL_imp_PO intY1_is_cl)
- apply force
-apply (simp add: induced_def intY1_f_closed z_in_interval)
-apply (simp add: P_def fix_imp_eq [of _ f A] reflE
-                 fix_subset [of f A, THEN subsetD])
-done
-
-lemma CompleteLatticeI_simp:
-     "[| (| pset = A, order = r |) \<in> PartialOrder;
-         \<forall>S. S \<subseteq> A --> (\<exists>L. isLub S (| pset = A, order = r |)  L) |]
-    ==> (| pset = A, order = r |) \<in> CompleteLattice"
-by (simp add: CompleteLatticeI Rdual)
-
-
-(*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__Tarski_full" ]]
-  declare (in CLF) fixf_po[intro] P_def [simp] A_def [simp] r_def [simp]
-               Tarski.tarski_full_lemma [intro] cl_po [intro] cl_co [intro]
-               CompleteLatticeI_simp [intro]
-theorem (in CLF) Tarski_full:
-     "(| pset = P, order = induced P r|) \<in> CompleteLattice"
-(*sledgehammer*) 
-apply (rule CompleteLatticeI_simp)
-apply (rule fixf_po, clarify)
-(*never proved, 2007-01-22*)
-using [[ atp_problem_prefix = "Tarski__Tarski_full_simpler" ]]
-(*sledgehammer*) 
-apply (simp add: P_def A_def r_def)
-apply (blast intro!: Tarski.tarski_full_lemma [OF Tarski.intro, OF CLF.intro Tarski_axioms.intro,
-  OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro] cl_po cl_co f_cl)
-done
-  declare (in CLF) fixf_po[rule del] P_def [simp del] A_def [simp del] r_def [simp del]
-         Tarski.tarski_full_lemma [rule del] cl_po [rule del] cl_co [rule del]
-         CompleteLatticeI_simp [rule del]
-
-
-end
--- a/src/HOL/MetisExamples/TransClosure.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,62 +0,0 @@
-(*  Title:      HOL/MetisTest/TransClosure.thy
-    ID:         $Id$
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-
-Testing the metis method
-*)
-
-theory TransClosure
-imports Main
-begin
-
-types addr = nat
-
-datatype val
-  = Unit        -- "dummy result value of void expressions"
-  | Null        -- "null reference"
-  | Bool bool   -- "Boolean value"
-  | Intg int    -- "integer value" 
-  | Addr addr   -- "addresses of objects in the heap"
-
-consts R::"(addr \<times> addr) set"
-
-consts f:: "addr \<Rightarrow> val"
-
-declare [[ atp_problem_prefix = "TransClosure__test" ]]
-lemma "\<lbrakk> f c = Intg x; \<forall> y. f b = Intg y \<longrightarrow> y \<noteq> x; (a,b) \<in> R\<^sup>*; (b,c) \<in> R\<^sup>* \<rbrakk> 
-   \<Longrightarrow> \<exists> c. (b,c) \<in> R \<and> (a,c) \<in> R\<^sup>*"  
-by (metis Transitive_Closure.rtrancl_into_rtrancl converse_rtranclE trancl_reflcl)
-
-lemma "\<lbrakk> f c = Intg x; \<forall> y. f b = Intg y \<longrightarrow> y \<noteq> x; (a,b) \<in> R\<^sup>*; (b,c) \<in> R\<^sup>* \<rbrakk> 
-   \<Longrightarrow> \<exists> c. (b,c) \<in> R \<and> (a,c) \<in> R\<^sup>*"
-proof (neg_clausify)
-assume 0: "f c = Intg x"
-assume 1: "(a, b) \<in> R\<^sup>*"
-assume 2: "(b, c) \<in> R\<^sup>*"
-assume 3: "f b \<noteq> Intg x"
-assume 4: "\<And>A. (b, A) \<notin> R \<or> (a, A) \<notin> R\<^sup>*"
-have 5: "b = c \<or> b \<in> Domain R"
-  by (metis Not_Domain_rtrancl 2)
-have 6: "\<And>X1. (a, X1) \<in> R\<^sup>* \<or> (b, X1) \<notin> R"
-  by (metis Transitive_Closure.rtrancl_into_rtrancl 1)
-have 7: "\<And>X1. (b, X1) \<notin> R"
-  by (metis 6 4)
-have 8: "b \<notin> Domain R"
-  by (metis 7 DomainE)
-have 9: "c = b"
-  by (metis 5 8)
-have 10: "f b = Intg x"
-  by (metis 0 9)
-show "False"
-  by (metis 10 3)
-qed
-
-declare [[ atp_problem_prefix = "TransClosure__test_simpler" ]]
-lemma "\<lbrakk> f c = Intg x; \<forall> y. f b = Intg y \<longrightarrow> y \<noteq> x; (a,b) \<in> R\<^sup>*; (b,c) \<in> R\<^sup>* \<rbrakk> 
-   \<Longrightarrow> \<exists> c. (b,c) \<in> R \<and> (a,c) \<in> R\<^sup>*"
-apply (erule_tac x="b" in converse_rtranclE)
-apply (metis rel_pow_0_E rel_pow_0_I)
-apply (metis DomainE Domain_iff Transitive_Closure.rtrancl_into_rtrancl)
-done
-
-end
--- a/src/HOL/MetisExamples/set.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,283 +0,0 @@
-(*  Title:      HOL/MetisExamples/set.thy
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-
-Testing the metis method
-*)
-
-theory set imports Main
-
-begin
-
-lemma "EX x X. ALL y. EX z Z. (~P(y,y) | P(x,x) | ~S(z,x)) &
-               (S(x,y) | ~S(y,z) | Q(Z,Z))  &
-               (Q(X,y) | ~Q(y,Z) | S(X,X))" 
-by metis
-(*??But metis can't prove the single-step version...*)
-
-
-
-lemma "P(n::nat) ==> ~P(0) ==> n ~= 0"
-by metis
-
-declare [[sledgehammer_modulus = 1]]
-
-
-(*multiple versions of this example*)
-lemma (*equal_union: *)
-   "(X = Y \<union> Z) =
-    (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
-proof (neg_clausify)
-fix x
-assume 0: "Y \<subseteq> X \<or> X = Y \<union> Z"
-assume 1: "Z \<subseteq> X \<or> X = Y \<union> Z"
-assume 2: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Y \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
-assume 3: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Z \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
-assume 4: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> \<not> X \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
-assume 5: "\<And>V. ((\<not> Y \<subseteq> V \<or> \<not> Z \<subseteq> V) \<or> X \<subseteq> V) \<or> X = Y \<union> Z"
-have 6: "sup Y Z = X \<or> Y \<subseteq> X"
-  by (metis 0)
-have 7: "sup Y Z = X \<or> Z \<subseteq> X"
-  by (metis 1)
-have 8: "\<And>X3. sup Y Z = X \<or> X \<subseteq> X3 \<or> \<not> Y \<subseteq> X3 \<or> \<not> Z \<subseteq> X3"
-  by (metis 5)
-have 9: "Y \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X"
-  by (metis 2)
-have 10: "Z \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X"
-  by (metis 3)
-have 11: "sup Y Z \<noteq> X \<or> \<not> X \<subseteq> x \<or> \<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X"
-  by (metis 4)
-have 12: "Z \<subseteq> X"
-  by (metis Un_upper2 7)
-have 13: "\<And>X3. sup Y Z = X \<or> X \<subseteq> sup X3 Z \<or> \<not> Y \<subseteq> sup X3 Z"
-  by (metis 8 Un_upper2)
-have 14: "Y \<subseteq> X"
-  by (metis Un_upper1 6)
-have 15: "Z \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X"
-  by (metis 10 12)
-have 16: "Y \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X"
-  by (metis 9 12)
-have 17: "sup Y Z \<noteq> X \<or> \<not> X \<subseteq> x \<or> \<not> Y \<subseteq> X"
-  by (metis 11 12)
-have 18: "sup Y Z \<noteq> X \<or> \<not> X \<subseteq> x"
-  by (metis 17 14)
-have 19: "Z \<subseteq> x \<or> sup Y Z \<noteq> X"
-  by (metis 15 14)
-have 20: "Y \<subseteq> x \<or> sup Y Z \<noteq> X"
-  by (metis 16 14)
-have 21: "sup Y Z = X \<or> X \<subseteq> sup Y Z"
-  by (metis 13 Un_upper1)
-have 22: "sup Y Z = X \<or> \<not> sup Y Z \<subseteq> X"
-  by (metis equalityI 21)
-have 23: "sup Y Z = X \<or> \<not> Z \<subseteq> X \<or> \<not> Y \<subseteq> X"
-  by (metis 22 Un_least)
-have 24: "sup Y Z = X \<or> \<not> Y \<subseteq> X"
-  by (metis 23 12)
-have 25: "sup Y Z = X"
-  by (metis 24 14)
-have 26: "\<And>X3. X \<subseteq> X3 \<or> \<not> Z \<subseteq> X3 \<or> \<not> Y \<subseteq> X3"
-  by (metis Un_least 25)
-have 27: "Y \<subseteq> x"
-  by (metis 20 25)
-have 28: "Z \<subseteq> x"
-  by (metis 19 25)
-have 29: "\<not> X \<subseteq> x"
-  by (metis 18 25)
-have 30: "X \<subseteq> x \<or> \<not> Y \<subseteq> x"
-  by (metis 26 28)
-have 31: "X \<subseteq> x"
-  by (metis 30 27)
-show "False"
-  by (metis 31 29)
-qed
-
-declare [[sledgehammer_modulus = 2]]
-
-lemma (*equal_union: *)
-   "(X = Y \<union> Z) =
-    (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
-proof (neg_clausify)
-fix x
-assume 0: "Y \<subseteq> X \<or> X = Y \<union> Z"
-assume 1: "Z \<subseteq> X \<or> X = Y \<union> Z"
-assume 2: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Y \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
-assume 3: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Z \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
-assume 4: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> \<not> X \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
-assume 5: "\<And>V. ((\<not> Y \<subseteq> V \<or> \<not> Z \<subseteq> V) \<or> X \<subseteq> V) \<or> X = Y \<union> Z"
-have 6: "sup Y Z = X \<or> Y \<subseteq> X"
-  by (metis 0)
-have 7: "Y \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X"
-  by (metis 2)
-have 8: "sup Y Z \<noteq> X \<or> \<not> X \<subseteq> x \<or> \<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X"
-  by (metis 4)
-have 9: "\<And>X3. sup Y Z = X \<or> X \<subseteq> sup X3 Z \<or> \<not> Y \<subseteq> sup X3 Z"
-  by (metis 5 Un_upper2)
-have 10: "Z \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X"
-  by (metis 3 Un_upper2)
-have 11: "sup Y Z \<noteq> X \<or> \<not> X \<subseteq> x \<or> \<not> Y \<subseteq> X"
-  by (metis 8 Un_upper2)
-have 12: "Z \<subseteq> x \<or> sup Y Z \<noteq> X"
-  by (metis 10 Un_upper1)
-have 13: "sup Y Z = X \<or> X \<subseteq> sup Y Z"
-  by (metis 9 Un_upper1)
-have 14: "sup Y Z = X \<or> \<not> Z \<subseteq> X \<or> \<not> Y \<subseteq> X"
-  by (metis equalityI 13 Un_least)
-have 15: "sup Y Z = X"
-  by (metis 14 1 6)
-have 16: "Y \<subseteq> x"
-  by (metis 7 Un_upper2 Un_upper1 15)
-have 17: "\<not> X \<subseteq> x"
-  by (metis 11 Un_upper1 15)
-have 18: "X \<subseteq> x"
-  by (metis Un_least 15 12 15 16)
-show "False"
-  by (metis 18 17)
-qed
-
-declare [[sledgehammer_modulus = 3]]
-
-lemma (*equal_union: *)
-   "(X = Y \<union> Z) =
-    (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
-proof (neg_clausify)
-fix x
-assume 0: "Y \<subseteq> X \<or> X = Y \<union> Z"
-assume 1: "Z \<subseteq> X \<or> X = Y \<union> Z"
-assume 2: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Y \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
-assume 3: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Z \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
-assume 4: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> \<not> X \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
-assume 5: "\<And>V. ((\<not> Y \<subseteq> V \<or> \<not> Z \<subseteq> V) \<or> X \<subseteq> V) \<or> X = Y \<union> Z"
-have 6: "Z \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X"
-  by (metis 3)
-have 7: "\<And>X3. sup Y Z = X \<or> X \<subseteq> sup X3 Z \<or> \<not> Y \<subseteq> sup X3 Z"
-  by (metis 5 Un_upper2)
-have 8: "Y \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X"
-  by (metis 2 Un_upper2)
-have 9: "Z \<subseteq> x \<or> sup Y Z \<noteq> X"
-  by (metis 6 Un_upper2 Un_upper1)
-have 10: "sup Y Z = X \<or> \<not> sup Y Z \<subseteq> X"
-  by (metis equalityI 7 Un_upper1)
-have 11: "sup Y Z = X"
-  by (metis 10 Un_least 1 0)
-have 12: "Z \<subseteq> x"
-  by (metis 9 11)
-have 13: "X \<subseteq> x"
-  by (metis Un_least 11 12 8 Un_upper1 11)
-show "False"
-  by (metis 13 4 Un_upper2 Un_upper1 11)
-qed
-
-(*Example included in TPHOLs paper*)
-
-declare [[sledgehammer_modulus = 4]]
-
-lemma (*equal_union: *)
-   "(X = Y \<union> Z) =
-    (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
-proof (neg_clausify)
-fix x
-assume 0: "Y \<subseteq> X \<or> X = Y \<union> Z"
-assume 1: "Z \<subseteq> X \<or> X = Y \<union> Z"
-assume 2: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Y \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
-assume 3: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Z \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
-assume 4: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> \<not> X \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
-assume 5: "\<And>V. ((\<not> Y \<subseteq> V \<or> \<not> Z \<subseteq> V) \<or> X \<subseteq> V) \<or> X = Y \<union> Z"
-have 6: "sup Y Z \<noteq> X \<or> \<not> X \<subseteq> x \<or> \<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X"
-  by (metis 4)
-have 7: "Z \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X"
-  by (metis 3 Un_upper2)
-have 8: "Z \<subseteq> x \<or> sup Y Z \<noteq> X"
-  by (metis 7 Un_upper1)
-have 9: "sup Y Z = X \<or> \<not> Z \<subseteq> X \<or> \<not> Y \<subseteq> X"
-  by (metis equalityI 5 Un_upper2 Un_upper1 Un_least)
-have 10: "Y \<subseteq> x"
-  by (metis 2 Un_upper2 1 Un_upper1 0 9 Un_upper2 1 Un_upper1 0)
-have 11: "X \<subseteq> x"
-  by (metis Un_least 9 Un_upper2 1 Un_upper1 0 8 9 Un_upper2 1 Un_upper1 0 10)
-show "False"
-  by (metis 11 6 Un_upper2 1 Un_upper1 0 9 Un_upper2 1 Un_upper1 0)
-qed 
-
-declare [[ atp_problem_prefix = "set__equal_union" ]]
-lemma (*equal_union: *)
-   "(X = Y \<union> Z) =
-    (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" 
-(*One shot proof: hand-reduced. Metis can't do the full proof any more.*)
-by (metis Un_least Un_upper1 Un_upper2 set_eq_subset)
-
-
-declare [[ atp_problem_prefix = "set__equal_inter" ]]
-lemma "(X = Y \<inter> Z) =
-    (X \<subseteq> Y \<and> X \<subseteq> Z \<and> (\<forall>V. V \<subseteq> Y \<and> V \<subseteq> Z \<longrightarrow> V \<subseteq> X))"
-by (metis Int_greatest Int_lower1 Int_lower2 set_eq_subset)
-
-declare [[ atp_problem_prefix = "set__fixedpoint" ]]
-lemma fixedpoint:
-    "\<exists>!x. f (g x) = x \<Longrightarrow> \<exists>!y. g (f y) = y"
-by metis
-
-lemma (*fixedpoint:*)
-    "\<exists>!x. f (g x) = x \<Longrightarrow> \<exists>!y. g (f y) = y"
-proof (neg_clausify)
-fix x xa
-assume 0: "f (g x) = x"
-assume 1: "\<And>y. y = x \<or> f (g y) \<noteq> y"
-assume 2: "\<And>x. g (f (xa x)) = xa x \<or> g (f x) \<noteq> x"
-assume 3: "\<And>x. g (f x) \<noteq> x \<or> xa x \<noteq> x"
-have 4: "\<And>X1. g (f X1) \<noteq> X1 \<or> g x \<noteq> X1"
-  by (metis 3 1 2)
-show "False"
-  by (metis 4 0)
-qed
-
-declare [[ atp_problem_prefix = "set__singleton_example" ]]
-lemma (*singleton_example_2:*)
-     "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
-by (metis Set.subsetI Union_upper insertCI set_eq_subset)
-  --{*found by SPASS*}
-
-lemma (*singleton_example_2:*)
-     "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
-by (metis Set.subsetI Union_upper insert_iff set_eq_subset)
-
-lemma singleton_example_2:
-     "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
-proof (neg_clausify)
-assume 0: "\<And>x. \<not> S \<subseteq> {x}"
-assume 1: "\<And>x. x \<notin> S \<or> \<Union>S \<subseteq> x"
-have 2: "\<And>X3. X3 = \<Union>S \<or> \<not> X3 \<subseteq> \<Union>S \<or> X3 \<notin> S"
-  by (metis set_eq_subset 1)
-have 3: "\<And>X3. S \<subseteq> insert (\<Union>S) X3"
-  by (metis insert_iff Set.subsetI Union_upper 2 Set.subsetI)
-show "False"
-  by (metis 3 0)
-qed
-
-
-
-text {*
-  From W. W. Bledsoe and Guohui Feng, SET-VAR. JAR 11 (3), 1993, pages
-  293-314.
-*}
-
-declare [[ atp_problem_prefix = "set__Bledsoe_Fung" ]]
-(*Notes: 1, the numbering doesn't completely agree with the paper. 
-2, we must rename set variables to avoid type clashes.*)
-lemma "\<exists>B. (\<forall>x \<in> B. x \<le> (0::int))"
-      "D \<in> F \<Longrightarrow> \<exists>G. \<forall>A \<in> G. \<exists>B \<in> F. A \<subseteq> B"
-      "P a \<Longrightarrow> \<exists>A. (\<forall>x \<in> A. P x) \<and> (\<exists>y. y \<in> A)"
-      "a < b \<and> b < (c::int) \<Longrightarrow> \<exists>B. a \<notin> B \<and> b \<in> B \<and> c \<notin> B"
-      "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A"
-      "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A"
-      "\<exists>A. a \<notin> A"
-      "(\<forall>C. (0, 0) \<in> C \<and> (\<forall>x y. (x, y) \<in> C \<longrightarrow> (Suc x, Suc y) \<in> C) \<longrightarrow> (n, m) \<in> C) \<and> Q n \<longrightarrow> Q m" 
-apply (metis atMost_iff)
-apply (metis emptyE)
-apply (metis insert_iff singletonE)
-apply (metis insertCI singletonE zless_le)
-apply (metis Collect_def Collect_mem_eq)
-apply (metis Collect_def Collect_mem_eq)
-apply (metis DiffE)
-apply (metis pair_in_Id_conv) 
-done
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Metis_Examples/Abstraction.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,297 @@
+(*  Title:      HOL/Metis_Examples/Abstraction.thy
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+
+Testing the metis method.
+*)
+
+theory Abstraction
+imports Main FuncSet
+begin
+
+(*For Christoph Benzmueller*)
+lemma "x<1 & ((op=) = (op=)) ==> ((op=) = (op=)) & (x<(2::nat))";
+  by (metis One_nat_def less_Suc0 not_less0 not_less_eq numeral_2_eq_2)
+
+(*this is a theorem, but we can't prove it unless ext is applied explicitly
+lemma "(op=) = (%x y. y=x)"
+*)
+
+consts
+  monotone :: "['a => 'a, 'a set, ('a *'a)set] => bool"
+  pset  :: "'a set => 'a set"
+  order :: "'a set => ('a * 'a) set"
+
+declare [[ atp_problem_prefix = "Abstraction__Collect_triv" ]]
+lemma (*Collect_triv:*) "a \<in> {x. P x} ==> P a"
+proof (neg_clausify)
+assume 0: "(a\<Colon>'a\<Colon>type) \<in> Collect (P\<Colon>'a\<Colon>type \<Rightarrow> bool)"
+assume 1: "\<not> (P\<Colon>'a\<Colon>type \<Rightarrow> bool) (a\<Colon>'a\<Colon>type)"
+have 2: "(P\<Colon>'a\<Colon>type \<Rightarrow> bool) (a\<Colon>'a\<Colon>type)"
+  by (metis CollectD 0)
+show "False"
+  by (metis 2 1)
+qed
+
+lemma Collect_triv: "a \<in> {x. P x} ==> P a"
+by (metis mem_Collect_eq)
+
+
+declare [[ atp_problem_prefix = "Abstraction__Collect_mp" ]]
+lemma "a \<in> {x. P x --> Q x} ==> a \<in> {x. P x} ==> a \<in> {x. Q x}"
+  by (metis CollectI Collect_imp_eq ComplD UnE mem_Collect_eq);
+  --{*34 secs*}
+
+declare [[ atp_problem_prefix = "Abstraction__Sigma_triv" ]]
+lemma "(a,b) \<in> Sigma A B ==> a \<in> A & b \<in> B a"
+proof (neg_clausify)
+assume 0: "(a\<Colon>'a\<Colon>type, b\<Colon>'b\<Colon>type) \<in> Sigma (A\<Colon>'a\<Colon>type set) (B\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>type set)"
+assume 1: "(a\<Colon>'a\<Colon>type) \<notin> (A\<Colon>'a\<Colon>type set) \<or> (b\<Colon>'b\<Colon>type) \<notin> (B\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>type set) a"
+have 2: "(a\<Colon>'a\<Colon>type) \<in> (A\<Colon>'a\<Colon>type set)"
+  by (metis SigmaD1 0)
+have 3: "(b\<Colon>'b\<Colon>type) \<in> (B\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>type set) (a\<Colon>'a\<Colon>type)"
+  by (metis SigmaD2 0)
+have 4: "(b\<Colon>'b\<Colon>type) \<notin> (B\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>type set) (a\<Colon>'a\<Colon>type)"
+  by (metis 1 2)
+show "False"
+  by (metis 3 4)
+qed
+
+lemma Sigma_triv: "(a,b) \<in> Sigma A B ==> a \<in> A & b \<in> B a"
+by (metis SigmaD1 SigmaD2)
+
+declare [[ atp_problem_prefix = "Abstraction__Sigma_Collect" ]]
+lemma "(a,b) \<in> (SIGMA x: A. {y. x = f y}) ==> a \<in> A & a = f b"
+(*???metis says this is satisfiable!
+by (metis CollectD SigmaD1 SigmaD2)
+*)
+by (meson CollectD SigmaD1 SigmaD2)
+
+
+(*single-step*)
+lemma "(a,b) \<in> (SIGMA x: A. {y. x = f y}) ==> a \<in> A & a = f b"
+by (metis SigmaD1 SigmaD2 insert_def singleton_conv2 Un_empty_right vimage_Collect_eq vimage_def vimage_singleton_eq)
+
+
+lemma "(a,b) \<in> (SIGMA x: A. {y. x = f y}) ==> a \<in> A & a = f b"
+proof (neg_clausify)
+assume 0: "(a\<Colon>'a\<Colon>type, b\<Colon>'b\<Colon>type)
+\<in> Sigma (A\<Colon>'a\<Colon>type set)
+   (COMBB Collect (COMBC (COMBB COMBB op =) (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>type)))"
+assume 1: "(a\<Colon>'a\<Colon>type) \<notin> (A\<Colon>'a\<Colon>type set) \<or> a \<noteq> (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>type) (b\<Colon>'b\<Colon>type)"
+have 2: "(a\<Colon>'a\<Colon>type) \<in> (A\<Colon>'a\<Colon>type set)"
+  by (metis 0 SigmaD1)
+have 3: "(b\<Colon>'b\<Colon>type)
+\<in> COMBB Collect (COMBC (COMBB COMBB op =) (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>type)) (a\<Colon>'a\<Colon>type)"
+  by (metis 0 SigmaD2) 
+have 4: "(b\<Colon>'b\<Colon>type) \<in> Collect (COMBB (op = (a\<Colon>'a\<Colon>type)) (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>type))"
+  by (metis 3)
+have 5: "(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>type) (b\<Colon>'b\<Colon>type) \<noteq> (a\<Colon>'a\<Colon>type)"
+  by (metis 1 2)
+have 6: "(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>type) (b\<Colon>'b\<Colon>type) = (a\<Colon>'a\<Colon>type)"
+  by (metis 4 vimage_singleton_eq insert_def singleton_conv2 Un_empty_right vimage_Collect_eq vimage_def)
+show "False"
+  by (metis 5 6)
+qed
+
+(*Alternative structured proof, untyped*)
+lemma "(a,b) \<in> (SIGMA x: A. {y. x = f y}) ==> a \<in> A & a = f b"
+proof (neg_clausify)
+assume 0: "(a, b) \<in> Sigma A (COMBB Collect (COMBC (COMBB COMBB op =) f))"
+have 1: "b \<in> Collect (COMBB (op = a) f)"
+  by (metis 0 SigmaD2)
+have 2: "f b = a"
+  by (metis 1 vimage_Collect_eq singleton_conv2 insert_def Un_empty_right vimage_singleton_eq vimage_def)
+assume 3: "a \<notin> A \<or> a \<noteq> f b"
+have 4: "a \<in> A"
+  by (metis 0 SigmaD1)
+have 5: "f b \<noteq> a"
+  by (metis 4 3)
+show "False"
+  by (metis 5 2)
+qed
+
+
+declare [[ atp_problem_prefix = "Abstraction__CLF_eq_in_pp" ]]
+lemma "(cl,f) \<in> CLF ==> CLF = (SIGMA cl: CL.{f. f \<in> pset cl}) ==> f \<in> pset cl"
+by (metis Collect_mem_eq SigmaD2)
+
+lemma "(cl,f) \<in> CLF ==> CLF = (SIGMA cl: CL.{f. f \<in> pset cl}) ==> f \<in> pset cl"
+proof (neg_clausify)
+assume 0: "(cl, f) \<in> CLF"
+assume 1: "CLF = Sigma CL (COMBB Collect (COMBB (COMBC op \<in>) pset))"
+assume 2: "f \<notin> pset cl"
+have 3: "\<And>X1 X2. X2 \<in> COMBB Collect (COMBB (COMBC op \<in>) pset) X1 \<or> (X1, X2) \<notin> CLF"
+  by (metis SigmaD2 1)
+have 4: "\<And>X1 X2. X2 \<in> pset X1 \<or> (X1, X2) \<notin> CLF"
+  by (metis 3 Collect_mem_eq)
+have 5: "(cl, f) \<notin> CLF"
+  by (metis 2 4)
+show "False"
+  by (metis 5 0)
+qed
+
+declare [[ atp_problem_prefix = "Abstraction__Sigma_Collect_Pi" ]]
+lemma
+    "(cl,f) \<in> (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl}) ==> 
+    f \<in> pset cl \<rightarrow> pset cl"
+proof (neg_clausify)
+assume 0: "f \<notin> Pi (pset cl) (COMBK (pset cl))"
+assume 1: "(cl, f)
+\<in> Sigma CL
+   (COMBB Collect
+     (COMBB (COMBC op \<in>) (COMBS (COMBB Pi pset) (COMBB COMBK pset))))"
+show "False"
+(*  by (metis 0 Collect_mem_eq SigmaD2 1) ??doesn't terminate*)
+  by (insert 0 1, simp add: COMBB_def COMBS_def COMBC_def)
+qed
+
+
+declare [[ atp_problem_prefix = "Abstraction__Sigma_Collect_Int" ]]
+lemma
+    "(cl,f) \<in> (SIGMA cl: CL. {f. f \<in> pset cl \<inter> cl}) ==>
+   f \<in> pset cl \<inter> cl"
+proof (neg_clausify)
+assume 0: "(cl, f)
+\<in> Sigma CL
+   (COMBB Collect (COMBB (COMBC op \<in>) (COMBS (COMBB op \<inter> pset) COMBI)))"
+assume 1: "f \<notin> pset cl \<inter> cl"
+have 2: "f \<in> COMBB Collect (COMBB (COMBC op \<in>) (COMBS (COMBB op \<inter> pset) COMBI)) cl" 
+  by (insert 0, simp add: COMBB_def) 
+(*  by (metis SigmaD2 0)  ??doesn't terminate*)
+have 3: "f \<in> COMBS (COMBB op \<inter> pset) COMBI cl"
+  by (metis 2 Collect_mem_eq)
+have 4: "f \<notin> cl \<inter> pset cl"
+  by (metis 1 Int_commute)
+have 5: "f \<in> cl \<inter> pset cl"
+  by (metis 3 Int_commute)
+show "False"
+  by (metis 5 4)
+qed
+
+
+declare [[ atp_problem_prefix = "Abstraction__Sigma_Collect_Pi_mono" ]]
+lemma
+    "(cl,f) \<in> (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl & monotone f (pset cl) (order cl)}) ==>
+   (f \<in> pset cl \<rightarrow> pset cl)  &  (monotone f (pset cl) (order cl))"
+by auto
+
+declare [[ atp_problem_prefix = "Abstraction__CLF_subset_Collect_Int" ]]
+lemma "(cl,f) \<in> CLF ==> 
+   CLF \<subseteq> (SIGMA cl: CL. {f. f \<in> pset cl \<inter> cl}) ==>
+   f \<in> pset cl \<inter> cl"
+by auto
+
+(*??no longer terminates, with combinators
+by (metis Collect_mem_eq Int_def SigmaD2 UnCI Un_absorb1)
+  --{*@{text Int_def} is redundant*}
+*)
+
+declare [[ atp_problem_prefix = "Abstraction__CLF_eq_Collect_Int" ]]
+lemma "(cl,f) \<in> CLF ==> 
+   CLF = (SIGMA cl: CL. {f. f \<in> pset cl \<inter> cl}) ==>
+   f \<in> pset cl \<inter> cl"
+by auto
+(*??no longer terminates, with combinators
+by (metis Collect_mem_eq Int_commute SigmaD2)
+*)
+
+declare [[ atp_problem_prefix = "Abstraction__CLF_subset_Collect_Pi" ]]
+lemma 
+   "(cl,f) \<in> CLF ==> 
+    CLF \<subseteq> (SIGMA cl': CL. {f. f \<in> pset cl' \<rightarrow> pset cl'}) ==> 
+    f \<in> pset cl \<rightarrow> pset cl"
+by fast
+(*??no longer terminates, with combinators
+by (metis Collect_mem_eq SigmaD2 subsetD)
+*)
+
+declare [[ atp_problem_prefix = "Abstraction__CLF_eq_Collect_Pi" ]]
+lemma 
+  "(cl,f) \<in> CLF ==> 
+   CLF = (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl}) ==> 
+   f \<in> pset cl \<rightarrow> pset cl"
+by auto
+(*??no longer terminates, with combinators
+by (metis Collect_mem_eq SigmaD2 contra_subsetD equalityE)
+*)
+
+declare [[ atp_problem_prefix = "Abstraction__CLF_eq_Collect_Pi_mono" ]]
+lemma 
+  "(cl,f) \<in> CLF ==> 
+   CLF = (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl & monotone f (pset cl) (order cl)}) ==>
+   (f \<in> pset cl \<rightarrow> pset cl)  &  (monotone f (pset cl) (order cl))"
+by auto
+
+declare [[ atp_problem_prefix = "Abstraction__map_eq_zipA" ]]
+lemma "map (%x. (f x, g x)) xs = zip (map f xs) (map g xs)"
+apply (induct xs)
+(*sledgehammer*)  
+apply auto
+done
+
+declare [[ atp_problem_prefix = "Abstraction__map_eq_zipB" ]]
+lemma "map (%w. (w -> w, w \<times> w)) xs = 
+       zip (map (%w. w -> w) xs) (map (%w. w \<times> w) xs)"
+apply (induct xs)
+(*sledgehammer*)  
+apply auto
+done
+
+declare [[ atp_problem_prefix = "Abstraction__image_evenA" ]]
+lemma "(%x. Suc(f x)) ` {x. even x} <= A ==> (\<forall>x. even x --> Suc(f x) \<in> A)";
+(*sledgehammer*)  
+by auto
+
+declare [[ atp_problem_prefix = "Abstraction__image_evenB" ]]
+lemma "(%x. f (f x)) ` ((%x. Suc(f x)) ` {x. even x}) <= A 
+       ==> (\<forall>x. even x --> f (f (Suc(f x))) \<in> A)";
+(*sledgehammer*)  
+by auto
+
+declare [[ atp_problem_prefix = "Abstraction__image_curry" ]]
+lemma "f \<in> (%u v. b \<times> u \<times> v) ` A ==> \<forall>u v. P (b \<times> u \<times> v) ==> P(f y)" 
+(*sledgehammer*)  
+by auto
+
+declare [[ atp_problem_prefix = "Abstraction__image_TimesA" ]]
+lemma image_TimesA: "(%(x,y). (f x, g y)) ` (A \<times> B) = (f`A) \<times> (g`B)"
+(*sledgehammer*) 
+apply (rule equalityI)
+(***Even the two inclusions are far too difficult
+using [[ atp_problem_prefix = "Abstraction__image_TimesA_simpler"]]
+***)
+apply (rule subsetI)
+apply (erule imageE)
+(*V manages from here with help: Abstraction__image_TimesA_simpler_1_b.p*)
+apply (erule ssubst)
+apply (erule SigmaE)
+(*V manages from here: Abstraction__image_TimesA_simpler_1_a.p*)
+apply (erule ssubst)
+apply (subst split_conv)
+apply (rule SigmaI) 
+apply (erule imageI) +
+txt{*subgoal 2*}
+apply (clarify );
+apply (simp add: );  
+apply (rule rev_image_eqI)  
+apply (blast intro: elim:); 
+apply (simp add: );
+done
+
+(*Given the difficulty of the previous problem, these two are probably
+impossible*)
+
+declare [[ atp_problem_prefix = "Abstraction__image_TimesB" ]]
+lemma image_TimesB:
+    "(%(x,y,z). (f x, g y, h z)) ` (A \<times> B \<times> C) = (f`A) \<times> (g`B) \<times> (h`C)" 
+(*sledgehammer*) 
+by force
+
+declare [[ atp_problem_prefix = "Abstraction__image_TimesC" ]]
+lemma image_TimesC:
+    "(%(x,y). (x \<rightarrow> x, y \<times> y)) ` (A \<times> B) = 
+     ((%x. x \<rightarrow> x) ` A) \<times> ((%y. y \<times> y) ` B)" 
+(*sledgehammer*) 
+by auto
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Metis_Examples/BT.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,243 @@
+(*  Title:      HOL/MetisTest/BT.thy
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+
+Testing the metis method
+*)
+
+header {* Binary trees *}
+
+theory BT
+imports Main
+begin
+
+
+datatype 'a bt =
+    Lf
+  | Br 'a  "'a bt"  "'a bt"
+
+consts
+  n_nodes   :: "'a bt => nat"
+  n_leaves  :: "'a bt => nat"
+  depth     :: "'a bt => nat"
+  reflect   :: "'a bt => 'a bt"
+  bt_map    :: "('a => 'b) => ('a bt => 'b bt)"
+  preorder  :: "'a bt => 'a list"
+  inorder   :: "'a bt => 'a list"
+  postorder :: "'a bt => 'a list"
+  appnd    :: "'a bt => 'a bt => 'a bt"
+
+primrec
+  "n_nodes Lf = 0"
+  "n_nodes (Br a t1 t2) = Suc (n_nodes t1 + n_nodes t2)"
+
+primrec
+  "n_leaves Lf = Suc 0"
+  "n_leaves (Br a t1 t2) = n_leaves t1 + n_leaves t2"
+
+primrec
+  "depth Lf = 0"
+  "depth (Br a t1 t2) = Suc (max (depth t1) (depth t2))"
+
+primrec
+  "reflect Lf = Lf"
+  "reflect (Br a t1 t2) = Br a (reflect t2) (reflect t1)"
+
+primrec
+  "bt_map f Lf = Lf"
+  "bt_map f (Br a t1 t2) = Br (f a) (bt_map f t1) (bt_map f t2)"
+
+primrec
+  "preorder Lf = []"
+  "preorder (Br a t1 t2) = [a] @ (preorder t1) @ (preorder t2)"
+
+primrec
+  "inorder Lf = []"
+  "inorder (Br a t1 t2) = (inorder t1) @ [a] @ (inorder t2)"
+
+primrec
+  "postorder Lf = []"
+  "postorder (Br a t1 t2) = (postorder t1) @ (postorder t2) @ [a]"
+
+primrec
+  "appnd Lf t = t"
+  "appnd (Br a t1 t2) t = Br a (appnd t1 t) (appnd t2 t)"
+
+
+text {* \medskip BT simplification *}
+
+declare [[ atp_problem_prefix = "BT__n_leaves_reflect" ]]
+lemma n_leaves_reflect: "n_leaves (reflect t) = n_leaves t"
+  apply (induct t)
+  apply (metis add_right_cancel n_leaves.simps(1) reflect.simps(1))
+  apply (metis add_commute n_leaves.simps(2) reflect.simps(2))
+  done
+
+declare [[ atp_problem_prefix = "BT__n_nodes_reflect" ]]
+lemma n_nodes_reflect: "n_nodes (reflect t) = n_nodes t"
+  apply (induct t)
+  apply (metis reflect.simps(1))
+  apply (metis n_nodes.simps(2) nat_add_commute reflect.simps(2))
+  done
+
+declare [[ atp_problem_prefix = "BT__depth_reflect" ]]
+lemma depth_reflect: "depth (reflect t) = depth t"
+  apply (induct t)
+  apply (metis depth.simps(1) reflect.simps(1))
+  apply (metis depth.simps(2) min_max.sup_commute reflect.simps(2))
+  done
+
+text {*
+  The famous relationship between the numbers of leaves and nodes.
+*}
+
+declare [[ atp_problem_prefix = "BT__n_leaves_nodes" ]]
+lemma n_leaves_nodes: "n_leaves t = Suc (n_nodes t)"
+  apply (induct t)
+  apply (metis n_leaves.simps(1) n_nodes.simps(1))
+  apply auto
+  done
+
+declare [[ atp_problem_prefix = "BT__reflect_reflect_ident" ]]
+lemma reflect_reflect_ident: "reflect (reflect t) = t"
+  apply (induct t)
+  apply (metis add_right_cancel reflect.simps(1));
+  apply (metis reflect.simps(2))
+  done
+
+declare [[ atp_problem_prefix = "BT__bt_map_ident" ]]
+lemma bt_map_ident: "bt_map (%x. x) = (%y. y)"
+apply (rule ext) 
+apply (induct_tac y)
+  apply (metis bt_map.simps(1))
+txt{*BUG involving flex-flex pairs*}
+(*  apply (metis bt_map.simps(2)) *)
+apply auto
+done
+
+
+declare [[ atp_problem_prefix = "BT__bt_map_appnd" ]]
+lemma bt_map_appnd: "bt_map f (appnd t u) = appnd (bt_map f t) (bt_map f u)"
+apply (induct t)
+  apply (metis appnd.simps(1) bt_map.simps(1))
+  apply (metis appnd.simps(2) bt_map.simps(2))  (*slow!!*)
+done
+
+
+declare [[ atp_problem_prefix = "BT__bt_map_compose" ]]
+lemma bt_map_compose: "bt_map (f o g) t = bt_map f (bt_map g t)"
+apply (induct t) 
+  apply (metis bt_map.simps(1))
+txt{*Metis runs forever*}
+(*  apply (metis bt_map.simps(2) o_apply)*)
+apply auto
+done
+
+
+declare [[ atp_problem_prefix = "BT__bt_map_reflect" ]]
+lemma bt_map_reflect: "bt_map f (reflect t) = reflect (bt_map f t)"
+  apply (induct t)
+  apply (metis add_right_cancel bt_map.simps(1) reflect.simps(1))
+  apply (metis add_right_cancel bt_map.simps(2) reflect.simps(2))
+  done
+
+declare [[ atp_problem_prefix = "BT__preorder_bt_map" ]]
+lemma preorder_bt_map: "preorder (bt_map f t) = map f (preorder t)"
+  apply (induct t)
+  apply (metis bt_map.simps(1) map.simps(1) preorder.simps(1))
+   apply simp
+  done
+
+declare [[ atp_problem_prefix = "BT__inorder_bt_map" ]]
+lemma inorder_bt_map: "inorder (bt_map f t) = map f (inorder t)"
+  apply (induct t)
+  apply (metis bt_map.simps(1) inorder.simps(1) map.simps(1))
+  apply simp
+  done
+
+declare [[ atp_problem_prefix = "BT__postorder_bt_map" ]]
+lemma postorder_bt_map: "postorder (bt_map f t) = map f (postorder t)"
+  apply (induct t)
+  apply (metis bt_map.simps(1) map.simps(1) postorder.simps(1))
+   apply simp
+  done
+
+declare [[ atp_problem_prefix = "BT__depth_bt_map" ]]
+lemma depth_bt_map [simp]: "depth (bt_map f t) = depth t"
+  apply (induct t)
+  apply (metis bt_map.simps(1) depth.simps(1))
+   apply simp
+  done
+
+declare [[ atp_problem_prefix = "BT__n_leaves_bt_map" ]]
+lemma n_leaves_bt_map [simp]: "n_leaves (bt_map f t) = n_leaves t"
+  apply (induct t)
+  apply (metis One_nat_def Suc_eq_plus1 bt_map.simps(1) less_add_one less_antisym linorder_neq_iff n_leaves.simps(1))
+  apply (metis bt_map.simps(2) n_leaves.simps(2))
+  done
+
+
+declare [[ atp_problem_prefix = "BT__preorder_reflect" ]]
+lemma preorder_reflect: "preorder (reflect t) = rev (postorder t)"
+  apply (induct t)
+  apply (metis postorder.simps(1) preorder.simps(1) reflect.simps(1) rev_is_Nil_conv)
+  apply (metis append_Nil Cons_eq_append_conv postorder.simps(2) preorder.simps(2) reflect.simps(2) rev.simps(2) rev_append rev_rev_ident)
+  done
+
+declare [[ atp_problem_prefix = "BT__inorder_reflect" ]]
+lemma inorder_reflect: "inorder (reflect t) = rev (inorder t)"
+  apply (induct t)
+  apply (metis inorder.simps(1) reflect.simps(1) rev.simps(1))
+  apply simp
+  done
+
+declare [[ atp_problem_prefix = "BT__postorder_reflect" ]]
+lemma postorder_reflect: "postorder (reflect t) = rev (preorder t)"
+  apply (induct t)
+  apply (metis postorder.simps(1) preorder.simps(1) reflect.simps(1) rev.simps(1))
+  apply (metis Cons_eq_appendI postorder.simps(2) preorder.simps(2) reflect.simps(2) rev.simps(2) rev_append self_append_conv2)
+  done
+
+text {*
+ Analogues of the standard properties of the append function for lists.
+*}
+
+declare [[ atp_problem_prefix = "BT__appnd_assoc" ]]
+lemma appnd_assoc [simp]:
+     "appnd (appnd t1 t2) t3 = appnd t1 (appnd t2 t3)"
+  apply (induct t1)
+  apply (metis appnd.simps(1))
+  apply (metis appnd.simps(2))
+  done
+
+declare [[ atp_problem_prefix = "BT__appnd_Lf2" ]]
+lemma appnd_Lf2 [simp]: "appnd t Lf = t"
+  apply (induct t)
+  apply (metis appnd.simps(1))
+  apply (metis appnd.simps(2))
+  done
+
+declare [[ atp_problem_prefix = "BT__depth_appnd" ]]
+  declare max_add_distrib_left [simp]
+lemma depth_appnd [simp]: "depth (appnd t1 t2) = depth t1 + depth t2"
+  apply (induct t1)
+  apply (metis add_0 appnd.simps(1) depth.simps(1))
+apply (simp add: ); 
+  done
+
+declare [[ atp_problem_prefix = "BT__n_leaves_appnd" ]]
+lemma n_leaves_appnd [simp]:
+     "n_leaves (appnd t1 t2) = n_leaves t1 * n_leaves t2"
+  apply (induct t1)
+  apply (metis One_nat_def appnd.simps(1) less_irrefl less_linear n_leaves.simps(1) nat_mult_1) 
+  apply (simp add: left_distrib)
+  done
+
+declare [[ atp_problem_prefix = "BT__bt_map_appnd" ]]
+lemma (*bt_map_appnd:*)
+     "bt_map f (appnd t1 t2) = appnd (bt_map f t1) (bt_map f t2)"
+  apply (induct t1)
+  apply (metis appnd.simps(1) bt_map_appnd)
+  apply (metis bt_map_appnd)
+  done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Metis_Examples/BigO.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,1231 @@
+(*  Title:      HOL/Metis_Examples/BigO.thy
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+
+Testing the metis method.
+*)
+
+header {* Big O notation *}
+
+theory BigO
+imports "~~/src/HOL/Decision_Procs/Dense_Linear_Order" Main SetsAndFunctions 
+begin
+
+subsection {* Definitions *}
+
+definition bigo :: "('a => 'b::ordered_idom) => ('a => 'b) set"    ("(1O'(_'))") where
+  "O(f::('a => 'b)) ==   {h. EX c. ALL x. abs (h x) <= c * abs (f x)}"
+
+declare [[ atp_problem_prefix = "BigO__bigo_pos_const" ]]
+lemma bigo_pos_const: "(EX (c::'a::ordered_idom). 
+    ALL x. (abs (h x)) <= (c * (abs (f x))))
+      = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))"
+  apply auto
+  apply (case_tac "c = 0", simp)
+  apply (rule_tac x = "1" in exI, simp)
+  apply (rule_tac x = "abs c" in exI, auto)
+  apply (metis abs_ge_minus_self abs_ge_zero abs_minus_cancel abs_of_nonneg equation_minus_iff Orderings.xt1(6) abs_mult)
+  done
+
+(*** Now various verions with an increasing modulus ***)
+
+declare [[sledgehammer_modulus = 1]]
+
+lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_idom). 
+    ALL x. (abs (h x)) <= (c * (abs (f x))))
+      = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))"
+  apply auto
+  apply (case_tac "c = 0", simp)
+  apply (rule_tac x = "1" in exI, simp)
+  apply (rule_tac x = "abs c" in exI, auto)
+proof (neg_clausify)
+fix c x
+have 0: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. \<bar>X1 * X2\<bar> = \<bar>X2 * X1\<bar>"
+  by (metis abs_mult mult_commute)
+have 1: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom.
+   X1 \<le> (0\<Colon>'a\<Colon>ordered_idom) \<or> \<bar>X2\<bar> * X1 = \<bar>X2 * X1\<bar>"
+  by (metis abs_mult_pos linorder_linear)
+have 2: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom.
+   \<not> (0\<Colon>'a\<Colon>ordered_idom) < X1 * X2 \<or>
+   \<not> (0\<Colon>'a\<Colon>ordered_idom) \<le> X2 \<or> \<not> X1 \<le> (0\<Colon>'a\<Colon>ordered_idom)"
+  by (metis linorder_not_less mult_nonneg_nonpos2)
+assume 3: "\<And>x\<Colon>'b\<Colon>type.
+   \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>
+   \<le> (c\<Colon>'a\<Colon>ordered_idom) * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>"
+assume 4: "\<not> \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar>
+  \<le> \<bar>c\<Colon>'a\<Colon>ordered_idom\<bar> * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>"
+have 5: "\<not> \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar>
+  \<le> \<bar>(c\<Colon>'a\<Colon>ordered_idom) * (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>"
+  by (metis 4 abs_mult)
+have 6: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom.
+   \<not> X1 \<le> (0\<Colon>'a\<Colon>ordered_idom) \<or> X1 \<le> \<bar>X2\<bar>"
+  by (metis abs_ge_zero xt1(6))
+have 7: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom.
+   X1 \<le> \<bar>X2\<bar> \<or> (0\<Colon>'a\<Colon>ordered_idom) < X1"
+  by (metis not_leE 6)
+have 8: "(0\<Colon>'a\<Colon>ordered_idom) < \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar>"
+  by (metis 5 7)
+have 9: "\<And>X1\<Colon>'a\<Colon>ordered_idom.
+   \<not> \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar> \<le> X1 \<or>
+   (0\<Colon>'a\<Colon>ordered_idom) < X1"
+  by (metis 8 order_less_le_trans)
+have 10: "(0\<Colon>'a\<Colon>ordered_idom)
+< (c\<Colon>'a\<Colon>ordered_idom) * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar>"
+  by (metis 3 9)
+have 11: "\<not> (c\<Colon>'a\<Colon>ordered_idom) \<le> (0\<Colon>'a\<Colon>ordered_idom)"
+  by (metis abs_ge_zero 2 10)
+have 12: "\<And>X1\<Colon>'a\<Colon>ordered_idom. (c\<Colon>'a\<Colon>ordered_idom) * \<bar>X1\<bar> = \<bar>X1 * c\<bar>"
+  by (metis mult_commute 1 11)
+have 13: "\<And>X1\<Colon>'b\<Colon>type.
+   - (h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1
+   \<le> (c\<Colon>'a\<Colon>ordered_idom) * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1\<bar>"
+  by (metis 3 abs_le_D2)
+have 14: "\<And>X1\<Colon>'b\<Colon>type.
+   - (h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1
+   \<le> \<bar>(c\<Colon>'a\<Colon>ordered_idom) * (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1\<bar>"
+  by (metis 0 12 13)
+have 15: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. \<bar>X1 * \<bar>X2\<bar>\<bar> = \<bar>X1 * X2\<bar>"
+  by (metis abs_mult abs_mult_pos abs_ge_zero)
+have 16: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. X1 \<le> \<bar>X2\<bar> \<or> \<not> X1 \<le> X2"
+  by (metis xt1(6) abs_ge_self)
+have 17: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. \<not> \<bar>X1\<bar> \<le> X2 \<or> X1 \<le> \<bar>X2\<bar>"
+  by (metis 16 abs_le_D1)
+have 18: "\<And>X1\<Colon>'b\<Colon>type.
+   (h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1
+   \<le> \<bar>(c\<Colon>'a\<Colon>ordered_idom) * (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1\<bar>"
+  by (metis 17 3 15)
+show "False"
+  by (metis abs_le_iff 5 18 14)
+qed
+
+declare [[sledgehammer_modulus = 2]]
+
+lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_idom). 
+    ALL x. (abs (h x)) <= (c * (abs (f x))))
+      = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))"
+  apply auto
+  apply (case_tac "c = 0", simp)
+  apply (rule_tac x = "1" in exI, simp)
+  apply (rule_tac x = "abs c" in exI, auto);
+proof (neg_clausify)
+fix c x
+have 0: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. \<bar>X1 * X2\<bar> = \<bar>X2 * X1\<bar>"
+  by (metis abs_mult mult_commute)
+assume 1: "\<And>x\<Colon>'b\<Colon>type.
+   \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>
+   \<le> (c\<Colon>'a\<Colon>ordered_idom) * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>"
+assume 2: "\<not> \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar>
+  \<le> \<bar>c\<Colon>'a\<Colon>ordered_idom\<bar> * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>"
+have 3: "\<not> \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar>
+  \<le> \<bar>(c\<Colon>'a\<Colon>ordered_idom) * (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>"
+  by (metis 2 abs_mult)
+have 4: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom.
+   \<not> X1 \<le> (0\<Colon>'a\<Colon>ordered_idom) \<or> X1 \<le> \<bar>X2\<bar>"
+  by (metis abs_ge_zero xt1(6))
+have 5: "(0\<Colon>'a\<Colon>ordered_idom) < \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar>"
+  by (metis not_leE 4 3)
+have 6: "(0\<Colon>'a\<Colon>ordered_idom)
+< (c\<Colon>'a\<Colon>ordered_idom) * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar>"
+  by (metis 1 order_less_le_trans 5)
+have 7: "\<And>X1\<Colon>'a\<Colon>ordered_idom. (c\<Colon>'a\<Colon>ordered_idom) * \<bar>X1\<bar> = \<bar>X1 * c\<bar>"
+  by (metis abs_ge_zero linorder_not_less mult_nonneg_nonpos2 6 linorder_linear abs_mult_pos mult_commute)
+have 8: "\<And>X1\<Colon>'b\<Colon>type.
+   - (h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1
+   \<le> \<bar>(c\<Colon>'a\<Colon>ordered_idom) * (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1\<bar>"
+  by (metis 0 7 abs_le_D2 1)
+have 9: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. \<not> \<bar>X1\<bar> \<le> X2 \<or> X1 \<le> \<bar>X2\<bar>"
+  by (metis abs_ge_self xt1(6) abs_le_D1)
+show "False"
+  by (metis 8 abs_ge_zero abs_mult_pos abs_mult 1 9 3 abs_le_iff)
+qed
+
+declare [[sledgehammer_modulus = 3]]
+
+lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_idom). 
+    ALL x. (abs (h x)) <= (c * (abs (f x))))
+      = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))"
+  apply auto
+  apply (case_tac "c = 0", simp)
+  apply (rule_tac x = "1" in exI, simp)
+  apply (rule_tac x = "abs c" in exI, auto);
+proof (neg_clausify)
+fix c x
+assume 0: "\<And>x\<Colon>'b\<Colon>type.
+   \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>
+   \<le> (c\<Colon>'a\<Colon>ordered_idom) * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>"
+assume 1: "\<not> \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar>
+  \<le> \<bar>c\<Colon>'a\<Colon>ordered_idom\<bar> * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>"
+have 2: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom.
+   X1 \<le> \<bar>X2\<bar> \<or> (0\<Colon>'a\<Colon>ordered_idom) < X1"
+  by (metis abs_ge_zero xt1(6) not_leE)
+have 3: "\<not> (c\<Colon>'a\<Colon>ordered_idom) \<le> (0\<Colon>'a\<Colon>ordered_idom)"
+  by (metis abs_ge_zero mult_nonneg_nonpos2 linorder_not_less order_less_le_trans 1 abs_mult 2 0)
+have 4: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. \<bar>X1 * \<bar>X2\<bar>\<bar> = \<bar>X1 * X2\<bar>"
+  by (metis abs_ge_zero abs_mult_pos abs_mult)
+have 5: "\<And>X1\<Colon>'b\<Colon>type.
+   (h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1
+   \<le> \<bar>(c\<Colon>'a\<Colon>ordered_idom) * (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1\<bar>"
+  by (metis 4 0 xt1(6) abs_ge_self abs_le_D1)
+show "False"
+  by (metis abs_mult mult_commute 3 abs_mult_pos linorder_linear 0 abs_le_D2 5 1 abs_le_iff)
+qed
+
+
+declare [[sledgehammer_modulus = 1]]
+
+lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_idom). 
+    ALL x. (abs (h x)) <= (c * (abs (f x))))
+      = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))"
+  apply auto
+  apply (case_tac "c = 0", simp)
+  apply (rule_tac x = "1" in exI, simp)
+  apply (rule_tac x = "abs c" in exI, auto);
+proof (neg_clausify)
+fix c x  (*sort/type constraint inserted by hand!*)
+have 0: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2. \<bar>X1 * \<bar>X2\<bar>\<bar> = \<bar>X1 * X2\<bar>"
+  by (metis abs_ge_zero abs_mult_pos abs_mult)
+assume 1: "\<And>A. \<bar>h A\<bar> \<le> c * \<bar>f A\<bar>"
+have 2: "\<And>X1 X2. \<not> \<bar>X1\<bar> \<le> X2 \<or> (0\<Colon>'a) \<le> X2"
+  by (metis abs_ge_zero order_trans)
+have 3: "\<And>X1. (0\<Colon>'a) \<le> c * \<bar>f X1\<bar>"
+  by (metis 1 2)
+have 4: "\<And>X1. c * \<bar>f X1\<bar> = \<bar>c * f X1\<bar>"
+  by (metis 0 abs_of_nonneg 3)
+have 5: "\<And>X1. - h X1 \<le> c * \<bar>f X1\<bar>"
+  by (metis 1 abs_le_D2)
+have 6: "\<And>X1. - h X1 \<le> \<bar>c * f X1\<bar>"
+  by (metis 4 5)
+have 7: "\<And>X1. h X1 \<le> c * \<bar>f X1\<bar>"
+  by (metis 1 abs_le_D1)
+have 8: "\<And>X1. h X1 \<le> \<bar>c * f X1\<bar>"
+  by (metis 4 7)
+assume 9: "\<not> \<bar>h x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>"
+have 10: "\<not> \<bar>h x\<bar> \<le> \<bar>c * f x\<bar>"
+  by (metis abs_mult 9)
+show "False"
+  by (metis 6 8 10 abs_leI)
+qed
+
+
+declare [[sledgehammer_sorts = true]]
+
+lemma bigo_alt_def: "O(f) = 
+    {h. EX c. (0 < c & (ALL x. abs (h x) <= c * abs (f x)))}"
+by (auto simp add: bigo_def bigo_pos_const)
+
+declare [[ atp_problem_prefix = "BigO__bigo_elt_subset" ]]
+lemma bigo_elt_subset [intro]: "f : O(g) ==> O(f) <= O(g)"
+  apply (auto simp add: bigo_alt_def)
+  apply (rule_tac x = "ca * c" in exI)
+  apply (rule conjI)
+  apply (rule mult_pos_pos)
+  apply (assumption)+ 
+(*sledgehammer*);
+  apply (rule allI)
+  apply (drule_tac x = "xa" in spec)+
+  apply (subgoal_tac "ca * abs(f xa) <= ca * (c * abs(g xa))");
+  apply (erule order_trans)
+  apply (simp add: mult_ac)
+  apply (rule mult_left_mono, assumption)
+  apply (rule order_less_imp_le, assumption);
+done
+
+
+declare [[ atp_problem_prefix = "BigO__bigo_refl" ]]
+lemma bigo_refl [intro]: "f : O(f)"
+  apply(auto simp add: bigo_def)
+proof (neg_clausify)
+fix x
+assume 0: "\<And>xa. \<not> \<bar>f (x xa)\<bar> \<le> xa * \<bar>f (x xa)\<bar>"
+have 1: "\<And>X2. X2 \<le> (1\<Colon>'b) * X2 \<or> \<not> (1\<Colon>'b) \<le> (1\<Colon>'b)"
+  by (metis mult_le_cancel_right1 order_eq_iff)
+have 2: "\<And>X2. X2 \<le> (1\<Colon>'b) * X2"
+  by (metis order_eq_iff 1)
+show "False"
+  by (metis 0 2)
+qed
+
+declare [[ atp_problem_prefix = "BigO__bigo_zero" ]]
+lemma bigo_zero: "0 : O(g)"
+  apply (auto simp add: bigo_def func_zero)
+proof (neg_clausify)
+fix x
+assume 0: "\<And>xa. \<not> (0\<Colon>'b) \<le> xa * \<bar>g (x xa)\<bar>"
+have 1: "\<not> (0\<Colon>'b) \<le> (0\<Colon>'b)"
+  by (metis 0 mult_eq_0_iff)
+show "False"
+  by (metis 1 linorder_neq_iff linorder_antisym_conv1)
+qed
+
+lemma bigo_zero2: "O(%x.0) = {%x.0}"
+  apply (auto simp add: bigo_def) 
+  apply (rule ext)
+  apply auto
+done
+
+lemma bigo_plus_self_subset [intro]: 
+  "O(f) \<oplus> O(f) <= O(f)"
+  apply (auto simp add: bigo_alt_def set_plus_def)
+  apply (rule_tac x = "c + ca" in exI)
+  apply auto
+  apply (simp add: ring_distribs func_plus)
+  apply (blast intro:order_trans abs_triangle_ineq add_mono elim:) 
+done
+
+lemma bigo_plus_idemp [simp]: "O(f) \<oplus> O(f) = O(f)"
+  apply (rule equalityI)
+  apply (rule bigo_plus_self_subset)
+  apply (rule set_zero_plus2) 
+  apply (rule bigo_zero)
+done
+
+lemma bigo_plus_subset [intro]: "O(f + g) <= O(f) \<oplus> O(g)"
+  apply (rule subsetI)
+  apply (auto simp add: bigo_def bigo_pos_const func_plus set_plus_def)
+  apply (subst bigo_pos_const [symmetric])+
+  apply (rule_tac x = 
+    "%n. if abs (g n) <= (abs (f n)) then x n else 0" in exI)
+  apply (rule conjI)
+  apply (rule_tac x = "c + c" in exI)
+  apply (clarsimp)
+  apply (auto)
+  apply (subgoal_tac "c * abs (f xa + g xa) <= (c + c) * abs (f xa)")
+  apply (erule_tac x = xa in allE)
+  apply (erule order_trans)
+  apply (simp)
+  apply (subgoal_tac "c * abs (f xa + g xa) <= c * (abs (f xa) + abs (g xa))")
+  apply (erule order_trans)
+  apply (simp add: ring_distribs)
+  apply (rule mult_left_mono)
+  apply assumption
+  apply (simp add: order_less_le)
+  apply (rule mult_left_mono)
+  apply (simp add: abs_triangle_ineq)
+  apply (simp add: order_less_le)
+  apply (rule mult_nonneg_nonneg)
+  apply (rule add_nonneg_nonneg)
+  apply auto
+  apply (rule_tac x = "%n. if (abs (f n)) <  abs (g n) then x n else 0" 
+     in exI)
+  apply (rule conjI)
+  apply (rule_tac x = "c + c" in exI)
+  apply auto
+  apply (subgoal_tac "c * abs (f xa + g xa) <= (c + c) * abs (g xa)")
+  apply (erule_tac x = xa in allE)
+  apply (erule order_trans)
+  apply (simp)
+  apply (subgoal_tac "c * abs (f xa + g xa) <= c * (abs (f xa) + abs (g xa))")
+  apply (erule order_trans)
+  apply (simp add: ring_distribs)
+  apply (rule mult_left_mono)
+  apply (simp add: order_less_le)
+  apply (simp add: order_less_le)
+  apply (rule mult_left_mono)
+  apply (rule abs_triangle_ineq)
+  apply (simp add: order_less_le)
+apply (metis abs_not_less_zero double_less_0_iff less_not_permute linorder_not_less mult_less_0_iff)
+  apply (rule ext)
+  apply (auto simp add: if_splits linorder_not_le)
+done
+
+lemma bigo_plus_subset2 [intro]: "A <= O(f) ==> B <= O(f) ==> A \<oplus> B <= O(f)"
+  apply (subgoal_tac "A \<oplus> B <= O(f) \<oplus> O(f)")
+  apply (erule order_trans)
+  apply simp
+  apply (auto del: subsetI simp del: bigo_plus_idemp)
+done
+
+declare [[ atp_problem_prefix = "BigO__bigo_plus_eq" ]]
+lemma bigo_plus_eq: "ALL x. 0 <= f x ==> ALL x. 0 <= g x ==> 
+  O(f + g) = O(f) \<oplus> O(g)"
+  apply (rule equalityI)
+  apply (rule bigo_plus_subset)
+  apply (simp add: bigo_alt_def set_plus_def func_plus)
+  apply clarify 
+(*sledgehammer*); 
+  apply (rule_tac x = "max c ca" in exI)
+  apply (rule conjI)
+   apply (metis Orderings.less_max_iff_disj)
+  apply clarify
+  apply (drule_tac x = "xa" in spec)+
+  apply (subgoal_tac "0 <= f xa + g xa")
+  apply (simp add: ring_distribs)
+  apply (subgoal_tac "abs(a xa + b xa) <= abs(a xa) + abs(b xa)")
+  apply (subgoal_tac "abs(a xa) + abs(b xa) <= 
+      max c ca * f xa + max c ca * g xa")
+  apply (blast intro: order_trans)
+  defer 1
+  apply (rule abs_triangle_ineq)
+  apply (metis add_nonneg_nonneg)
+  apply (rule add_mono)
+using [[ atp_problem_prefix = "BigO__bigo_plus_eq_simpler" ]] 
+(*Found by SPASS; SLOW*)
+apply (metis le_maxI2 linorder_linear linorder_not_le min_max.sup_absorb1 mult_le_cancel_right order_trans)
+apply (metis le_maxI2 linorder_not_le mult_le_cancel_right order_trans)
+done
+
+declare [[ atp_problem_prefix = "BigO__bigo_bounded_alt" ]]
+lemma bigo_bounded_alt: "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> 
+    f : O(g)" 
+  apply (auto simp add: bigo_def)
+(*Version 1: one-shot proof*)
+  apply (metis OrderedGroup.abs_le_D1 linorder_class.not_less  order_less_le  Orderings.xt1(12)  Ring_and_Field.abs_mult)
+  done
+
+lemma (*bigo_bounded_alt:*) "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> 
+    f : O(g)" 
+  apply (auto simp add: bigo_def)
+(*Version 2: single-step proof*)
+proof (neg_clausify)
+fix x
+assume 0: "\<And>x. f x \<le> c * g x"
+assume 1: "\<And>xa. \<not> f (x xa) \<le> xa * \<bar>g (x xa)\<bar>"
+have 2: "\<And>X3. c * g X3 = f X3 \<or> \<not> c * g X3 \<le> f X3"
+  by (metis 0 order_antisym_conv)
+have 3: "\<And>X3. \<not> f (x \<bar>X3\<bar>) \<le> \<bar>X3 * g (x \<bar>X3\<bar>)\<bar>"
+  by (metis 1 abs_mult)
+have 4: "\<And>X1 X3\<Colon>'b\<Colon>ordered_idom. X3 \<le> X1 \<or> X1 \<le> \<bar>X3\<bar>"
+  by (metis linorder_linear abs_le_D1)
+have 5: "\<And>X3::'b. \<bar>X3\<bar> * \<bar>X3\<bar> = X3 * X3"
+  by (metis abs_mult_self)
+have 6: "\<And>X3. \<not> X3 * X3 < (0\<Colon>'b\<Colon>ordered_idom)"
+  by (metis not_square_less_zero)
+have 7: "\<And>X1 X3::'b. \<bar>X1\<bar> * \<bar>X3\<bar> = \<bar>X3 * X1\<bar>"
+  by (metis abs_mult mult_commute)
+have 8: "\<And>X3::'b. X3 * X3 = \<bar>X3 * X3\<bar>"
+  by (metis abs_mult 5)
+have 9: "\<And>X3. X3 * g (x \<bar>X3\<bar>) \<le> f (x \<bar>X3\<bar>)"
+  by (metis 3 4)
+have 10: "c * g (x \<bar>c\<bar>) = f (x \<bar>c\<bar>)"
+  by (metis 2 9)
+have 11: "\<And>X3::'b. \<bar>X3\<bar> * \<bar>\<bar>X3\<bar>\<bar> = \<bar>X3\<bar> * \<bar>X3\<bar>"
+  by (metis abs_idempotent abs_mult 8)
+have 12: "\<And>X3::'b. \<bar>X3 * \<bar>X3\<bar>\<bar> = \<bar>X3\<bar> * \<bar>X3\<bar>"
+  by (metis mult_commute 7 11)
+have 13: "\<And>X3::'b. \<bar>X3 * \<bar>X3\<bar>\<bar> = X3 * X3"
+  by (metis 8 7 12)
+have 14: "\<And>X3. X3 \<le> \<bar>X3\<bar> \<or> X3 < (0\<Colon>'b)"
+  by (metis abs_ge_self abs_le_D1 abs_if)
+have 15: "\<And>X3. X3 \<le> \<bar>X3\<bar> \<or> \<bar>X3\<bar> < (0\<Colon>'b)"
+  by (metis abs_ge_self abs_le_D1 abs_if)
+have 16: "\<And>X3. X3 * X3 < (0\<Colon>'b) \<or> X3 * \<bar>X3\<bar> \<le> X3 * X3"
+  by (metis 15 13)
+have 17: "\<And>X3::'b. X3 * \<bar>X3\<bar> \<le> X3 * X3"
+  by (metis 16 6)
+have 18: "\<And>X3. X3 \<le> \<bar>X3\<bar> \<or> \<not> X3 < (0\<Colon>'b)"
+  by (metis mult_le_cancel_left 17)
+have 19: "\<And>X3::'b. X3 \<le> \<bar>X3\<bar>"
+  by (metis 18 14)
+have 20: "\<not> f (x \<bar>c\<bar>) \<le> \<bar>f (x \<bar>c\<bar>)\<bar>"
+  by (metis 3 10)
+show "False"
+  by (metis 20 19)
+qed
+
+
+text{*So here is the easier (and more natural) problem using transitivity*}
+declare [[ atp_problem_prefix = "BigO__bigo_bounded_alt_trans" ]]
+lemma "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)" 
+  apply (auto simp add: bigo_def)
+  (*Version 1: one-shot proof*) 
+  apply (metis Orderings.leD Orderings.leI abs_ge_self abs_le_D1 abs_mult abs_of_nonneg order_le_less)
+  done
+
+text{*So here is the easier (and more natural) problem using transitivity*}
+declare [[ atp_problem_prefix = "BigO__bigo_bounded_alt_trans" ]]
+lemma "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)" 
+  apply (auto simp add: bigo_def)
+(*Version 2: single-step proof*)
+proof (neg_clausify)
+fix x
+assume 0: "\<And>A\<Colon>'a\<Colon>type.
+   (f\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) A
+   \<le> (c\<Colon>'b\<Colon>ordered_idom) * (g\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) A"
+assume 1: "\<And>A\<Colon>'b\<Colon>ordered_idom.
+   \<not> (f\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) ((x\<Colon>'b\<Colon>ordered_idom \<Rightarrow> 'a\<Colon>type) A)
+     \<le> A * \<bar>(g\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) (x A)\<bar>"
+have 2: "\<And>X2\<Colon>'a\<Colon>type.
+   \<not> (c\<Colon>'b\<Colon>ordered_idom) * (g\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) X2
+     < (f\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) X2"
+  by (metis 0 linorder_not_le)
+have 3: "\<And>X2\<Colon>'b\<Colon>ordered_idom.
+   \<not> (f\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) ((x\<Colon>'b\<Colon>ordered_idom \<Rightarrow> 'a\<Colon>type) \<bar>X2\<bar>)
+     \<le> \<bar>X2 * (g\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) (x \<bar>X2\<bar>)\<bar>"
+  by (metis abs_mult 1)
+have 4: "\<And>X2\<Colon>'b\<Colon>ordered_idom.
+   \<bar>X2 * (g\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) ((x\<Colon>'b\<Colon>ordered_idom \<Rightarrow> 'a\<Colon>type) \<bar>X2\<bar>)\<bar>
+   < (f\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) (x \<bar>X2\<bar>)"
+  by (metis 3 linorder_not_less)
+have 5: "\<And>X2\<Colon>'b\<Colon>ordered_idom.
+   X2 * (g\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) ((x\<Colon>'b\<Colon>ordered_idom \<Rightarrow> 'a\<Colon>type) \<bar>X2\<bar>)
+   < (f\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) (x \<bar>X2\<bar>)"
+  by (metis abs_less_iff 4)
+show "False"
+  by (metis 2 5)
+qed
+
+
+lemma bigo_bounded: "ALL x. 0 <= f x ==> ALL x. f x <= g x ==> 
+    f : O(g)" 
+  apply (erule bigo_bounded_alt [of f 1 g])
+  apply simp
+done
+
+declare [[ atp_problem_prefix = "BigO__bigo_bounded2" ]]
+lemma bigo_bounded2: "ALL x. lb x <= f x ==> ALL x. f x <= lb x + g x ==>
+    f : lb +o O(g)"
+  apply (rule set_minus_imp_plus)
+  apply (rule bigo_bounded)
+  apply (auto simp add: diff_minus fun_Compl_def func_plus)
+  prefer 2
+  apply (drule_tac x = x in spec)+ 
+  apply arith (*not clear that it's provable otherwise*) 
+proof (neg_clausify)
+fix x
+assume 0: "\<And>y. lb y \<le> f y"
+assume 1: "\<not> (0\<Colon>'b) \<le> f x + - lb x"
+have 2: "\<And>X3. (0\<Colon>'b) + X3 = X3"
+  by (metis diff_eq_eq right_minus_eq)
+have 3: "\<not> (0\<Colon>'b) \<le> f x - lb x"
+  by (metis 1 diff_minus)
+have 4: "\<not> (0\<Colon>'b) + lb x \<le> f x"
+  by (metis 3 le_diff_eq)
+show "False"
+  by (metis 4 2 0)
+qed
+
+declare [[ atp_problem_prefix = "BigO__bigo_abs" ]]
+lemma bigo_abs: "(%x. abs(f x)) =o O(f)" 
+  apply (unfold bigo_def)
+  apply auto
+proof (neg_clausify)
+fix x
+assume 0: "\<And>xa. \<not> \<bar>f (x xa)\<bar> \<le> xa * \<bar>f (x xa)\<bar>"
+have 1: "\<And>X2. X2 \<le> (1\<Colon>'b) * X2 \<or> \<not> (1\<Colon>'b) \<le> (1\<Colon>'b)"
+  by (metis mult_le_cancel_right1 order_eq_iff)
+have 2: "\<And>X2. X2 \<le> (1\<Colon>'b) * X2"
+  by (metis order_eq_iff 1)
+show "False"
+  by (metis 0 2)
+qed
+
+declare [[ atp_problem_prefix = "BigO__bigo_abs2" ]]
+lemma bigo_abs2: "f =o O(%x. abs(f x))"
+  apply (unfold bigo_def)
+  apply auto
+proof (neg_clausify)
+fix x
+assume 0: "\<And>xa. \<not> \<bar>f (x xa)\<bar> \<le> xa * \<bar>f (x xa)\<bar>"
+have 1: "\<And>X2. X2 \<le> (1\<Colon>'b) * X2 \<or> \<not> (1\<Colon>'b) \<le> (1\<Colon>'b)"
+  by (metis mult_le_cancel_right1 order_eq_iff)
+have 2: "\<And>X2. X2 \<le> (1\<Colon>'b) * X2"
+  by (metis order_eq_iff 1)
+show "False"
+  by (metis 0 2)
+qed
+ 
+lemma bigo_abs3: "O(f) = O(%x. abs(f x))"
+  apply (rule equalityI)
+  apply (rule bigo_elt_subset)
+  apply (rule bigo_abs2)
+  apply (rule bigo_elt_subset)
+  apply (rule bigo_abs)
+done
+
+lemma bigo_abs4: "f =o g +o O(h) ==> 
+    (%x. abs (f x)) =o (%x. abs (g x)) +o O(h)"
+  apply (drule set_plus_imp_minus)
+  apply (rule set_minus_imp_plus)
+  apply (subst fun_diff_def)
+proof -
+  assume a: "f - g : O(h)"
+  have "(%x. abs (f x) - abs (g x)) =o O(%x. abs(abs (f x) - abs (g x)))"
+    by (rule bigo_abs2)
+  also have "... <= O(%x. abs (f x - g x))"
+    apply (rule bigo_elt_subset)
+    apply (rule bigo_bounded)
+    apply force
+    apply (rule allI)
+    apply (rule abs_triangle_ineq3)
+    done
+  also have "... <= O(f - g)"
+    apply (rule bigo_elt_subset)
+    apply (subst fun_diff_def)
+    apply (rule bigo_abs)
+    done
+  also have "... <= O(h)"
+    using a by (rule bigo_elt_subset)
+  finally show "(%x. abs (f x) - abs (g x)) : O(h)".
+qed
+
+lemma bigo_abs5: "f =o O(g) ==> (%x. abs(f x)) =o O(g)" 
+by (unfold bigo_def, auto)
+
+lemma bigo_elt_subset2 [intro]: "f : g +o O(h) ==> O(f) <= O(g) \<oplus> O(h)"
+proof -
+  assume "f : g +o O(h)"
+  also have "... <= O(g) \<oplus> O(h)"
+    by (auto del: subsetI)
+  also have "... = O(%x. abs(g x)) \<oplus> O(%x. abs(h x))"
+    apply (subst bigo_abs3 [symmetric])+
+    apply (rule refl)
+    done
+  also have "... = O((%x. abs(g x)) + (%x. abs(h x)))"
+    by (rule bigo_plus_eq [symmetric], auto)
+  finally have "f : ...".
+  then have "O(f) <= ..."
+    by (elim bigo_elt_subset)
+  also have "... = O(%x. abs(g x)) \<oplus> O(%x. abs(h x))"
+    by (rule bigo_plus_eq, auto)
+  finally show ?thesis
+    by (simp add: bigo_abs3 [symmetric])
+qed
+
+declare [[ atp_problem_prefix = "BigO__bigo_mult" ]]
+lemma bigo_mult [intro]: "O(f)\<otimes>O(g) <= O(f * g)"
+  apply (rule subsetI)
+  apply (subst bigo_def)
+  apply (auto simp del: abs_mult mult_ac
+              simp add: bigo_alt_def set_times_def func_times)
+(*sledgehammer*); 
+  apply (rule_tac x = "c * ca" in exI)
+  apply(rule allI)
+  apply(erule_tac x = x in allE)+
+  apply(subgoal_tac "c * ca * abs(f x * g x) = 
+      (c * abs(f x)) * (ca * abs(g x))")
+using [[ atp_problem_prefix = "BigO__bigo_mult_simpler" ]]
+prefer 2 
+apply (metis mult_assoc mult_left_commute
+  OrderedGroup.abs_of_pos OrderedGroup.mult_left_commute
+  Ring_and_Field.abs_mult Ring_and_Field.mult_pos_pos)
+  apply (erule ssubst) 
+  apply (subst abs_mult)
+(*not qute BigO__bigo_mult_simpler_1 (a hard problem!) as abs_mult has
+  just been done*)
+proof (neg_clausify)
+fix a c b ca x
+assume 0: "(0\<Colon>'b\<Colon>ordered_idom) < (c\<Colon>'b\<Colon>ordered_idom)"
+assume 1: "\<bar>(a\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) (x\<Colon>'a)\<bar>
+\<le> (c\<Colon>'b\<Colon>ordered_idom) * \<bar>(f\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>"
+assume 2: "\<bar>(b\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) (x\<Colon>'a)\<bar>
+\<le> (ca\<Colon>'b\<Colon>ordered_idom) * \<bar>(g\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>"
+assume 3: "\<not> \<bar>(a\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) (x\<Colon>'a)\<bar> *
+  \<bar>(b\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>
+  \<le> (c\<Colon>'b\<Colon>ordered_idom) * \<bar>(f\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar> *
+    ((ca\<Colon>'b\<Colon>ordered_idom) * \<bar>(g\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>)"
+have 4: "\<bar>c\<Colon>'b\<Colon>ordered_idom\<bar> = c"
+  by (metis OrderedGroup.abs_of_pos 0)
+have 5: "\<And>X1\<Colon>'b\<Colon>ordered_idom. (c\<Colon>'b\<Colon>ordered_idom) * \<bar>X1\<bar> = \<bar>c * X1\<bar>"
+  by (metis Ring_and_Field.abs_mult 4)
+have 6: "(0\<Colon>'b\<Colon>ordered_idom) = (1\<Colon>'b\<Colon>ordered_idom) \<or>
+(0\<Colon>'b\<Colon>ordered_idom) < (1\<Colon>'b\<Colon>ordered_idom)"
+  by (metis OrderedGroup.abs_not_less_zero Ring_and_Field.abs_one Ring_and_Field.linorder_neqE_ordered_idom)
+have 7: "(0\<Colon>'b\<Colon>ordered_idom) < (1\<Colon>'b\<Colon>ordered_idom)"
+  by (metis 6 Ring_and_Field.one_neq_zero)
+have 8: "\<bar>1\<Colon>'b\<Colon>ordered_idom\<bar> = (1\<Colon>'b\<Colon>ordered_idom)"
+  by (metis OrderedGroup.abs_of_pos 7)
+have 9: "\<And>X1\<Colon>'b\<Colon>ordered_idom. (0\<Colon>'b\<Colon>ordered_idom) \<le> (c\<Colon>'b\<Colon>ordered_idom) * \<bar>X1\<bar>"
+  by (metis OrderedGroup.abs_ge_zero 5)
+have 10: "\<And>X1\<Colon>'b\<Colon>ordered_idom. X1 * (1\<Colon>'b\<Colon>ordered_idom) = X1"
+  by (metis Ring_and_Field.mult_cancel_right2 mult_commute)
+have 11: "\<And>X1\<Colon>'b\<Colon>ordered_idom. \<bar>\<bar>X1\<bar>\<bar> = \<bar>X1\<bar> * \<bar>1\<Colon>'b\<Colon>ordered_idom\<bar>"
+  by (metis Ring_and_Field.abs_mult OrderedGroup.abs_idempotent 10)
+have 12: "\<And>X1\<Colon>'b\<Colon>ordered_idom. \<bar>\<bar>X1\<bar>\<bar> = \<bar>X1\<bar>"
+  by (metis 11 8 10)
+have 13: "\<And>X1\<Colon>'b\<Colon>ordered_idom. (0\<Colon>'b\<Colon>ordered_idom) \<le> \<bar>X1\<bar>"
+  by (metis OrderedGroup.abs_ge_zero 12)
+have 14: "\<not> (0\<Colon>'b\<Colon>ordered_idom)
+  \<le> (c\<Colon>'b\<Colon>ordered_idom) * \<bar>(f\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) (x\<Colon>'a)\<bar> \<or>
+\<not> (0\<Colon>'b\<Colon>ordered_idom) \<le> \<bar>(b\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar> \<or>
+\<not> \<bar>b x\<bar> \<le> (ca\<Colon>'b\<Colon>ordered_idom) * \<bar>(g\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar> \<or>
+\<not> \<bar>(a\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar> \<le> c * \<bar>f x\<bar>"
+  by (metis 3 Ring_and_Field.mult_mono)
+have 15: "\<not> (0\<Colon>'b\<Colon>ordered_idom) \<le> \<bar>(b\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) (x\<Colon>'a)\<bar> \<or>
+\<not> \<bar>b x\<bar> \<le> (ca\<Colon>'b\<Colon>ordered_idom) * \<bar>(g\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar> \<or>
+\<not> \<bar>(a\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>
+  \<le> (c\<Colon>'b\<Colon>ordered_idom) * \<bar>(f\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>"
+  by (metis 14 9)
+have 16: "\<not> \<bar>(b\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) (x\<Colon>'a)\<bar>
+  \<le> (ca\<Colon>'b\<Colon>ordered_idom) * \<bar>(g\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar> \<or>
+\<not> \<bar>(a\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>
+  \<le> (c\<Colon>'b\<Colon>ordered_idom) * \<bar>(f\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>"
+  by (metis 15 13)
+have 17: "\<not> \<bar>(a\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) (x\<Colon>'a)\<bar>
+  \<le> (c\<Colon>'b\<Colon>ordered_idom) * \<bar>(f\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>"
+  by (metis 16 2)
+show 18: "False"
+  by (metis 17 1)
+qed
+
+
+declare [[ atp_problem_prefix = "BigO__bigo_mult2" ]]
+lemma bigo_mult2 [intro]: "f *o O(g) <= O(f * g)"
+  apply (auto simp add: bigo_def elt_set_times_def func_times abs_mult)
+(*sledgehammer*); 
+  apply (rule_tac x = c in exI)
+  apply clarify
+  apply (drule_tac x = x in spec)
+using [[ atp_problem_prefix = "BigO__bigo_mult2_simpler" ]]
+(*sledgehammer [no luck]*); 
+  apply (subgoal_tac "abs(f x) * abs(b x) <= abs(f x) * (c * abs(g x))")
+  apply (simp add: mult_ac)
+  apply (rule mult_left_mono, assumption)
+  apply (rule abs_ge_zero)
+done
+
+declare [[ atp_problem_prefix = "BigO__bigo_mult3" ]]
+lemma bigo_mult3: "f : O(h) ==> g : O(j) ==> f * g : O(h * j)"
+by (metis bigo_mult set_times_intro subset_iff)
+
+declare [[ atp_problem_prefix = "BigO__bigo_mult4" ]]
+lemma bigo_mult4 [intro]:"f : k +o O(h) ==> g * f : (g * k) +o O(g * h)"
+by (metis bigo_mult2 set_plus_mono_b set_times_intro2 set_times_plus_distrib)
+
+
+lemma bigo_mult5: "ALL x. f x ~= 0 ==>
+    O(f * g) <= (f::'a => ('b::ordered_field)) *o O(g)"
+proof -
+  assume "ALL x. f x ~= 0"
+  show "O(f * g) <= f *o O(g)"
+  proof
+    fix h
+    assume "h : O(f * g)"
+    then have "(%x. 1 / (f x)) * h : (%x. 1 / f x) *o O(f * g)"
+      by auto
+    also have "... <= O((%x. 1 / f x) * (f * g))"
+      by (rule bigo_mult2)
+    also have "(%x. 1 / f x) * (f * g) = g"
+      apply (simp add: func_times) 
+      apply (rule ext)
+      apply (simp add: prems nonzero_divide_eq_eq mult_ac)
+      done
+    finally have "(%x. (1::'b) / f x) * h : O(g)".
+    then have "f * ((%x. (1::'b) / f x) * h) : f *o O(g)"
+      by auto
+    also have "f * ((%x. (1::'b) / f x) * h) = h"
+      apply (simp add: func_times) 
+      apply (rule ext)
+      apply (simp add: prems nonzero_divide_eq_eq mult_ac)
+      done
+    finally show "h : f *o O(g)".
+  qed
+qed
+
+declare [[ atp_problem_prefix = "BigO__bigo_mult6" ]]
+lemma bigo_mult6: "ALL x. f x ~= 0 ==>
+    O(f * g) = (f::'a => ('b::ordered_field)) *o O(g)"
+by (metis bigo_mult2 bigo_mult5 order_antisym)
+
+(*proof requires relaxing relevance: 2007-01-25*)
+declare [[ atp_problem_prefix = "BigO__bigo_mult7" ]]
+  declare bigo_mult6 [simp]
+lemma bigo_mult7: "ALL x. f x ~= 0 ==>
+    O(f * g) <= O(f::'a => ('b::ordered_field)) \<otimes> O(g)"
+(*sledgehammer*)
+  apply (subst bigo_mult6)
+  apply assumption
+  apply (rule set_times_mono3) 
+  apply (rule bigo_refl)
+done
+  declare bigo_mult6 [simp del]
+
+declare [[ atp_problem_prefix = "BigO__bigo_mult8" ]]
+  declare bigo_mult7[intro!]
+lemma bigo_mult8: "ALL x. f x ~= 0 ==>
+    O(f * g) = O(f::'a => ('b::ordered_field)) \<otimes> O(g)"
+by (metis bigo_mult bigo_mult7 order_antisym_conv)
+
+lemma bigo_minus [intro]: "f : O(g) ==> - f : O(g)"
+  by (auto simp add: bigo_def fun_Compl_def)
+
+lemma bigo_minus2: "f : g +o O(h) ==> -f : -g +o O(h)"
+  apply (rule set_minus_imp_plus)
+  apply (drule set_plus_imp_minus)
+  apply (drule bigo_minus)
+  apply (simp add: diff_minus)
+done
+
+lemma bigo_minus3: "O(-f) = O(f)"
+  by (auto simp add: bigo_def fun_Compl_def abs_minus_cancel)
+
+lemma bigo_plus_absorb_lemma1: "f : O(g) ==> f +o O(g) <= O(g)"
+proof -
+  assume a: "f : O(g)"
+  show "f +o O(g) <= O(g)"
+  proof -
+    have "f : O(f)" by auto
+    then have "f +o O(g) <= O(f) \<oplus> O(g)"
+      by (auto del: subsetI)
+    also have "... <= O(g) \<oplus> O(g)"
+    proof -
+      from a have "O(f) <= O(g)" by (auto del: subsetI)
+      thus ?thesis by (auto del: subsetI)
+    qed
+    also have "... <= O(g)" by (simp add: bigo_plus_idemp)
+    finally show ?thesis .
+  qed
+qed
+
+lemma bigo_plus_absorb_lemma2: "f : O(g) ==> O(g) <= f +o O(g)"
+proof -
+  assume a: "f : O(g)"
+  show "O(g) <= f +o O(g)"
+  proof -
+    from a have "-f : O(g)" by auto
+    then have "-f +o O(g) <= O(g)" by (elim bigo_plus_absorb_lemma1)
+    then have "f +o (-f +o O(g)) <= f +o O(g)" by auto
+    also have "f +o (-f +o O(g)) = O(g)"
+      by (simp add: set_plus_rearranges)
+    finally show ?thesis .
+  qed
+qed
+
+declare [[ atp_problem_prefix = "BigO__bigo_plus_absorb" ]]
+lemma bigo_plus_absorb [simp]: "f : O(g) ==> f +o O(g) = O(g)"
+by (metis bigo_plus_absorb_lemma1 bigo_plus_absorb_lemma2 order_eq_iff);
+
+lemma bigo_plus_absorb2 [intro]: "f : O(g) ==> A <= O(g) ==> f +o A <= O(g)"
+  apply (subgoal_tac "f +o A <= f +o O(g)")
+  apply force+
+done
+
+lemma bigo_add_commute_imp: "f : g +o O(h) ==> g : f +o O(h)"
+  apply (subst set_minus_plus [symmetric])
+  apply (subgoal_tac "g - f = - (f - g)")
+  apply (erule ssubst)
+  apply (rule bigo_minus)
+  apply (subst set_minus_plus)
+  apply assumption
+  apply  (simp add: diff_minus add_ac)
+done
+
+lemma bigo_add_commute: "(f : g +o O(h)) = (g : f +o O(h))"
+  apply (rule iffI)
+  apply (erule bigo_add_commute_imp)+
+done
+
+lemma bigo_const1: "(%x. c) : O(%x. 1)"
+by (auto simp add: bigo_def mult_ac)
+
+declare [[ atp_problem_prefix = "BigO__bigo_const2" ]]
+lemma (*bigo_const2 [intro]:*) "O(%x. c) <= O(%x. 1)"
+by (metis bigo_const1 bigo_elt_subset);
+
+lemma bigo_const2 [intro]: "O(%x. c::'b::ordered_idom) <= O(%x. 1)";
+(*??FAILS because the two occurrences of COMBK have different polymorphic types
+proof (neg_clausify)
+assume 0: "\<not> O(COMBK (c\<Colon>'b\<Colon>ordered_idom)) \<subseteq> O(COMBK (1\<Colon>'b\<Colon>ordered_idom))"
+have 1: "COMBK (c\<Colon>'b\<Colon>ordered_idom) \<notin> O(COMBK (1\<Colon>'b\<Colon>ordered_idom))"
+apply (rule notI) 
+apply (rule 0 [THEN notE]) 
+apply (rule bigo_elt_subset) 
+apply assumption; 
+sorry
+  by (metis 0 bigo_elt_subset)  loops??
+show "False"
+  by (metis 1 bigo_const1)
+qed
+*)
+  apply (rule bigo_elt_subset)
+  apply (rule bigo_const1)
+done
+
+declare [[ atp_problem_prefix = "BigO__bigo_const3" ]]
+lemma bigo_const3: "(c::'a::ordered_field) ~= 0 ==> (%x. 1) : O(%x. c)"
+apply (simp add: bigo_def)
+proof (neg_clausify)
+assume 0: "(c\<Colon>'a\<Colon>ordered_field) \<noteq> (0\<Colon>'a\<Colon>ordered_field)"
+assume 1: "\<And>A\<Colon>'a\<Colon>ordered_field. \<not> (1\<Colon>'a\<Colon>ordered_field) \<le> A * \<bar>c\<Colon>'a\<Colon>ordered_field\<bar>"
+have 2: "(0\<Colon>'a\<Colon>ordered_field) = \<bar>c\<Colon>'a\<Colon>ordered_field\<bar> \<or>
+\<not> (1\<Colon>'a\<Colon>ordered_field) \<le> (1\<Colon>'a\<Colon>ordered_field)"
+  by (metis 1 field_inverse)
+have 3: "\<bar>c\<Colon>'a\<Colon>ordered_field\<bar> = (0\<Colon>'a\<Colon>ordered_field)"
+  by (metis linorder_neq_iff linorder_antisym_conv1 2)
+have 4: "(0\<Colon>'a\<Colon>ordered_field) = (c\<Colon>'a\<Colon>ordered_field)"
+  by (metis 3 abs_eq_0)
+show "False"
+  by (metis 0 4)
+qed
+
+lemma bigo_const4: "(c::'a::ordered_field) ~= 0 ==> O(%x. 1) <= O(%x. c)"
+by (rule bigo_elt_subset, rule bigo_const3, assumption)
+
+lemma bigo_const [simp]: "(c::'a::ordered_field) ~= 0 ==> 
+    O(%x. c) = O(%x. 1)"
+by (rule equalityI, rule bigo_const2, rule bigo_const4, assumption)
+
+declare [[ atp_problem_prefix = "BigO__bigo_const_mult1" ]]
+lemma bigo_const_mult1: "(%x. c * f x) : O(f)"
+  apply (simp add: bigo_def abs_mult)
+proof (neg_clausify)
+fix x
+assume 0: "\<And>xa\<Colon>'b\<Colon>ordered_idom.
+   \<not> \<bar>c\<Colon>'b\<Colon>ordered_idom\<bar> *
+     \<bar>(f\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) ((x\<Colon>'b\<Colon>ordered_idom \<Rightarrow> 'a\<Colon>type) xa)\<bar>
+     \<le> xa * \<bar>f (x xa)\<bar>"
+show "False"
+  by (metis linorder_neq_iff linorder_antisym_conv1 0)
+qed
+
+lemma bigo_const_mult2: "O(%x. c * f x) <= O(f)"
+by (rule bigo_elt_subset, rule bigo_const_mult1)
+
+declare [[ atp_problem_prefix = "BigO__bigo_const_mult3" ]]
+lemma bigo_const_mult3: "(c::'a::ordered_field) ~= 0 ==> f : O(%x. c * f x)"
+  apply (simp add: bigo_def)
+(*sledgehammer [no luck]*); 
+  apply (rule_tac x = "abs(inverse c)" in exI)
+  apply (simp only: abs_mult [symmetric] mult_assoc [symmetric])
+apply (subst left_inverse) 
+apply (auto ); 
+done
+
+lemma bigo_const_mult4: "(c::'a::ordered_field) ~= 0 ==> 
+    O(f) <= O(%x. c * f x)"
+by (rule bigo_elt_subset, rule bigo_const_mult3, assumption)
+
+lemma bigo_const_mult [simp]: "(c::'a::ordered_field) ~= 0 ==> 
+    O(%x. c * f x) = O(f)"
+by (rule equalityI, rule bigo_const_mult2, erule bigo_const_mult4)
+
+declare [[ atp_problem_prefix = "BigO__bigo_const_mult5" ]]
+lemma bigo_const_mult5 [simp]: "(c::'a::ordered_field) ~= 0 ==> 
+    (%x. c) *o O(f) = O(f)"
+  apply (auto del: subsetI)
+  apply (rule order_trans)
+  apply (rule bigo_mult2)
+  apply (simp add: func_times)
+  apply (auto intro!: subsetI simp add: bigo_def elt_set_times_def func_times)
+  apply (rule_tac x = "%y. inverse c * x y" in exI)
+  apply (rename_tac g d) 
+  apply safe
+  apply (rule_tac [2] ext) 
+   prefer 2 
+   apply simp
+  apply (simp add: mult_assoc [symmetric] abs_mult)
+  (*couldn't get this proof without the step above; SLOW*)
+  apply (metis mult_assoc abs_ge_zero mult_left_mono)
+done
+
+
+declare [[ atp_problem_prefix = "BigO__bigo_const_mult6" ]]
+lemma bigo_const_mult6 [intro]: "(%x. c) *o O(f) <= O(f)"
+  apply (auto intro!: subsetI
+    simp add: bigo_def elt_set_times_def func_times
+    simp del: abs_mult mult_ac)
+(*sledgehammer*); 
+  apply (rule_tac x = "ca * (abs c)" in exI)
+  apply (rule allI)
+  apply (subgoal_tac "ca * abs(c) * abs(f x) = abs(c) * (ca * abs(f x))")
+  apply (erule ssubst)
+  apply (subst abs_mult)
+  apply (rule mult_left_mono)
+  apply (erule spec)
+  apply simp
+  apply(simp add: mult_ac)
+done
+
+lemma bigo_const_mult7 [intro]: "f =o O(g) ==> (%x. c * f x) =o O(g)"
+proof -
+  assume "f =o O(g)"
+  then have "(%x. c) * f =o (%x. c) *o O(g)"
+    by auto
+  also have "(%x. c) * f = (%x. c * f x)"
+    by (simp add: func_times)
+  also have "(%x. c) *o O(g) <= O(g)"
+    by (auto del: subsetI)
+  finally show ?thesis .
+qed
+
+lemma bigo_compose1: "f =o O(g) ==> (%x. f(k x)) =o O(%x. g(k x))"
+by (unfold bigo_def, auto)
+
+lemma bigo_compose2: "f =o g +o O(h) ==> (%x. f(k x)) =o (%x. g(k x)) +o 
+    O(%x. h(k x))"
+  apply (simp only: set_minus_plus [symmetric] diff_minus fun_Compl_def
+      func_plus)
+  apply (erule bigo_compose1)
+done
+
+subsection {* Setsum *}
+
+lemma bigo_setsum_main: "ALL x. ALL y : A x. 0 <= h x y ==> 
+    EX c. ALL x. ALL y : A x. abs(f x y) <= c * (h x y) ==>
+      (%x. SUM y : A x. f x y) =o O(%x. SUM y : A x. h x y)"  
+  apply (auto simp add: bigo_def)
+  apply (rule_tac x = "abs c" in exI)
+  apply (subst abs_of_nonneg) back back
+  apply (rule setsum_nonneg)
+  apply force
+  apply (subst setsum_right_distrib)
+  apply (rule allI)
+  apply (rule order_trans)
+  apply (rule setsum_abs)
+  apply (rule setsum_mono)
+apply (blast intro: order_trans mult_right_mono abs_ge_self) 
+done
+
+declare [[ atp_problem_prefix = "BigO__bigo_setsum1" ]]
+lemma bigo_setsum1: "ALL x y. 0 <= h x y ==> 
+    EX c. ALL x y. abs(f x y) <= c * (h x y) ==>
+      (%x. SUM y : A x. f x y) =o O(%x. SUM y : A x. h x y)"
+  apply (rule bigo_setsum_main)
+(*sledgehammer*); 
+  apply force
+  apply clarsimp
+  apply (rule_tac x = c in exI)
+  apply force
+done
+
+lemma bigo_setsum2: "ALL y. 0 <= h y ==> 
+    EX c. ALL y. abs(f y) <= c * (h y) ==>
+      (%x. SUM y : A x. f y) =o O(%x. SUM y : A x. h y)"
+by (rule bigo_setsum1, auto)  
+
+declare [[ atp_problem_prefix = "BigO__bigo_setsum3" ]]
+lemma bigo_setsum3: "f =o O(h) ==>
+    (%x. SUM y : A x. (l x y) * f(k x y)) =o
+      O(%x. SUM y : A x. abs(l x y * h(k x y)))"
+  apply (rule bigo_setsum1)
+  apply (rule allI)+
+  apply (rule abs_ge_zero)
+  apply (unfold bigo_def)
+  apply (auto simp add: abs_mult);
+(*sledgehammer*); 
+  apply (rule_tac x = c in exI)
+  apply (rule allI)+
+  apply (subst mult_left_commute)
+  apply (rule mult_left_mono)
+  apply (erule spec)
+  apply (rule abs_ge_zero)
+done
+
+lemma bigo_setsum4: "f =o g +o O(h) ==>
+    (%x. SUM y : A x. l x y * f(k x y)) =o
+      (%x. SUM y : A x. l x y * g(k x y)) +o
+        O(%x. SUM y : A x. abs(l x y * h(k x y)))"
+  apply (rule set_minus_imp_plus)
+  apply (subst fun_diff_def)
+  apply (subst setsum_subtractf [symmetric])
+  apply (subst right_diff_distrib [symmetric])
+  apply (rule bigo_setsum3)
+  apply (subst fun_diff_def [symmetric])
+  apply (erule set_plus_imp_minus)
+done
+
+declare [[ atp_problem_prefix = "BigO__bigo_setsum5" ]]
+lemma bigo_setsum5: "f =o O(h) ==> ALL x y. 0 <= l x y ==> 
+    ALL x. 0 <= h x ==>
+      (%x. SUM y : A x. (l x y) * f(k x y)) =o
+        O(%x. SUM y : A x. (l x y) * h(k x y))" 
+  apply (subgoal_tac "(%x. SUM y : A x. (l x y) * h(k x y)) = 
+      (%x. SUM y : A x. abs((l x y) * h(k x y)))")
+  apply (erule ssubst)
+  apply (erule bigo_setsum3)
+  apply (rule ext)
+  apply (rule setsum_cong2)
+  apply (thin_tac "f \<in> O(h)") 
+apply (metis abs_of_nonneg zero_le_mult_iff)
+done
+
+lemma bigo_setsum6: "f =o g +o O(h) ==> ALL x y. 0 <= l x y ==>
+    ALL x. 0 <= h x ==>
+      (%x. SUM y : A x. (l x y) * f(k x y)) =o
+        (%x. SUM y : A x. (l x y) * g(k x y)) +o
+          O(%x. SUM y : A x. (l x y) * h(k x y))" 
+  apply (rule set_minus_imp_plus)
+  apply (subst fun_diff_def)
+  apply (subst setsum_subtractf [symmetric])
+  apply (subst right_diff_distrib [symmetric])
+  apply (rule bigo_setsum5)
+  apply (subst fun_diff_def [symmetric])
+  apply (drule set_plus_imp_minus)
+  apply auto
+done
+
+subsection {* Misc useful stuff *}
+
+lemma bigo_useful_intro: "A <= O(f) ==> B <= O(f) ==>
+  A \<oplus> B <= O(f)"
+  apply (subst bigo_plus_idemp [symmetric])
+  apply (rule set_plus_mono2)
+  apply assumption+
+done
+
+lemma bigo_useful_add: "f =o O(h) ==> g =o O(h) ==> f + g =o O(h)"
+  apply (subst bigo_plus_idemp [symmetric])
+  apply (rule set_plus_intro)
+  apply assumption+
+done
+  
+lemma bigo_useful_const_mult: "(c::'a::ordered_field) ~= 0 ==> 
+    (%x. c) * f =o O(h) ==> f =o O(h)"
+  apply (rule subsetD)
+  apply (subgoal_tac "(%x. 1 / c) *o O(h) <= O(h)")
+  apply assumption
+  apply (rule bigo_const_mult6)
+  apply (subgoal_tac "f = (%x. 1 / c) * ((%x. c) * f)")
+  apply (erule ssubst)
+  apply (erule set_times_intro2)
+  apply (simp add: func_times) 
+done
+
+declare [[ atp_problem_prefix = "BigO__bigo_fix" ]]
+lemma bigo_fix: "(%x. f ((x::nat) + 1)) =o O(%x. h(x + 1)) ==> f 0 = 0 ==>
+    f =o O(h)"
+  apply (simp add: bigo_alt_def)
+(*sledgehammer*); 
+  apply clarify
+  apply (rule_tac x = c in exI)
+  apply safe
+  apply (case_tac "x = 0")
+apply (metis OrderedGroup.abs_ge_zero  OrderedGroup.abs_zero  order_less_le  Ring_and_Field.split_mult_pos_le) 
+  apply (subgoal_tac "x = Suc (x - 1)")
+  apply metis
+  apply simp
+  done
+
+
+lemma bigo_fix2: 
+    "(%x. f ((x::nat) + 1)) =o (%x. g(x + 1)) +o O(%x. h(x + 1)) ==> 
+       f 0 = g 0 ==> f =o g +o O(h)"
+  apply (rule set_minus_imp_plus)
+  apply (rule bigo_fix)
+  apply (subst fun_diff_def)
+  apply (subst fun_diff_def [symmetric])
+  apply (rule set_plus_imp_minus)
+  apply simp
+  apply (simp add: fun_diff_def)
+done
+
+subsection {* Less than or equal to *}
+
+constdefs 
+  lesso :: "('a => 'b::ordered_idom) => ('a => 'b) => ('a => 'b)"
+      (infixl "<o" 70)
+  "f <o g == (%x. max (f x - g x) 0)"
+
+lemma bigo_lesseq1: "f =o O(h) ==> ALL x. abs (g x) <= abs (f x) ==>
+    g =o O(h)"
+  apply (unfold bigo_def)
+  apply clarsimp
+apply (blast intro: order_trans) 
+done
+
+lemma bigo_lesseq2: "f =o O(h) ==> ALL x. abs (g x) <= f x ==>
+      g =o O(h)"
+  apply (erule bigo_lesseq1)
+apply (blast intro: abs_ge_self order_trans) 
+done
+
+lemma bigo_lesseq3: "f =o O(h) ==> ALL x. 0 <= g x ==> ALL x. g x <= f x ==>
+      g =o O(h)"
+  apply (erule bigo_lesseq2)
+  apply (rule allI)
+  apply (subst abs_of_nonneg)
+  apply (erule spec)+
+done
+
+lemma bigo_lesseq4: "f =o O(h) ==>
+    ALL x. 0 <= g x ==> ALL x. g x <= abs (f x) ==>
+      g =o O(h)"
+  apply (erule bigo_lesseq1)
+  apply (rule allI)
+  apply (subst abs_of_nonneg)
+  apply (erule spec)+
+done
+
+declare [[ atp_problem_prefix = "BigO__bigo_lesso1" ]]
+lemma bigo_lesso1: "ALL x. f x <= g x ==> f <o g =o O(h)"
+  apply (unfold lesso_def)
+  apply (subgoal_tac "(%x. max (f x - g x) 0) = 0")
+(*??Translation of TSTP raised an exception: Type unification failed: Variable ?'X2.0::type not of sort ord*)
+apply (metis bigo_zero)
+  apply (unfold func_zero)
+  apply (rule ext)
+  apply (simp split: split_max)
+done
+
+
+declare [[ atp_problem_prefix = "BigO__bigo_lesso2" ]]
+lemma bigo_lesso2: "f =o g +o O(h) ==>
+    ALL x. 0 <= k x ==> ALL x. k x <= f x ==>
+      k <o g =o O(h)"
+  apply (unfold lesso_def)
+  apply (rule bigo_lesseq4)
+  apply (erule set_plus_imp_minus)
+  apply (rule allI)
+  apply (rule le_maxI2)
+  apply (rule allI)
+  apply (subst fun_diff_def)
+apply (erule thin_rl)
+(*sledgehammer*);  
+  apply (case_tac "0 <= k x - g x")
+  prefer 2 (*re-order subgoals because I don't know what to put after a structured proof*)
+   apply (metis abs_ge_zero abs_minus_commute linorder_linear min_max.sup_absorb1 min_max.sup_commute)
+proof (neg_clausify)
+fix x
+assume 0: "\<And>A. k A \<le> f A"
+have 1: "\<And>(X1\<Colon>'b\<Colon>ordered_idom) X2. \<not> max X1 X2 < X1"
+  by (metis linorder_not_less le_maxI1)  (*sort inserted by hand*)
+assume 2: "(0\<Colon>'b) \<le> k x - g x"
+have 3: "\<not> k x - g x < (0\<Colon>'b)"
+  by (metis 2 linorder_not_less)
+have 4: "\<And>X1 X2. min X1 (k X2) \<le> f X2"
+  by (metis min_max.inf_le2 min_max.le_inf_iff min_max.le_iff_inf 0)
+have 5: "\<bar>g x - f x\<bar> = f x - g x"
+  by (metis abs_minus_commute combine_common_factor mult_zero_right minus_add_cancel minus_zero abs_if diff_less_eq min_max.inf_commute 4 linorder_not_le min_max.le_iff_inf 3 diff_less_0_iff_less linorder_not_less)
+have 6: "max (0\<Colon>'b) (k x - g x) = k x - g x"
+  by (metis min_max.le_iff_sup 2)
+assume 7: "\<not> max (k x - g x) (0\<Colon>'b) \<le> \<bar>f x - g x\<bar>"
+have 8: "\<not> k x - g x \<le> f x - g x"
+  by (metis 5 abs_minus_commute 7 min_max.sup_commute 6)
+show "False"
+  by (metis min_max.sup_commute min_max.inf_commute min_max.sup_inf_absorb min_max.le_iff_inf 0 max_diff_distrib_left 1 linorder_not_le 8)
+qed
+
+declare [[ atp_problem_prefix = "BigO__bigo_lesso3" ]]
+lemma bigo_lesso3: "f =o g +o O(h) ==>
+    ALL x. 0 <= k x ==> ALL x. g x <= k x ==>
+      f <o k =o O(h)"
+  apply (unfold lesso_def)
+  apply (rule bigo_lesseq4)
+  apply (erule set_plus_imp_minus)
+  apply (rule allI)
+  apply (rule le_maxI2)
+  apply (rule allI)
+  apply (subst fun_diff_def)
+apply (erule thin_rl) 
+(*sledgehammer*); 
+  apply (case_tac "0 <= f x - k x")
+  apply (simp)
+  apply (subst abs_of_nonneg)
+  apply (drule_tac x = x in spec) back
+using [[ atp_problem_prefix = "BigO__bigo_lesso3_simpler" ]]
+apply (metis diff_less_0_iff_less linorder_not_le not_leE uminus_add_conv_diff xt1(12) xt1(6))
+apply (metis add_minus_cancel diff_le_eq le_diff_eq uminus_add_conv_diff)
+apply (metis abs_ge_zero linorder_linear min_max.sup_absorb1 min_max.sup_commute)
+done
+
+lemma bigo_lesso4: "f <o g =o O(k::'a=>'b::ordered_field) ==>
+    g =o h +o O(k) ==> f <o h =o O(k)"
+  apply (unfold lesso_def)
+  apply (drule set_plus_imp_minus)
+  apply (drule bigo_abs5) back
+  apply (simp add: fun_diff_def)
+  apply (drule bigo_useful_add)
+  apply assumption
+  apply (erule bigo_lesseq2) back
+  apply (rule allI)
+  apply (auto simp add: func_plus fun_diff_def algebra_simps
+    split: split_max abs_split)
+done
+
+declare [[ atp_problem_prefix = "BigO__bigo_lesso5" ]]
+lemma bigo_lesso5: "f <o g =o O(h) ==>
+    EX C. ALL x. f x <= g x + C * abs(h x)"
+  apply (simp only: lesso_def bigo_alt_def)
+  apply clarsimp
+  apply (metis abs_if abs_mult add_commute diff_le_eq less_not_permute)  
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Metis_Examples/Message.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,812 @@
+(*  Title:      HOL/MetisTest/Message.thy
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+
+Testing the metis method.
+*)
+
+theory Message imports Main begin
+
+(*Needed occasionally with spy_analz_tac, e.g. in analz_insert_Key_newK*)
+lemma strange_Un_eq [simp]: "A \<union> (B \<union> A) = B \<union> A"
+by blast
+
+types 
+  key = nat
+
+consts
+  all_symmetric :: bool        --{*true if all keys are symmetric*}
+  invKey        :: "key=>key"  --{*inverse of a symmetric key*}
+
+specification (invKey)
+  invKey [simp]: "invKey (invKey K) = K"
+  invKey_symmetric: "all_symmetric --> invKey = id"
+    by (rule exI [of _ id], auto)
+
+
+text{*The inverse of a symmetric key is itself; that of a public key
+      is the private key and vice versa*}
+
+constdefs
+  symKeys :: "key set"
+  "symKeys == {K. invKey K = K}"
+
+datatype  --{*We allow any number of friendly agents*}
+  agent = Server | Friend nat | Spy
+
+datatype
+     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*}
+
+
+text{*Concrete syntax: messages appear as {|A,B,NA|}, etc...*}
+syntax
+  "@MTuple"      :: "['a, args] => 'a * 'b"       ("(2{|_,/ _|})")
+
+syntax (xsymbols)
+  "@MTuple"      :: "['a, args] => 'a * 'b"       ("(2\<lbrace>_,/ _\<rbrace>)")
+
+translations
+  "{|x, y, z|}"   == "{|x, {|y, z|}|}"
+  "{|x, y|}"      == "MPair x y"
+
+
+constdefs
+  HPair :: "[msg,msg] => msg"                       ("(4Hash[_] /_)" [0, 1000])
+    --{*Message Y paired with a MAC computed with the help of X*}
+    "Hash[X] Y == {| Hash{|X,Y|}, Y|}"
+
+  keysFor :: "msg set => key set"
+    --{*Keys useful to decrypt elements of a message set*}
+  "keysFor H == invKey ` {K. \<exists>X. Crypt K X \<in> H}"
+
+
+subsubsection{*Inductive Definition of All Parts" of a Message*}
+
+inductive_set
+  parts :: "msg set => msg set"
+  for H :: "msg set"
+  where
+    Inj [intro]:               "X \<in> H ==> X \<in> parts H"
+  | Fst:         "{|X,Y|}   \<in> parts H ==> X \<in> parts H"
+  | Snd:         "{|X,Y|}   \<in> parts H ==> Y \<in> parts H"
+  | Body:        "Crypt K X \<in> parts H ==> X \<in> parts H"
+
+
+declare [[ atp_problem_prefix = "Message__parts_mono" ]]
+lemma parts_mono: "G \<subseteq> H ==> parts(G) \<subseteq> parts(H)"
+apply auto
+apply (erule parts.induct) 
+apply (metis Inj set_mp)
+apply (metis Fst)
+apply (metis Snd)
+apply (metis Body)
+done
+
+
+text{*Equations hold because constructors are injective.*}
+lemma Friend_image_eq [simp]: "(Friend x \<in> Friend`A) = (x:A)"
+by auto
+
+lemma Key_image_eq [simp]: "(Key x \<in> Key`A) = (x\<in>A)"
+by auto
+
+lemma Nonce_Key_image_eq [simp]: "(Nonce x \<notin> Key`A)"
+by auto
+
+
+subsubsection{*Inverse of keys *}
+
+declare [[ atp_problem_prefix = "Message__invKey_eq" ]]
+lemma invKey_eq [simp]: "(invKey K = invKey K') = (K=K')"
+by (metis invKey)
+
+
+subsection{*keysFor operator*}
+
+lemma keysFor_empty [simp]: "keysFor {} = {}"
+by (unfold keysFor_def, blast)
+
+lemma keysFor_Un [simp]: "keysFor (H \<union> H') = keysFor H \<union> keysFor H'"
+by (unfold keysFor_def, blast)
+
+lemma keysFor_UN [simp]: "keysFor (\<Union>i\<in>A. H i) = (\<Union>i\<in>A. keysFor (H i))"
+by (unfold keysFor_def, blast)
+
+text{*Monotonicity*}
+lemma keysFor_mono: "G \<subseteq> H ==> keysFor(G) \<subseteq> keysFor(H)"
+by (unfold keysFor_def, blast)
+
+lemma keysFor_insert_Agent [simp]: "keysFor (insert (Agent A) H) = keysFor H"
+by (unfold keysFor_def, auto)
+
+lemma keysFor_insert_Nonce [simp]: "keysFor (insert (Nonce N) H) = keysFor H"
+by (unfold keysFor_def, auto)
+
+lemma keysFor_insert_Number [simp]: "keysFor (insert (Number N) H) = keysFor H"
+by (unfold keysFor_def, auto)
+
+lemma keysFor_insert_Key [simp]: "keysFor (insert (Key K) H) = keysFor H"
+by (unfold keysFor_def, auto)
+
+lemma keysFor_insert_Hash [simp]: "keysFor (insert (Hash X) H) = keysFor H"
+by (unfold keysFor_def, auto)
+
+lemma keysFor_insert_MPair [simp]: "keysFor (insert {|X,Y|} H) = keysFor H"
+by (unfold keysFor_def, auto)
+
+lemma keysFor_insert_Crypt [simp]: 
+    "keysFor (insert (Crypt K X) H) = insert (invKey K) (keysFor H)"
+by (unfold keysFor_def, auto)
+
+lemma keysFor_image_Key [simp]: "keysFor (Key`E) = {}"
+by (unfold keysFor_def, auto)
+
+lemma Crypt_imp_invKey_keysFor: "Crypt K X \<in> H ==> invKey K \<in> keysFor H"
+by (unfold keysFor_def, blast)
+
+
+subsection{*Inductive relation "parts"*}
+
+lemma MPair_parts:
+     "[| {|X,Y|} \<in> parts H;        
+         [| X \<in> parts H; Y \<in> parts H |] ==> P |] ==> P"
+by (blast dest: parts.Fst parts.Snd) 
+
+    declare MPair_parts [elim!]  parts.Body [dest!]
+text{*NB These two rules are UNSAFE in the formal sense, as they discard the
+     compound message.  They work well on THIS FILE.  
+  @{text MPair_parts} is left as SAFE because it speeds up proofs.
+  The Crypt rule is normally kept UNSAFE to avoid breaking up certificates.*}
+
+lemma parts_increasing: "H \<subseteq> parts(H)"
+by blast
+
+lemmas parts_insertI = subset_insertI [THEN parts_mono, THEN subsetD, standard]
+
+lemma parts_empty [simp]: "parts{} = {}"
+apply safe
+apply (erule parts.induct)
+apply blast+
+done
+
+lemma parts_emptyE [elim!]: "X\<in> parts{} ==> P"
+by simp
+
+text{*WARNING: loops if H = {Y}, therefore must not be repeated!*}
+lemma parts_singleton: "X\<in> parts H ==> \<exists>Y\<in>H. X\<in> parts {Y}"
+apply (erule parts.induct)
+apply fast+
+done
+
+
+subsubsection{*Unions *}
+
+lemma parts_Un_subset1: "parts(G) \<union> parts(H) \<subseteq> parts(G \<union> H)"
+by (intro Un_least parts_mono Un_upper1 Un_upper2)
+
+lemma parts_Un_subset2: "parts(G \<union> H) \<subseteq> parts(G) \<union> parts(H)"
+apply (rule subsetI)
+apply (erule parts.induct, blast+)
+done
+
+lemma parts_Un [simp]: "parts(G \<union> H) = parts(G) \<union> parts(H)"
+by (intro equalityI parts_Un_subset1 parts_Un_subset2)
+
+lemma parts_insert: "parts (insert X H) = parts {X} \<union> parts H"
+apply (subst insert_is_Un [of _ H])
+apply (simp only: parts_Un)
+done
+
+declare [[ atp_problem_prefix = "Message__parts_insert_two" ]]
+lemma parts_insert2:
+     "parts (insert X (insert Y H)) = parts {X} \<union> parts {Y} \<union> parts H"
+by (metis Un_commute Un_empty_left Un_empty_right Un_insert_left Un_insert_right parts_Un)
+
+
+lemma parts_UN_subset1: "(\<Union>x\<in>A. parts(H x)) \<subseteq> parts(\<Union>x\<in>A. H x)"
+by (intro UN_least parts_mono UN_upper)
+
+lemma parts_UN_subset2: "parts(\<Union>x\<in>A. H x) \<subseteq> (\<Union>x\<in>A. parts(H x))"
+apply (rule subsetI)
+apply (erule parts.induct, blast+)
+done
+
+lemma parts_UN [simp]: "parts(\<Union>x\<in>A. H x) = (\<Union>x\<in>A. parts(H x))"
+by (intro equalityI parts_UN_subset1 parts_UN_subset2)
+
+text{*Added to simplify arguments to parts, analz and synth.
+  NOTE: the UN versions are no longer used!*}
+
+
+text{*This allows @{text blast} to simplify occurrences of 
+  @{term "parts(G\<union>H)"} in the assumption.*}
+lemmas in_parts_UnE = parts_Un [THEN equalityD1, THEN subsetD, THEN UnE] 
+declare in_parts_UnE [elim!]
+
+lemma parts_insert_subset: "insert X (parts H) \<subseteq> parts(insert X H)"
+by (blast intro: parts_mono [THEN [2] rev_subsetD])
+
+subsubsection{*Idempotence and transitivity *}
+
+lemma parts_partsD [dest!]: "X\<in> parts (parts H) ==> X\<in> parts H"
+by (erule parts.induct, blast+)
+
+lemma parts_idem [simp]: "parts (parts H) = parts H"
+by blast
+
+declare [[ atp_problem_prefix = "Message__parts_subset_iff" ]]
+lemma parts_subset_iff [simp]: "(parts G \<subseteq> parts H) = (G \<subseteq> parts H)"
+apply (rule iffI) 
+apply (metis Un_absorb1 Un_subset_iff parts_Un parts_increasing)
+apply (metis parts_idem parts_mono)
+done
+
+lemma parts_trans: "[| X\<in> parts G;  G \<subseteq> parts H |] ==> X\<in> parts H"
+by (blast dest: parts_mono); 
+
+
+declare [[ atp_problem_prefix = "Message__parts_cut" ]]
+lemma parts_cut: "[|Y\<in> parts(insert X G);  X\<in> parts H|] ==> Y\<in> parts(G \<union> H)"
+by (metis Un_subset_iff insert_subset parts_increasing parts_trans) 
+
+
+
+subsubsection{*Rewrite rules for pulling out atomic messages *}
+
+lemmas parts_insert_eq_I = equalityI [OF subsetI parts_insert_subset]
+
+
+lemma parts_insert_Agent [simp]:
+     "parts (insert (Agent agt) H) = insert (Agent agt) (parts H)"
+apply (rule parts_insert_eq_I) 
+apply (erule parts.induct, auto) 
+done
+
+lemma parts_insert_Nonce [simp]:
+     "parts (insert (Nonce N) H) = insert (Nonce N) (parts H)"
+apply (rule parts_insert_eq_I) 
+apply (erule parts.induct, auto) 
+done
+
+lemma parts_insert_Number [simp]:
+     "parts (insert (Number N) H) = insert (Number N) (parts H)"
+apply (rule parts_insert_eq_I) 
+apply (erule parts.induct, auto) 
+done
+
+lemma parts_insert_Key [simp]:
+     "parts (insert (Key K) H) = insert (Key K) (parts H)"
+apply (rule parts_insert_eq_I) 
+apply (erule parts.induct, auto) 
+done
+
+lemma parts_insert_Hash [simp]:
+     "parts (insert (Hash X) H) = insert (Hash X) (parts H)"
+apply (rule parts_insert_eq_I) 
+apply (erule parts.induct, auto) 
+done
+
+lemma parts_insert_Crypt [simp]:
+     "parts (insert (Crypt K X) H) =  
+          insert (Crypt K X) (parts (insert X H))"
+apply (rule equalityI)
+apply (rule subsetI)
+apply (erule parts.induct, auto)
+apply (blast intro: parts.Body)
+done
+
+lemma parts_insert_MPair [simp]:
+     "parts (insert {|X,Y|} H) =  
+          insert {|X,Y|} (parts (insert X (insert Y H)))"
+apply (rule equalityI)
+apply (rule subsetI)
+apply (erule parts.induct, auto)
+apply (blast intro: parts.Fst parts.Snd)+
+done
+
+lemma parts_image_Key [simp]: "parts (Key`N) = Key`N"
+apply auto
+apply (erule parts.induct, auto)
+done
+
+
+declare [[ atp_problem_prefix = "Message__msg_Nonce_supply" ]]
+lemma msg_Nonce_supply: "\<exists>N. \<forall>n. N\<le>n --> Nonce n \<notin> parts {msg}"
+apply (induct_tac "msg") 
+apply (simp_all add: parts_insert2)
+apply (metis Suc_n_not_le_n)
+apply (metis le_trans linorder_linear)
+done
+
+subsection{*Inductive relation "analz"*}
+
+text{*Inductive definition of "analz" -- what can be broken down from a set of
+    messages, including keys.  A form of downward closure.  Pairs can
+    be taken apart; messages decrypted with known keys.  *}
+
+inductive_set
+  analz :: "msg set => msg set"
+  for H :: "msg set"
+  where
+    Inj [intro,simp] :    "X \<in> H ==> X \<in> analz H"
+  | Fst:     "{|X,Y|} \<in> analz H ==> X \<in> analz H"
+  | Snd:     "{|X,Y|} \<in> analz H ==> Y \<in> analz H"
+  | Decrypt [dest]: 
+             "[|Crypt K X \<in> analz H; Key(invKey K): analz H|] ==> X \<in> analz H"
+
+
+text{*Monotonicity; Lemma 1 of Lowe's paper*}
+lemma analz_mono: "G\<subseteq>H ==> analz(G) \<subseteq> analz(H)"
+apply auto
+apply (erule analz.induct) 
+apply (auto dest: analz.Fst analz.Snd) 
+done
+
+text{*Making it safe speeds up proofs*}
+lemma MPair_analz [elim!]:
+     "[| {|X,Y|} \<in> analz H;        
+             [| X \<in> analz H; Y \<in> analz H |] ==> P   
+          |] ==> P"
+by (blast dest: analz.Fst analz.Snd)
+
+lemma analz_increasing: "H \<subseteq> analz(H)"
+by blast
+
+lemma analz_subset_parts: "analz H \<subseteq> parts H"
+apply (rule subsetI)
+apply (erule analz.induct, blast+)
+done
+
+lemmas analz_into_parts = analz_subset_parts [THEN subsetD, standard]
+
+lemmas not_parts_not_analz = analz_subset_parts [THEN contra_subsetD, standard]
+
+
+declare [[ atp_problem_prefix = "Message__parts_analz" ]]
+lemma parts_analz [simp]: "parts (analz H) = parts H"
+apply (rule equalityI)
+apply (metis analz_subset_parts parts_subset_iff)
+apply (metis analz_increasing parts_mono)
+done
+
+
+lemma analz_parts [simp]: "analz (parts H) = parts H"
+apply auto
+apply (erule analz.induct, auto)
+done
+
+lemmas analz_insertI = subset_insertI [THEN analz_mono, THEN [2] rev_subsetD, standard]
+
+subsubsection{*General equational properties *}
+
+lemma analz_empty [simp]: "analz{} = {}"
+apply safe
+apply (erule analz.induct, blast+)
+done
+
+text{*Converse fails: we can analz more from the union than from the 
+  separate parts, as a key in one might decrypt a message in the other*}
+lemma analz_Un: "analz(G) \<union> analz(H) \<subseteq> analz(G \<union> H)"
+by (intro Un_least analz_mono Un_upper1 Un_upper2)
+
+lemma analz_insert: "insert X (analz H) \<subseteq> analz(insert X H)"
+by (blast intro: analz_mono [THEN [2] rev_subsetD])
+
+subsubsection{*Rewrite rules for pulling out atomic messages *}
+
+lemmas analz_insert_eq_I = equalityI [OF subsetI analz_insert]
+
+lemma analz_insert_Agent [simp]:
+     "analz (insert (Agent agt) H) = insert (Agent agt) (analz H)"
+apply (rule analz_insert_eq_I) 
+apply (erule analz.induct, auto) 
+done
+
+lemma analz_insert_Nonce [simp]:
+     "analz (insert (Nonce N) H) = insert (Nonce N) (analz H)"
+apply (rule analz_insert_eq_I) 
+apply (erule analz.induct, auto) 
+done
+
+lemma analz_insert_Number [simp]:
+     "analz (insert (Number N) H) = insert (Number N) (analz H)"
+apply (rule analz_insert_eq_I) 
+apply (erule analz.induct, auto) 
+done
+
+lemma analz_insert_Hash [simp]:
+     "analz (insert (Hash X) H) = insert (Hash X) (analz H)"
+apply (rule analz_insert_eq_I) 
+apply (erule analz.induct, auto) 
+done
+
+text{*Can only pull out Keys if they are not needed to decrypt the rest*}
+lemma analz_insert_Key [simp]: 
+    "K \<notin> keysFor (analz H) ==>   
+          analz (insert (Key K) H) = insert (Key K) (analz H)"
+apply (unfold keysFor_def)
+apply (rule analz_insert_eq_I) 
+apply (erule analz.induct, auto) 
+done
+
+lemma analz_insert_MPair [simp]:
+     "analz (insert {|X,Y|} H) =  
+          insert {|X,Y|} (analz (insert X (insert Y H)))"
+apply (rule equalityI)
+apply (rule subsetI)
+apply (erule analz.induct, auto)
+apply (erule analz.induct)
+apply (blast intro: analz.Fst analz.Snd)+
+done
+
+text{*Can pull out enCrypted message if the Key is not known*}
+lemma analz_insert_Crypt:
+     "Key (invKey K) \<notin> analz H 
+      ==> analz (insert (Crypt K X) H) = insert (Crypt K X) (analz H)"
+apply (rule analz_insert_eq_I) 
+apply (erule analz.induct, auto) 
+
+done
+
+lemma lemma1: "Key (invKey K) \<in> analz H ==>   
+               analz (insert (Crypt K X) H) \<subseteq>  
+               insert (Crypt K X) (analz (insert X H))" 
+apply (rule subsetI)
+apply (erule_tac x = x in analz.induct, auto)
+done
+
+lemma lemma2: "Key (invKey K) \<in> analz H ==>   
+               insert (Crypt K X) (analz (insert X H)) \<subseteq>  
+               analz (insert (Crypt K X) H)"
+apply auto
+apply (erule_tac x = x in analz.induct, auto)
+apply (blast intro: analz_insertI analz.Decrypt)
+done
+
+lemma analz_insert_Decrypt:
+     "Key (invKey K) \<in> analz H ==>   
+               analz (insert (Crypt K X) H) =  
+               insert (Crypt K X) (analz (insert X H))"
+by (intro equalityI lemma1 lemma2)
+
+text{*Case analysis: either the message is secure, or it is not! Effective,
+but can cause subgoals to blow up! Use with @{text "split_if"}; apparently
+@{text "split_tac"} does not cope with patterns such as @{term"analz (insert
+(Crypt K X) H)"} *} 
+lemma analz_Crypt_if [simp]:
+     "analz (insert (Crypt K X) H) =                 
+          (if (Key (invKey K) \<in> analz H)                 
+           then insert (Crypt K X) (analz (insert X H))  
+           else insert (Crypt K X) (analz H))"
+by (simp add: analz_insert_Crypt analz_insert_Decrypt)
+
+
+text{*This rule supposes "for the sake of argument" that we have the key.*}
+lemma analz_insert_Crypt_subset:
+     "analz (insert (Crypt K X) H) \<subseteq>   
+           insert (Crypt K X) (analz (insert X H))"
+apply (rule subsetI)
+apply (erule analz.induct, auto)
+done
+
+
+lemma analz_image_Key [simp]: "analz (Key`N) = Key`N"
+apply auto
+apply (erule analz.induct, auto)
+done
+
+
+subsubsection{*Idempotence and transitivity *}
+
+lemma analz_analzD [dest!]: "X\<in> analz (analz H) ==> X\<in> analz H"
+by (erule analz.induct, blast+)
+
+lemma analz_idem [simp]: "analz (analz H) = analz H"
+by blast
+
+lemma analz_subset_iff [simp]: "(analz G \<subseteq> analz H) = (G \<subseteq> analz H)"
+apply (rule iffI)
+apply (iprover intro: subset_trans analz_increasing)  
+apply (frule analz_mono, simp) 
+done
+
+lemma analz_trans: "[| X\<in> analz G;  G \<subseteq> analz H |] ==> X\<in> analz H"
+by (drule analz_mono, blast)
+
+
+declare [[ atp_problem_prefix = "Message__analz_cut" ]]
+    declare analz_trans[intro]
+lemma analz_cut: "[| Y\<in> analz (insert X H);  X\<in> analz H |] ==> Y\<in> analz H"
+(*TOO SLOW
+by (metis analz_idem analz_increasing analz_mono insert_absorb insert_mono insert_subset) --{*317s*}
+??*)
+by (erule analz_trans, blast)
+
+
+text{*This rewrite rule helps in the simplification of messages that involve
+  the forwarding of unknown components (X).  Without it, removing occurrences
+  of X can be very complicated. *}
+lemma analz_insert_eq: "X\<in> analz H ==> analz (insert X H) = analz H"
+by (blast intro: analz_cut analz_insertI)
+
+
+text{*A congruence rule for "analz" *}
+
+declare [[ atp_problem_prefix = "Message__analz_subset_cong" ]]
+lemma analz_subset_cong:
+     "[| analz G \<subseteq> analz G'; analz H \<subseteq> analz H' |] 
+      ==> analz (G \<union> H) \<subseteq> analz (G' \<union> H')"
+apply simp
+apply (metis Un_absorb2 Un_commute Un_subset_iff Un_upper1 Un_upper2 analz_mono)
+done
+
+
+lemma analz_cong:
+     "[| analz G = analz G'; analz H = analz H'  
+               |] ==> analz (G \<union> H) = analz (G' \<union> H')"
+by (intro equalityI analz_subset_cong, simp_all) 
+
+lemma analz_insert_cong:
+     "analz H = analz H' ==> analz(insert X H) = analz(insert X H')"
+by (force simp only: insert_def intro!: analz_cong)
+
+text{*If there are no pairs or encryptions then analz does nothing*}
+lemma analz_trivial:
+     "[| \<forall>X Y. {|X,Y|} \<notin> H;  \<forall>X K. Crypt K X \<notin> H |] ==> analz H = H"
+apply safe
+apply (erule analz.induct, blast+)
+done
+
+text{*These two are obsolete (with a single Spy) but cost little to prove...*}
+lemma analz_UN_analz_lemma:
+     "X\<in> analz (\<Union>i\<in>A. analz (H i)) ==> X\<in> analz (\<Union>i\<in>A. H i)"
+apply (erule analz.induct)
+apply (blast intro: analz_mono [THEN [2] rev_subsetD])+
+done
+
+lemma analz_UN_analz [simp]: "analz (\<Union>i\<in>A. analz (H i)) = analz (\<Union>i\<in>A. H i)"
+by (blast intro: analz_UN_analz_lemma analz_mono [THEN [2] rev_subsetD])
+
+
+subsection{*Inductive relation "synth"*}
+
+text{*Inductive definition of "synth" -- what can be built up from a set of
+    messages.  A form of upward closure.  Pairs can be built, messages
+    encrypted with known keys.  Agent names are public domain.
+    Numbers can be guessed, but Nonces cannot be.  *}
+
+inductive_set
+  synth :: "msg set => msg set"
+  for H :: "msg set"
+  where
+    Inj    [intro]:   "X \<in> H ==> X \<in> synth H"
+  | Agent  [intro]:   "Agent agt \<in> synth H"
+  | Number [intro]:   "Number n  \<in> synth H"
+  | Hash   [intro]:   "X \<in> synth H ==> Hash X \<in> synth H"
+  | MPair  [intro]:   "[|X \<in> synth H;  Y \<in> synth H|] ==> {|X,Y|} \<in> synth H"
+  | Crypt  [intro]:   "[|X \<in> synth H;  Key(K) \<in> H|] ==> Crypt K X \<in> synth H"
+
+text{*Monotonicity*}
+lemma synth_mono: "G\<subseteq>H ==> synth(G) \<subseteq> synth(H)"
+  by (auto, erule synth.induct, auto)  
+
+text{*NO @{text Agent_synth}, as any Agent name can be synthesized.  
+  The same holds for @{term Number}*}
+inductive_cases Nonce_synth [elim!]: "Nonce n \<in> synth H"
+inductive_cases Key_synth   [elim!]: "Key K \<in> synth H"
+inductive_cases Hash_synth  [elim!]: "Hash X \<in> synth H"
+inductive_cases MPair_synth [elim!]: "{|X,Y|} \<in> synth H"
+inductive_cases Crypt_synth [elim!]: "Crypt K X \<in> synth H"
+
+
+lemma synth_increasing: "H \<subseteq> synth(H)"
+by blast
+
+subsubsection{*Unions *}
+
+text{*Converse fails: we can synth more from the union than from the 
+  separate parts, building a compound message using elements of each.*}
+lemma synth_Un: "synth(G) \<union> synth(H) \<subseteq> synth(G \<union> H)"
+by (intro Un_least synth_mono Un_upper1 Un_upper2)
+
+
+declare [[ atp_problem_prefix = "Message__synth_insert" ]]
+ 
+lemma synth_insert: "insert X (synth H) \<subseteq> synth(insert X H)"
+by (metis insert_iff insert_subset subset_insertI synth.Inj synth_mono)
+
+subsubsection{*Idempotence and transitivity *}
+
+lemma synth_synthD [dest!]: "X\<in> synth (synth H) ==> X\<in> synth H"
+by (erule synth.induct, blast+)
+
+lemma synth_idem: "synth (synth H) = synth H"
+by blast
+
+lemma synth_subset_iff [simp]: "(synth G \<subseteq> synth H) = (G \<subseteq> synth H)"
+apply (rule iffI)
+apply (iprover intro: subset_trans synth_increasing)  
+apply (frule synth_mono, simp add: synth_idem) 
+done
+
+lemma synth_trans: "[| X\<in> synth G;  G \<subseteq> synth H |] ==> X\<in> synth H"
+by (drule synth_mono, blast)
+
+declare [[ atp_problem_prefix = "Message__synth_cut" ]]
+lemma synth_cut: "[| Y\<in> synth (insert X H);  X\<in> synth H |] ==> Y\<in> synth H"
+(*TOO SLOW
+by (metis insert_absorb insert_mono insert_subset synth_idem synth_increasing synth_mono)
+*)
+by (erule synth_trans, blast)
+
+
+lemma Agent_synth [simp]: "Agent A \<in> synth H"
+by blast
+
+lemma Number_synth [simp]: "Number n \<in> synth H"
+by blast
+
+lemma Nonce_synth_eq [simp]: "(Nonce N \<in> synth H) = (Nonce N \<in> H)"
+by blast
+
+lemma Key_synth_eq [simp]: "(Key K \<in> synth H) = (Key K \<in> H)"
+by blast
+
+lemma Crypt_synth_eq [simp]:
+     "Key K \<notin> H ==> (Crypt K X \<in> synth H) = (Crypt K X \<in> H)"
+by blast
+
+
+lemma keysFor_synth [simp]: 
+    "keysFor (synth H) = keysFor H \<union> invKey`{K. Key K \<in> H}"
+by (unfold keysFor_def, blast)
+
+
+subsubsection{*Combinations of parts, analz and synth *}
+
+declare [[ atp_problem_prefix = "Message__parts_synth" ]]
+lemma parts_synth [simp]: "parts (synth H) = parts H \<union> synth H"
+apply (rule equalityI)
+apply (rule subsetI)
+apply (erule parts.induct)
+apply (metis UnCI)
+apply (metis MPair_synth UnCI UnE insert_absorb insert_subset parts.Fst parts_increasing)
+apply (metis MPair_synth UnCI UnE insert_absorb insert_subset parts.Snd parts_increasing)
+apply (metis Body Crypt_synth UnCI UnE insert_absorb insert_subset parts_increasing)
+apply (metis Un_subset_iff parts_increasing parts_mono synth_increasing)
+done
+
+
+
+
+declare [[ atp_problem_prefix = "Message__analz_analz_Un" ]]
+lemma analz_analz_Un [simp]: "analz (analz G \<union> H) = analz (G \<union> H)"
+apply (rule equalityI);
+apply (metis analz_idem analz_subset_cong order_eq_refl)
+apply (metis analz_increasing analz_subset_cong order_eq_refl)
+done
+
+declare [[ atp_problem_prefix = "Message__analz_synth_Un" ]]
+    declare analz_mono [intro] analz.Fst [intro] analz.Snd [intro] Un_least [intro]
+lemma analz_synth_Un [simp]: "analz (synth G \<union> H) = analz (G \<union> H) \<union> synth G"
+apply (rule equalityI)
+apply (rule subsetI)
+apply (erule analz.induct)
+apply (metis UnCI UnE Un_commute analz.Inj)
+apply (metis MPair_synth UnCI UnE Un_commute Un_upper1 analz.Fst analz_increasing analz_mono insert_absorb insert_subset)
+apply (metis MPair_synth UnCI UnE Un_commute Un_upper1 analz.Snd analz_increasing analz_mono insert_absorb insert_subset)
+apply (blast intro: analz.Decrypt)
+apply blast
+done
+
+
+declare [[ atp_problem_prefix = "Message__analz_synth" ]]
+lemma analz_synth [simp]: "analz (synth H) = analz H \<union> synth H"
+proof (neg_clausify)
+assume 0: "analz (synth H) \<noteq> analz H \<union> synth H"
+have 1: "\<And>X1 X3. sup (analz (sup X3 X1)) (synth X3) = analz (sup (synth X3) X1)"
+  by (metis analz_synth_Un)
+have 2: "sup (analz H) (synth H) \<noteq> analz (synth H)"
+  by (metis 0)
+have 3: "\<And>X1 X3. sup (synth X3) (analz (sup X3 X1)) = analz (sup (synth X3) X1)"
+  by (metis 1 Un_commute)
+have 4: "\<And>X3. sup (synth X3) (analz X3) = analz (sup (synth X3) {})"
+  by (metis 3 Un_empty_right)
+have 5: "\<And>X3. sup (synth X3) (analz X3) = analz (synth X3)"
+  by (metis 4 Un_empty_right)
+have 6: "\<And>X3. sup (analz X3) (synth X3) = analz (synth X3)"
+  by (metis 5 Un_commute)
+show "False"
+  by (metis 2 6)
+qed
+
+
+subsubsection{*For reasoning about the Fake rule in traces *}
+
+declare [[ atp_problem_prefix = "Message__parts_insert_subset_Un" ]]
+lemma parts_insert_subset_Un: "X\<in> G ==> parts(insert X H) \<subseteq> parts G \<union> parts H"
+proof (neg_clausify)
+assume 0: "X \<in> G"
+assume 1: "\<not> parts (insert X H) \<subseteq> parts G \<union> parts H"
+have 2: "\<not> parts (insert X H) \<subseteq> parts (G \<union> H)"
+  by (metis 1 parts_Un)
+have 3: "\<not> insert X H \<subseteq> G \<union> H"
+  by (metis 2 parts_mono)
+have 4: "X \<notin> G \<union> H \<or> \<not> H \<subseteq> G \<union> H"
+  by (metis 3 insert_subset)
+have 5: "X \<notin> G \<union> H"
+  by (metis 4 Un_upper2)
+have 6: "X \<notin> G"
+  by (metis 5 UnCI)
+show "False"
+  by (metis 6 0)
+qed
+
+declare [[ atp_problem_prefix = "Message__Fake_parts_insert" ]]
+lemma Fake_parts_insert:
+     "X \<in> synth (analz H) ==>  
+      parts (insert X H) \<subseteq> synth (analz H) \<union> parts H"
+proof (neg_clausify)
+assume 0: "X \<in> synth (analz H)"
+assume 1: "\<not> parts (insert X H) \<subseteq> synth (analz H) \<union> parts H"
+have 2: "\<And>X3. parts X3 \<union> synth (analz X3) = parts (synth (analz X3))"
+  by (metis parts_synth parts_analz)
+have 3: "\<And>X3. analz X3 \<union> synth (analz X3) = analz (synth (analz X3))"
+  by (metis analz_synth analz_idem)
+have 4: "\<And>X3. analz X3 \<subseteq> analz (synth X3)"
+  by (metis Un_upper1 analz_synth)
+have 5: "\<not> parts (insert X H) \<subseteq> parts H \<union> synth (analz H)"
+  by (metis 1 Un_commute)
+have 6: "\<not> parts (insert X H) \<subseteq> parts (synth (analz H))"
+  by (metis 5 2)
+have 7: "\<not> insert X H \<subseteq> synth (analz H)"
+  by (metis 6 parts_mono)
+have 8: "X \<notin> synth (analz H) \<or> \<not> H \<subseteq> synth (analz H)"
+  by (metis 7 insert_subset)
+have 9: "\<not> H \<subseteq> synth (analz H)"
+  by (metis 8 0)
+have 10: "\<And>X3. X3 \<subseteq> analz (synth X3)"
+  by (metis analz_subset_iff 4)
+have 11: "\<And>X3. X3 \<subseteq> analz (synth (analz X3))"
+  by (metis analz_subset_iff 10)
+have 12: "\<And>X3. analz (synth (analz X3)) = synth (analz X3) \<or>
+     \<not> analz X3 \<subseteq> synth (analz X3)"
+  by (metis Un_absorb1 3)
+have 13: "\<And>X3. analz (synth (analz X3)) = synth (analz X3)"
+  by (metis 12 synth_increasing)
+have 14: "\<And>X3. X3 \<subseteq> synth (analz X3)"
+  by (metis 11 13)
+show "False"
+  by (metis 9 14)
+qed
+
+lemma Fake_parts_insert_in_Un:
+     "[|Z \<in> parts (insert X H);  X: synth (analz H)|] 
+      ==> Z \<in>  synth (analz H) \<union> parts H";
+by (blast dest: Fake_parts_insert  [THEN subsetD, dest])
+
+declare [[ atp_problem_prefix = "Message__Fake_analz_insert" ]]
+    declare analz_mono [intro] synth_mono [intro] 
+lemma Fake_analz_insert:
+     "X\<in> synth (analz G) ==>  
+      analz (insert X H) \<subseteq> synth (analz G) \<union> analz (G \<union> H)"
+by (metis Un_commute Un_insert_left Un_insert_right Un_upper1 analz_analz_Un analz_mono analz_synth_Un equalityE insert_absorb order_le_less xt1(12))
+
+declare [[ atp_problem_prefix = "Message__Fake_analz_insert_simpler" ]]
+(*simpler problems?  BUT METIS CAN'T PROVE
+lemma Fake_analz_insert_simpler:
+     "X\<in> synth (analz G) ==>  
+      analz (insert X H) \<subseteq> synth (analz G) \<union> analz (G \<union> H)"
+apply (rule subsetI)
+apply (subgoal_tac "x \<in> analz (synth (analz G) \<union> H) ")
+apply (metis Un_commute analz_analz_Un analz_synth_Un)
+apply (metis Un_commute Un_upper1 Un_upper2 analz_cut analz_increasing analz_mono insert_absorb insert_mono insert_subset)
+done
+*)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Metis_Examples/ROOT.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,8 @@
+(*  Title:      HOL/Metis_Examples/ROOT.ML
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+
+Testing the metis method.
+*)
+
+use_thys ["set", "BigO", "Abstraction", "BT", "Message", "Tarski", "TransClosure"];
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Metis_Examples/Tarski.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,1110 @@
+(*  Title:      HOL/MetisTest/Tarski.thy
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+
+Testing the metis method.
+*)
+
+header {* The Full Theorem of Tarski *}
+
+theory Tarski
+imports Main FuncSet
+begin
+
+(*Many of these higher-order problems appear to be impossible using the
+current linkup. They often seem to need either higher-order unification
+or explicit reasoning about connectives such as conjunction. The numerous
+set comprehensions are to blame.*)
+
+
+record 'a potype =
+  pset  :: "'a set"
+  order :: "('a * 'a) set"
+
+constdefs
+  monotone :: "['a => 'a, 'a set, ('a *'a)set] => bool"
+  "monotone f A r == \<forall>x\<in>A. \<forall>y\<in>A. (x, y): r --> ((f x), (f y)) : r"
+
+  least :: "['a => bool, 'a potype] => 'a"
+  "least P po == @ x. x: pset po & P x &
+                       (\<forall>y \<in> pset po. P y --> (x,y): order po)"
+
+  greatest :: "['a => bool, 'a potype] => 'a"
+  "greatest P po == @ x. x: pset po & P x &
+                          (\<forall>y \<in> pset po. P y --> (y,x): order po)"
+
+  lub  :: "['a set, 'a potype] => 'a"
+  "lub S po == least (%x. \<forall>y\<in>S. (y,x): order po) po"
+
+  glb  :: "['a set, 'a potype] => 'a"
+  "glb S po == greatest (%x. \<forall>y\<in>S. (x,y): order po) po"
+
+  isLub :: "['a set, 'a potype, 'a] => bool"
+  "isLub S po == %L. (L: pset po & (\<forall>y\<in>S. (y,L): order po) &
+                   (\<forall>z\<in>pset po. (\<forall>y\<in>S. (y,z): order po) --> (L,z): order po))"
+
+  isGlb :: "['a set, 'a potype, 'a] => bool"
+  "isGlb S po == %G. (G: pset po & (\<forall>y\<in>S. (G,y): order po) &
+                 (\<forall>z \<in> pset po. (\<forall>y\<in>S. (z,y): order po) --> (z,G): order po))"
+
+  "fix"    :: "[('a => 'a), 'a set] => 'a set"
+  "fix f A  == {x. x: A & f x = x}"
+
+  interval :: "[('a*'a) set,'a, 'a ] => 'a set"
+  "interval r a b == {x. (a,x): r & (x,b): r}"
+
+constdefs
+  Bot :: "'a potype => 'a"
+  "Bot po == least (%x. True) po"
+
+  Top :: "'a potype => 'a"
+  "Top po == greatest (%x. True) po"
+
+  PartialOrder :: "('a potype) set"
+  "PartialOrder == {P. refl_on (pset P) (order P) & antisym (order P) &
+                       trans (order P)}"
+
+  CompleteLattice :: "('a potype) set"
+  "CompleteLattice == {cl. cl: PartialOrder &
+                        (\<forall>S. S \<subseteq> pset cl --> (\<exists>L. isLub S cl L)) &
+                        (\<forall>S. S \<subseteq> pset cl --> (\<exists>G. isGlb S cl G))}"
+
+  induced :: "['a set, ('a * 'a) set] => ('a *'a)set"
+  "induced A r == {(a,b). a : A & b: A & (a,b): r}"
+
+constdefs
+  sublattice :: "('a potype * 'a set)set"
+  "sublattice ==
+      SIGMA cl: CompleteLattice.
+          {S. S \<subseteq> pset cl &
+           (| pset = S, order = induced S (order cl) |): CompleteLattice }"
+
+syntax
+  "@SL"  :: "['a set, 'a potype] => bool" ("_ <<= _" [51,50]50)
+
+translations
+  "S <<= cl" == "S : sublattice `` {cl}"
+
+constdefs
+  dual :: "'a potype => 'a potype"
+  "dual po == (| pset = pset po, order = converse (order po) |)"
+
+locale PO =
+  fixes cl :: "'a potype"
+    and A  :: "'a set"
+    and r  :: "('a * 'a) set"
+  assumes cl_po:  "cl : PartialOrder"
+  defines A_def: "A == pset cl"
+     and  r_def: "r == order cl"
+
+locale CL = PO +
+  assumes cl_co:  "cl : CompleteLattice"
+
+definition CLF_set :: "('a potype * ('a => 'a)) set" where
+  "CLF_set = (SIGMA cl: CompleteLattice.
+            {f. f: pset cl -> pset cl & monotone f (pset cl) (order cl)})"
+
+locale CLF = CL +
+  fixes f :: "'a => 'a"
+    and P :: "'a set"
+  assumes f_cl:  "(cl,f) : CLF_set" (*was the equivalent "f : CLF``{cl}"*)
+  defines P_def: "P == fix f A"
+
+
+locale Tarski = CLF +
+  fixes Y     :: "'a set"
+    and intY1 :: "'a set"
+    and v     :: "'a"
+  assumes
+    Y_ss: "Y \<subseteq> P"
+  defines
+    intY1_def: "intY1 == interval r (lub Y cl) (Top cl)"
+    and v_def: "v == glb {x. ((%x: intY1. f x) x, x): induced intY1 r &
+                             x: intY1}
+                      (| pset=intY1, order=induced intY1 r|)"
+
+
+subsection {* Partial Order *}
+
+lemma (in PO) PO_imp_refl_on: "refl_on A r"
+apply (insert cl_po)
+apply (simp add: PartialOrder_def A_def r_def)
+done
+
+lemma (in PO) PO_imp_sym: "antisym r"
+apply (insert cl_po)
+apply (simp add: PartialOrder_def r_def)
+done
+
+lemma (in PO) PO_imp_trans: "trans r"
+apply (insert cl_po)
+apply (simp add: PartialOrder_def r_def)
+done
+
+lemma (in PO) reflE: "x \<in> A ==> (x, x) \<in> r"
+apply (insert cl_po)
+apply (simp add: PartialOrder_def refl_on_def A_def r_def)
+done
+
+lemma (in PO) antisymE: "[| (a, b) \<in> r; (b, a) \<in> r |] ==> a = b"
+apply (insert cl_po)
+apply (simp add: PartialOrder_def antisym_def r_def)
+done
+
+lemma (in PO) transE: "[| (a, b) \<in> r; (b, c) \<in> r|] ==> (a,c) \<in> r"
+apply (insert cl_po)
+apply (simp add: PartialOrder_def r_def)
+apply (unfold trans_def, fast)
+done
+
+lemma (in PO) monotoneE:
+     "[| monotone f A r;  x \<in> A; y \<in> A; (x, y) \<in> r |] ==> (f x, f y) \<in> r"
+by (simp add: monotone_def)
+
+lemma (in PO) po_subset_po:
+     "S \<subseteq> A ==> (| pset = S, order = induced S r |) \<in> PartialOrder"
+apply (simp (no_asm) add: PartialOrder_def)
+apply auto
+-- {* refl *}
+apply (simp add: refl_on_def induced_def)
+apply (blast intro: reflE)
+-- {* antisym *}
+apply (simp add: antisym_def induced_def)
+apply (blast intro: antisymE)
+-- {* trans *}
+apply (simp add: trans_def induced_def)
+apply (blast intro: transE)
+done
+
+lemma (in PO) indE: "[| (x, y) \<in> induced S r; S \<subseteq> A |] ==> (x, y) \<in> r"
+by (simp add: add: induced_def)
+
+lemma (in PO) indI: "[| (x, y) \<in> r; x \<in> S; y \<in> S |] ==> (x, y) \<in> induced S r"
+by (simp add: add: induced_def)
+
+lemma (in CL) CL_imp_ex_isLub: "S \<subseteq> A ==> \<exists>L. isLub S cl L"
+apply (insert cl_co)
+apply (simp add: CompleteLattice_def A_def)
+done
+
+declare (in CL) cl_co [simp]
+
+lemma isLub_lub: "(\<exists>L. isLub S cl L) = isLub S cl (lub S cl)"
+by (simp add: lub_def least_def isLub_def some_eq_ex [symmetric])
+
+lemma isGlb_glb: "(\<exists>G. isGlb S cl G) = isGlb S cl (glb S cl)"
+by (simp add: glb_def greatest_def isGlb_def some_eq_ex [symmetric])
+
+lemma isGlb_dual_isLub: "isGlb S cl = isLub S (dual cl)"
+by (simp add: isLub_def isGlb_def dual_def converse_def)
+
+lemma isLub_dual_isGlb: "isLub S cl = isGlb S (dual cl)"
+by (simp add: isLub_def isGlb_def dual_def converse_def)
+
+lemma (in PO) dualPO: "dual cl \<in> PartialOrder"
+apply (insert cl_po)
+apply (simp add: PartialOrder_def dual_def refl_on_converse
+                 trans_converse antisym_converse)
+done
+
+lemma Rdual:
+     "\<forall>S. (S \<subseteq> A -->( \<exists>L. isLub S (| pset = A, order = r|) L))
+      ==> \<forall>S. (S \<subseteq> A --> (\<exists>G. isGlb S (| pset = A, order = r|) G))"
+apply safe
+apply (rule_tac x = "lub {y. y \<in> A & (\<forall>k \<in> S. (y, k) \<in> r)}
+                      (|pset = A, order = r|) " in exI)
+apply (drule_tac x = "{y. y \<in> A & (\<forall>k \<in> S. (y,k) \<in> r) }" in spec)
+apply (drule mp, fast)
+apply (simp add: isLub_lub isGlb_def)
+apply (simp add: isLub_def, blast)
+done
+
+lemma lub_dual_glb: "lub S cl = glb S (dual cl)"
+by (simp add: lub_def glb_def least_def greatest_def dual_def converse_def)
+
+lemma glb_dual_lub: "glb S cl = lub S (dual cl)"
+by (simp add: lub_def glb_def least_def greatest_def dual_def converse_def)
+
+lemma CL_subset_PO: "CompleteLattice \<subseteq> PartialOrder"
+by (simp add: PartialOrder_def CompleteLattice_def, fast)
+
+lemmas CL_imp_PO = CL_subset_PO [THEN subsetD]
+
+declare PO.PO_imp_refl_on  [OF PO.intro [OF CL_imp_PO], simp]
+declare PO.PO_imp_sym   [OF PO.intro [OF CL_imp_PO], simp]
+declare PO.PO_imp_trans [OF PO.intro [OF CL_imp_PO], simp]
+
+lemma (in CL) CO_refl_on: "refl_on A r"
+by (rule PO_imp_refl_on)
+
+lemma (in CL) CO_antisym: "antisym r"
+by (rule PO_imp_sym)
+
+lemma (in CL) CO_trans: "trans r"
+by (rule PO_imp_trans)
+
+lemma CompleteLatticeI:
+     "[| po \<in> PartialOrder; (\<forall>S. S \<subseteq> pset po --> (\<exists>L. isLub S po L));
+         (\<forall>S. S \<subseteq> pset po --> (\<exists>G. isGlb S po G))|]
+      ==> po \<in> CompleteLattice"
+apply (unfold CompleteLattice_def, blast)
+done
+
+lemma (in CL) CL_dualCL: "dual cl \<in> CompleteLattice"
+apply (insert cl_co)
+apply (simp add: CompleteLattice_def dual_def)
+apply (fold dual_def)
+apply (simp add: isLub_dual_isGlb [symmetric] isGlb_dual_isLub [symmetric]
+                 dualPO)
+done
+
+lemma (in PO) dualA_iff: "pset (dual cl) = pset cl"
+by (simp add: dual_def)
+
+lemma (in PO) dualr_iff: "((x, y) \<in> (order(dual cl))) = ((y, x) \<in> order cl)"
+by (simp add: dual_def)
+
+lemma (in PO) monotone_dual:
+     "monotone f (pset cl) (order cl) 
+     ==> monotone f (pset (dual cl)) (order(dual cl))"
+by (simp add: monotone_def dualA_iff dualr_iff)
+
+lemma (in PO) interval_dual:
+     "[| x \<in> A; y \<in> A|] ==> interval r x y = interval (order(dual cl)) y x"
+apply (simp add: interval_def dualr_iff)
+apply (fold r_def, fast)
+done
+
+lemma (in PO) interval_not_empty:
+     "[| trans r; interval r a b \<noteq> {} |] ==> (a, b) \<in> r"
+apply (simp add: interval_def)
+apply (unfold trans_def, blast)
+done
+
+lemma (in PO) interval_imp_mem: "x \<in> interval r a b ==> (a, x) \<in> r"
+by (simp add: interval_def)
+
+lemma (in PO) left_in_interval:
+     "[| a \<in> A; b \<in> A; interval r a b \<noteq> {} |] ==> a \<in> interval r a b"
+apply (simp (no_asm_simp) add: interval_def)
+apply (simp add: PO_imp_trans interval_not_empty)
+apply (simp add: reflE)
+done
+
+lemma (in PO) right_in_interval:
+     "[| a \<in> A; b \<in> A; interval r a b \<noteq> {} |] ==> b \<in> interval r a b"
+apply (simp (no_asm_simp) add: interval_def)
+apply (simp add: PO_imp_trans interval_not_empty)
+apply (simp add: reflE)
+done
+
+
+subsection {* sublattice *}
+
+lemma (in PO) sublattice_imp_CL:
+     "S <<= cl  ==> (| pset = S, order = induced S r |) \<in> CompleteLattice"
+by (simp add: sublattice_def CompleteLattice_def A_def r_def)
+
+lemma (in CL) sublatticeI:
+     "[| S \<subseteq> A; (| pset = S, order = induced S r |) \<in> CompleteLattice |]
+      ==> S <<= cl"
+by (simp add: sublattice_def A_def r_def)
+
+
+subsection {* lub *}
+
+lemma (in CL) lub_unique: "[| S \<subseteq> A; isLub S cl x; isLub S cl L|] ==> x = L"
+apply (rule antisymE)
+apply (auto simp add: isLub_def r_def)
+done
+
+lemma (in CL) lub_upper: "[|S \<subseteq> A; x \<in> S|] ==> (x, lub S cl) \<in> r"
+apply (rule CL_imp_ex_isLub [THEN exE], assumption)
+apply (unfold lub_def least_def)
+apply (rule some_equality [THEN ssubst])
+  apply (simp add: isLub_def)
+ apply (simp add: lub_unique A_def isLub_def)
+apply (simp add: isLub_def r_def)
+done
+
+lemma (in CL) lub_least:
+     "[| S \<subseteq> A; L \<in> A; \<forall>x \<in> S. (x,L) \<in> r |] ==> (lub S cl, L) \<in> r"
+apply (rule CL_imp_ex_isLub [THEN exE], assumption)
+apply (unfold lub_def least_def)
+apply (rule_tac s=x in some_equality [THEN ssubst])
+  apply (simp add: isLub_def)
+ apply (simp add: lub_unique A_def isLub_def)
+apply (simp add: isLub_def r_def A_def)
+done
+
+lemma (in CL) lub_in_lattice: "S \<subseteq> A ==> lub S cl \<in> A"
+apply (rule CL_imp_ex_isLub [THEN exE], assumption)
+apply (unfold lub_def least_def)
+apply (subst some_equality)
+apply (simp add: isLub_def)
+prefer 2 apply (simp add: isLub_def A_def)
+apply (simp add: lub_unique A_def isLub_def)
+done
+
+lemma (in CL) lubI:
+     "[| S \<subseteq> A; L \<in> A; \<forall>x \<in> S. (x,L) \<in> r;
+         \<forall>z \<in> A. (\<forall>y \<in> S. (y,z) \<in> r) --> (L,z) \<in> r |] ==> L = lub S cl"
+apply (rule lub_unique, assumption)
+apply (simp add: isLub_def A_def r_def)
+apply (unfold isLub_def)
+apply (rule conjI)
+apply (fold A_def r_def)
+apply (rule lub_in_lattice, assumption)
+apply (simp add: lub_upper lub_least)
+done
+
+lemma (in CL) lubIa: "[| S \<subseteq> A; isLub S cl L |] ==> L = lub S cl"
+by (simp add: lubI isLub_def A_def r_def)
+
+lemma (in CL) isLub_in_lattice: "isLub S cl L ==> L \<in> A"
+by (simp add: isLub_def  A_def)
+
+lemma (in CL) isLub_upper: "[|isLub S cl L; y \<in> S|] ==> (y, L) \<in> r"
+by (simp add: isLub_def r_def)
+
+lemma (in CL) isLub_least:
+     "[| isLub S cl L; z \<in> A; \<forall>y \<in> S. (y, z) \<in> r|] ==> (L, z) \<in> r"
+by (simp add: isLub_def A_def r_def)
+
+lemma (in CL) isLubI:
+     "[| L \<in> A; \<forall>y \<in> S. (y, L) \<in> r;
+         (\<forall>z \<in> A. (\<forall>y \<in> S. (y, z):r) --> (L, z) \<in> r)|] ==> isLub S cl L"
+by (simp add: isLub_def A_def r_def)
+
+
+
+subsection {* glb *}
+
+lemma (in CL) glb_in_lattice: "S \<subseteq> A ==> glb S cl \<in> A"
+apply (subst glb_dual_lub)
+apply (simp add: A_def)
+apply (rule dualA_iff [THEN subst])
+apply (rule CL.lub_in_lattice)
+apply (rule CL.intro)
+apply (rule PO.intro)
+apply (rule dualPO)
+apply (rule CL_axioms.intro)
+apply (rule CL_dualCL)
+apply (simp add: dualA_iff)
+done
+
+lemma (in CL) glb_lower: "[|S \<subseteq> A; x \<in> S|] ==> (glb S cl, x) \<in> r"
+apply (subst glb_dual_lub)
+apply (simp add: r_def)
+apply (rule dualr_iff [THEN subst])
+apply (rule CL.lub_upper)
+apply (rule CL.intro)
+apply (rule PO.intro)
+apply (rule dualPO)
+apply (rule CL_axioms.intro)
+apply (rule CL_dualCL)
+apply (simp add: dualA_iff A_def, assumption)
+done
+
+text {*
+  Reduce the sublattice property by using substructural properties;
+  abandoned see @{text "Tarski_4.ML"}.
+*}
+
+declare (in CLF) f_cl [simp]
+
+(*never proved, 2007-01-22: Tarski__CLF_unnamed_lemma
+  NOT PROVABLE because of the conjunction used in the definition: we don't
+  allow reasoning with rules like conjE, which is essential here.*)
+declare [[ atp_problem_prefix = "Tarski__CLF_unnamed_lemma" ]]
+lemma (in CLF) [simp]:
+    "f: pset cl -> pset cl & monotone f (pset cl) (order cl)" 
+apply (insert f_cl)
+apply (unfold CLF_set_def)
+apply (erule SigmaE2) 
+apply (erule CollectE) 
+apply assumption
+done
+
+lemma (in CLF) f_in_funcset: "f \<in> A -> A"
+by (simp add: A_def)
+
+lemma (in CLF) monotone_f: "monotone f A r"
+by (simp add: A_def r_def)
+
+(*never proved, 2007-01-22*)
+declare [[ atp_problem_prefix = "Tarski__CLF_CLF_dual" ]]
+declare (in CLF) CLF_set_def [simp] CL_dualCL [simp] monotone_dual [simp] dualA_iff [simp]
+
+lemma (in CLF) CLF_dual: "(dual cl, f) \<in> CLF_set" 
+apply (simp del: dualA_iff)
+apply (simp)
+done
+
+declare (in CLF) CLF_set_def[simp del] CL_dualCL[simp del] monotone_dual[simp del]
+          dualA_iff[simp del]
+
+
+subsection {* fixed points *}
+
+lemma fix_subset: "fix f A \<subseteq> A"
+by (simp add: fix_def, fast)
+
+lemma fix_imp_eq: "x \<in> fix f A ==> f x = x"
+by (simp add: fix_def)
+
+lemma fixf_subset:
+     "[| A \<subseteq> B; x \<in> fix (%y: A. f y) A |] ==> x \<in> fix f B"
+by (simp add: fix_def, auto)
+
+
+subsection {* lemmas for Tarski, lub *}
+
+(*never proved, 2007-01-22*)
+declare [[ atp_problem_prefix = "Tarski__CLF_lubH_le_flubH" ]]
+  declare CL.lub_least[intro] CLF.f_in_funcset[intro] funcset_mem[intro] CL.lub_in_lattice[intro] PO.transE[intro] PO.monotoneE[intro] CLF.monotone_f[intro] CL.lub_upper[intro] 
+lemma (in CLF) lubH_le_flubH:
+     "H = {x. (x, f x) \<in> r & x \<in> A} ==> (lub H cl, f (lub H cl)) \<in> r"
+apply (rule lub_least, fast)
+apply (rule f_in_funcset [THEN funcset_mem])
+apply (rule lub_in_lattice, fast)
+-- {* @{text "\<forall>x:H. (x, f (lub H r)) \<in> r"} *}
+apply (rule ballI)
+(*never proved, 2007-01-22*)
+using [[ atp_problem_prefix = "Tarski__CLF_lubH_le_flubH_simpler" ]]
+apply (rule transE)
+-- {* instantiates @{text "(x, ?z) \<in> order cl to (x, f x)"}, *}
+-- {* because of the def of @{text H} *}
+apply fast
+-- {* so it remains to show @{text "(f x, f (lub H cl)) \<in> r"} *}
+apply (rule_tac f = "f" in monotoneE)
+apply (rule monotone_f, fast)
+apply (rule lub_in_lattice, fast)
+apply (rule lub_upper, fast)
+apply assumption
+done
+  declare CL.lub_least[rule del] CLF.f_in_funcset[rule del] 
+          funcset_mem[rule del] CL.lub_in_lattice[rule del] 
+          PO.transE[rule del] PO.monotoneE[rule del] 
+          CLF.monotone_f[rule del] CL.lub_upper[rule del] 
+
+(*never proved, 2007-01-22*)
+declare [[ atp_problem_prefix = "Tarski__CLF_flubH_le_lubH" ]]
+  declare CLF.f_in_funcset[intro] funcset_mem[intro] CL.lub_in_lattice[intro]
+       PO.monotoneE[intro] CLF.monotone_f[intro] CL.lub_upper[intro] 
+       CLF.lubH_le_flubH[simp]
+lemma (in CLF) flubH_le_lubH:
+     "[|  H = {x. (x, f x) \<in> r & x \<in> A} |] ==> (f (lub H cl), lub H cl) \<in> r"
+apply (rule lub_upper, fast)
+apply (rule_tac t = "H" in ssubst, assumption)
+apply (rule CollectI)
+apply (rule conjI)
+using [[ atp_problem_prefix = "Tarski__CLF_flubH_le_lubH_simpler" ]]
+(*??no longer terminates, with combinators
+apply (metis CO_refl_on lubH_le_flubH monotone_def monotone_f reflD1 reflD2) 
+*)
+apply (metis CO_refl_on lubH_le_flubH monotoneE [OF monotone_f] refl_onD1 refl_onD2)
+apply (metis CO_refl_on lubH_le_flubH refl_onD2)
+done
+  declare CLF.f_in_funcset[rule del] funcset_mem[rule del] 
+          CL.lub_in_lattice[rule del] PO.monotoneE[rule del] 
+          CLF.monotone_f[rule del] CL.lub_upper[rule del] 
+          CLF.lubH_le_flubH[simp del]
+
+
+(*never proved, 2007-01-22*)
+declare [[ atp_problem_prefix = "Tarski__CLF_lubH_is_fixp" ]]
+(*Single-step version fails. The conjecture clauses refer to local abstraction
+functions (Frees), which prevents expand_defs_tac from removing those 
+"definitions" at the end of the proof. *)
+lemma (in CLF) lubH_is_fixp:
+     "H = {x. (x, f x) \<in> r & x \<in> A} ==> lub H cl \<in> fix f A"
+apply (simp add: fix_def)
+apply (rule conjI)
+proof (neg_clausify)
+assume 0: "H =
+Collect
+ (COMBS (COMBB op \<and> (COMBC (COMBB op \<in> (COMBS Pair f)) r)) (COMBC op \<in> A))"
+assume 1: "lub (Collect
+      (COMBS (COMBB op \<and> (COMBC (COMBB op \<in> (COMBS Pair f)) r))
+        (COMBC op \<in> A)))
+ cl
+\<notin> A"
+have 2: "lub H cl \<notin> A"
+  by (metis 1 0)
+have 3: "(lub H cl, f (lub H cl)) \<in> r"
+  by (metis lubH_le_flubH 0)
+have 4: "(f (lub H cl), lub H cl) \<in> r"
+  by (metis flubH_le_lubH 0)
+have 5: "lub H cl = f (lub H cl) \<or> (lub H cl, f (lub H cl)) \<notin> r"
+  by (metis antisymE 4)
+have 6: "lub H cl = f (lub H cl)"
+  by (metis 5 3)
+have 7: "(lub H cl, lub H cl) \<in> r"
+  by (metis 6 4)
+have 8: "\<And>X1. lub H cl \<in> X1 \<or> \<not> refl_on X1 r"
+  by (metis 7 refl_onD2)
+have 9: "\<not> refl_on A r"
+  by (metis 8 2)
+show "False"
+  by (metis CO_refl_on 9);
+next --{*apparently the way to insert a second structured proof*}
+  show "H = {x. (x, f x) \<in> r \<and> x \<in> A} \<Longrightarrow>
+  f (lub {x. (x, f x) \<in> r \<and> x \<in> A} cl) = lub {x. (x, f x) \<in> r \<and> x \<in> A} cl"
+  proof (neg_clausify)
+  assume 0: "H =
+  Collect
+   (COMBS (COMBB op \<and> (COMBC (COMBB op \<in> (COMBS Pair f)) r)) (COMBC op \<in> A))"
+  assume 1: "f (lub (Collect
+           (COMBS (COMBB op \<and> (COMBC (COMBB op \<in> (COMBS Pair f)) r))
+             (COMBC op \<in> A)))
+      cl) \<noteq>
+  lub (Collect
+        (COMBS (COMBB op \<and> (COMBC (COMBB op \<in> (COMBS Pair f)) r))
+          (COMBC op \<in> A)))
+   cl"
+  have 2: "f (lub H cl) \<noteq>
+  lub (Collect
+        (COMBS (COMBB op \<and> (COMBC (COMBB op \<in> (COMBS Pair f)) r))
+          (COMBC op \<in> A)))
+   cl"
+    by (metis 1 0)
+  have 3: "f (lub H cl) \<noteq> lub H cl"
+    by (metis 2 0)
+  have 4: "(lub H cl, f (lub H cl)) \<in> r"
+    by (metis lubH_le_flubH 0)
+  have 5: "(f (lub H cl), lub H cl) \<in> r"
+    by (metis flubH_le_lubH 0)
+  have 6: "lub H cl = f (lub H cl) \<or> (lub H cl, f (lub H cl)) \<notin> r"
+    by (metis antisymE 5)
+  have 7: "lub H cl = f (lub H cl)"
+    by (metis 6 4)
+  show "False"
+    by (metis 3 7)
+  qed
+qed
+
+lemma (in CLF) (*lubH_is_fixp:*)
+     "H = {x. (x, f x) \<in> r & x \<in> A} ==> lub H cl \<in> fix f A"
+apply (simp add: fix_def)
+apply (rule conjI)
+using [[ atp_problem_prefix = "Tarski__CLF_lubH_is_fixp_simpler" ]]
+apply (metis CO_refl_on lubH_le_flubH refl_onD1)
+apply (metis antisymE flubH_le_lubH lubH_le_flubH)
+done
+
+lemma (in CLF) fix_in_H:
+     "[| H = {x. (x, f x) \<in> r & x \<in> A};  x \<in> P |] ==> x \<in> H"
+by (simp add: P_def fix_imp_eq [of _ f A] reflE CO_refl_on
+                    fix_subset [of f A, THEN subsetD])
+
+
+lemma (in CLF) fixf_le_lubH:
+     "H = {x. (x, f x) \<in> r & x \<in> A} ==> \<forall>x \<in> fix f A. (x, lub H cl) \<in> r"
+apply (rule ballI)
+apply (rule lub_upper, fast)
+apply (rule fix_in_H)
+apply (simp_all add: P_def)
+done
+
+declare [[ atp_problem_prefix = "Tarski__CLF_lubH_least_fixf" ]]
+lemma (in CLF) lubH_least_fixf:
+     "H = {x. (x, f x) \<in> r & x \<in> A}
+      ==> \<forall>L. (\<forall>y \<in> fix f A. (y,L) \<in> r) --> (lub H cl, L) \<in> r"
+apply (metis P_def lubH_is_fixp)
+done
+
+subsection {* Tarski fixpoint theorem 1, first part *}
+declare [[ atp_problem_prefix = "Tarski__CLF_T_thm_1_lub" ]]
+  declare CL.lubI[intro] fix_subset[intro] CL.lub_in_lattice[intro] 
+          CLF.fixf_le_lubH[simp] CLF.lubH_least_fixf[simp]
+lemma (in CLF) T_thm_1_lub: "lub P cl = lub {x. (x, f x) \<in> r & x \<in> A} cl"
+(*sledgehammer;*)
+apply (rule sym)
+apply (simp add: P_def)
+apply (rule lubI)
+using [[ atp_problem_prefix = "Tarski__CLF_T_thm_1_lub_simpler" ]]
+apply (metis P_def fix_subset) 
+apply (metis Collect_conj_eq Collect_mem_eq Int_commute Int_lower1 lub_in_lattice vimage_def)
+(*??no longer terminates, with combinators
+apply (metis P_def fix_def fixf_le_lubH)
+apply (metis P_def fix_def lubH_least_fixf)
+*)
+apply (simp add: fixf_le_lubH)
+apply (simp add: lubH_least_fixf)
+done
+  declare CL.lubI[rule del] fix_subset[rule del] CL.lub_in_lattice[rule del] 
+          CLF.fixf_le_lubH[simp del] CLF.lubH_least_fixf[simp del]
+
+
+(*never proved, 2007-01-22*)
+declare [[ atp_problem_prefix = "Tarski__CLF_glbH_is_fixp" ]]
+  declare glb_dual_lub[simp] PO.dualA_iff[intro] CLF.lubH_is_fixp[intro] 
+          PO.dualPO[intro] CL.CL_dualCL[intro] PO.dualr_iff[simp]
+lemma (in CLF) glbH_is_fixp: "H = {x. (f x, x) \<in> r & x \<in> A} ==> glb H cl \<in> P"
+  -- {* Tarski for glb *}
+(*sledgehammer;*)
+apply (simp add: glb_dual_lub P_def A_def r_def)
+apply (rule dualA_iff [THEN subst])
+apply (rule CLF.lubH_is_fixp)
+apply (rule CLF.intro)
+apply (rule CL.intro)
+apply (rule PO.intro)
+apply (rule dualPO)
+apply (rule CL_axioms.intro)
+apply (rule CL_dualCL)
+apply (rule CLF_axioms.intro)
+apply (rule CLF_dual)
+apply (simp add: dualr_iff dualA_iff)
+done
+  declare glb_dual_lub[simp del] PO.dualA_iff[rule del] CLF.lubH_is_fixp[rule del] 
+          PO.dualPO[rule del] CL.CL_dualCL[rule del] PO.dualr_iff[simp del]
+
+
+(*never proved, 2007-01-22*)
+declare [[ atp_problem_prefix = "Tarski__T_thm_1_glb" ]]  (*ALL THEOREMS*)
+lemma (in CLF) T_thm_1_glb: "glb P cl = glb {x. (f x, x) \<in> r & x \<in> A} cl"
+(*sledgehammer;*)
+apply (simp add: glb_dual_lub P_def A_def r_def)
+apply (rule dualA_iff [THEN subst])
+(*never proved, 2007-01-22*)
+using [[ atp_problem_prefix = "Tarski__T_thm_1_glb_simpler" ]]  (*ALL THEOREMS*)
+(*sledgehammer;*)
+apply (simp add: CLF.T_thm_1_lub [of _ f, OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro,
+  OF dualPO CL_dualCL] dualPO CL_dualCL CLF_dual dualr_iff)
+done
+
+subsection {* interval *}
+
+
+declare [[ atp_problem_prefix = "Tarski__rel_imp_elem" ]]
+  declare (in CLF) CO_refl_on[simp] refl_on_def [simp]
+lemma (in CLF) rel_imp_elem: "(x, y) \<in> r ==> x \<in> A"
+by (metis CO_refl_on refl_onD1)
+  declare (in CLF) CO_refl_on[simp del]  refl_on_def [simp del]
+
+declare [[ atp_problem_prefix = "Tarski__interval_subset" ]]
+  declare (in CLF) rel_imp_elem[intro] 
+  declare interval_def [simp]
+lemma (in CLF) interval_subset: "[| a \<in> A; b \<in> A |] ==> interval r a b \<subseteq> A"
+by (metis CO_refl_on interval_imp_mem refl_onD refl_onD2 rel_imp_elem subset_eq)
+  declare (in CLF) rel_imp_elem[rule del] 
+  declare interval_def [simp del]
+
+
+lemma (in CLF) intervalI:
+     "[| (a, x) \<in> r; (x, b) \<in> r |] ==> x \<in> interval r a b"
+by (simp add: interval_def)
+
+lemma (in CLF) interval_lemma1:
+     "[| S \<subseteq> interval r a b; x \<in> S |] ==> (a, x) \<in> r"
+by (unfold interval_def, fast)
+
+lemma (in CLF) interval_lemma2:
+     "[| S \<subseteq> interval r a b; x \<in> S |] ==> (x, b) \<in> r"
+by (unfold interval_def, fast)
+
+lemma (in CLF) a_less_lub:
+     "[| S \<subseteq> A; S \<noteq> {};
+         \<forall>x \<in> S. (a,x) \<in> r; \<forall>y \<in> S. (y, L) \<in> r |] ==> (a,L) \<in> r"
+by (blast intro: transE)
+
+lemma (in CLF) glb_less_b:
+     "[| S \<subseteq> A; S \<noteq> {};
+         \<forall>x \<in> S. (x,b) \<in> r; \<forall>y \<in> S. (G, y) \<in> r |] ==> (G,b) \<in> r"
+by (blast intro: transE)
+
+lemma (in CLF) S_intv_cl:
+     "[| a \<in> A; b \<in> A; S \<subseteq> interval r a b |]==> S \<subseteq> A"
+by (simp add: subset_trans [OF _ interval_subset])
+
+declare [[ atp_problem_prefix = "Tarski__L_in_interval" ]]  (*ALL THEOREMS*)
+lemma (in CLF) L_in_interval:
+     "[| a \<in> A; b \<in> A; S \<subseteq> interval r a b;
+         S \<noteq> {}; isLub S cl L; interval r a b \<noteq> {} |] ==> L \<in> interval r a b" 
+(*WON'T TERMINATE
+apply (metis CO_trans intervalI interval_lemma1 interval_lemma2 isLub_least isLub_upper subset_empty subset_iff trans_def)
+*)
+apply (rule intervalI)
+apply (rule a_less_lub)
+prefer 2 apply assumption
+apply (simp add: S_intv_cl)
+apply (rule ballI)
+apply (simp add: interval_lemma1)
+apply (simp add: isLub_upper)
+-- {* @{text "(L, b) \<in> r"} *}
+apply (simp add: isLub_least interval_lemma2)
+done
+
+(*never proved, 2007-01-22*)
+declare [[ atp_problem_prefix = "Tarski__G_in_interval" ]]  (*ALL THEOREMS*)
+lemma (in CLF) G_in_interval:
+     "[| a \<in> A; b \<in> A; interval r a b \<noteq> {}; S \<subseteq> interval r a b; isGlb S cl G;
+         S \<noteq> {} |] ==> G \<in> interval r a b"
+apply (simp add: interval_dual)
+apply (simp add: CLF.L_in_interval [of _ f, OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro]
+                 dualA_iff A_def dualPO CL_dualCL CLF_dual isGlb_dual_isLub)
+done
+
+declare [[ atp_problem_prefix = "Tarski__intervalPO" ]]  (*ALL THEOREMS*)
+lemma (in CLF) intervalPO:
+     "[| a \<in> A; b \<in> A; interval r a b \<noteq> {} |]
+      ==> (| pset = interval r a b, order = induced (interval r a b) r |)
+          \<in> PartialOrder"
+proof (neg_clausify)
+assume 0: "a \<in> A"
+assume 1: "b \<in> A"
+assume 2: "\<lparr>pset = interval r a b, order = induced (interval r a b) r\<rparr> \<notin> PartialOrder"
+have 3: "\<not> interval r a b \<subseteq> A"
+  by (metis 2 po_subset_po)
+have 4: "b \<notin> A \<or> a \<notin> A"
+  by (metis 3 interval_subset)
+have 5: "a \<notin> A"
+  by (metis 4 1)
+show "False"
+  by (metis 5 0)
+qed
+
+lemma (in CLF) intv_CL_lub:
+ "[| a \<in> A; b \<in> A; interval r a b \<noteq> {} |]
+  ==> \<forall>S. S \<subseteq> interval r a b -->
+          (\<exists>L. isLub S (| pset = interval r a b,
+                          order = induced (interval r a b) r |)  L)"
+apply (intro strip)
+apply (frule S_intv_cl [THEN CL_imp_ex_isLub])
+prefer 2 apply assumption
+apply assumption
+apply (erule exE)
+-- {* define the lub for the interval as *}
+apply (rule_tac x = "if S = {} then a else L" in exI)
+apply (simp (no_asm_simp) add: isLub_def split del: split_if)
+apply (intro impI conjI)
+-- {* @{text "(if S = {} then a else L) \<in> interval r a b"} *}
+apply (simp add: CL_imp_PO L_in_interval)
+apply (simp add: left_in_interval)
+-- {* lub prop 1 *}
+apply (case_tac "S = {}")
+-- {* @{text "S = {}, y \<in> S = False => everything"} *}
+apply fast
+-- {* @{text "S \<noteq> {}"} *}
+apply simp
+-- {* @{text "\<forall>y:S. (y, L) \<in> induced (interval r a b) r"} *}
+apply (rule ballI)
+apply (simp add: induced_def  L_in_interval)
+apply (rule conjI)
+apply (rule subsetD)
+apply (simp add: S_intv_cl, assumption)
+apply (simp add: isLub_upper)
+-- {* @{text "\<forall>z:interval r a b. (\<forall>y:S. (y, z) \<in> induced (interval r a b) r \<longrightarrow> (if S = {} then a else L, z) \<in> induced (interval r a b) r"} *}
+apply (rule ballI)
+apply (rule impI)
+apply (case_tac "S = {}")
+-- {* @{text "S = {}"} *}
+apply simp
+apply (simp add: induced_def  interval_def)
+apply (rule conjI)
+apply (rule reflE, assumption)
+apply (rule interval_not_empty)
+apply (rule CO_trans)
+apply (simp add: interval_def)
+-- {* @{text "S \<noteq> {}"} *}
+apply simp
+apply (simp add: induced_def  L_in_interval)
+apply (rule isLub_least, assumption)
+apply (rule subsetD)
+prefer 2 apply assumption
+apply (simp add: S_intv_cl, fast)
+done
+
+lemmas (in CLF) intv_CL_glb = intv_CL_lub [THEN Rdual]
+
+(*never proved, 2007-01-22*)
+declare [[ atp_problem_prefix = "Tarski__interval_is_sublattice" ]]  (*ALL THEOREMS*)
+lemma (in CLF) interval_is_sublattice:
+     "[| a \<in> A; b \<in> A; interval r a b \<noteq> {} |]
+        ==> interval r a b <<= cl"
+(*sledgehammer *)
+apply (rule sublatticeI)
+apply (simp add: interval_subset)
+(*never proved, 2007-01-22*)
+using [[ atp_problem_prefix = "Tarski__interval_is_sublattice_simpler" ]]
+(*sledgehammer *)
+apply (rule CompleteLatticeI)
+apply (simp add: intervalPO)
+ apply (simp add: intv_CL_lub)
+apply (simp add: intv_CL_glb)
+done
+
+lemmas (in CLF) interv_is_compl_latt =
+    interval_is_sublattice [THEN sublattice_imp_CL]
+
+
+subsection {* Top and Bottom *}
+lemma (in CLF) Top_dual_Bot: "Top cl = Bot (dual cl)"
+by (simp add: Top_def Bot_def least_def greatest_def dualA_iff dualr_iff)
+
+lemma (in CLF) Bot_dual_Top: "Bot cl = Top (dual cl)"
+by (simp add: Top_def Bot_def least_def greatest_def dualA_iff dualr_iff)
+
+declare [[ atp_problem_prefix = "Tarski__Bot_in_lattice" ]]  (*ALL THEOREMS*)
+lemma (in CLF) Bot_in_lattice: "Bot cl \<in> A"
+(*sledgehammer; *)
+apply (simp add: Bot_def least_def)
+apply (rule_tac a="glb A cl" in someI2)
+apply (simp_all add: glb_in_lattice glb_lower 
+                     r_def [symmetric] A_def [symmetric])
+done
+
+(*first proved 2007-01-25 after relaxing relevance*)
+declare [[ atp_problem_prefix = "Tarski__Top_in_lattice" ]]  (*ALL THEOREMS*)
+lemma (in CLF) Top_in_lattice: "Top cl \<in> A"
+(*sledgehammer;*)
+apply (simp add: Top_dual_Bot A_def)
+(*first proved 2007-01-25 after relaxing relevance*)
+using [[ atp_problem_prefix = "Tarski__Top_in_lattice_simpler" ]]  (*ALL THEOREMS*)
+(*sledgehammer*)
+apply (rule dualA_iff [THEN subst])
+apply (blast intro!: CLF.Bot_in_lattice [OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro] dualPO CL_dualCL CLF_dual)
+done
+
+lemma (in CLF) Top_prop: "x \<in> A ==> (x, Top cl) \<in> r"
+apply (simp add: Top_def greatest_def)
+apply (rule_tac a="lub A cl" in someI2)
+apply (rule someI2)
+apply (simp_all add: lub_in_lattice lub_upper 
+                     r_def [symmetric] A_def [symmetric])
+done
+
+(*never proved, 2007-01-22*)
+declare [[ atp_problem_prefix = "Tarski__Bot_prop" ]]  (*ALL THEOREMS*) 
+lemma (in CLF) Bot_prop: "x \<in> A ==> (Bot cl, x) \<in> r"
+(*sledgehammer*) 
+apply (simp add: Bot_dual_Top r_def)
+apply (rule dualr_iff [THEN subst])
+apply (simp add: CLF.Top_prop [of _ f, OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro]
+                 dualA_iff A_def dualPO CL_dualCL CLF_dual)
+done
+
+declare [[ atp_problem_prefix = "Tarski__Bot_in_lattice" ]]  (*ALL THEOREMS*)
+lemma (in CLF) Top_intv_not_empty: "x \<in> A  ==> interval r x (Top cl) \<noteq> {}" 
+apply (metis Top_in_lattice Top_prop empty_iff intervalI reflE)
+done
+
+declare [[ atp_problem_prefix = "Tarski__Bot_intv_not_empty" ]]  (*ALL THEOREMS*)
+lemma (in CLF) Bot_intv_not_empty: "x \<in> A ==> interval r (Bot cl) x \<noteq> {}" 
+apply (metis Bot_prop ex_in_conv intervalI reflE rel_imp_elem)
+done
+
+
+subsection {* fixed points form a partial order *}
+
+lemma (in CLF) fixf_po: "(| pset = P, order = induced P r|) \<in> PartialOrder"
+by (simp add: P_def fix_subset po_subset_po)
+
+(*first proved 2007-01-25 after relaxing relevance*)
+declare [[ atp_problem_prefix = "Tarski__Y_subset_A" ]]
+  declare (in Tarski) P_def[simp] Y_ss [simp]
+  declare fix_subset [intro] subset_trans [intro]
+lemma (in Tarski) Y_subset_A: "Y \<subseteq> A"
+(*sledgehammer*) 
+apply (rule subset_trans [OF _ fix_subset])
+apply (rule Y_ss [simplified P_def])
+done
+  declare (in Tarski) P_def[simp del] Y_ss [simp del]
+  declare fix_subset [rule del] subset_trans [rule del]
+
+
+lemma (in Tarski) lubY_in_A: "lub Y cl \<in> A"
+  by (rule Y_subset_A [THEN lub_in_lattice])
+
+(*never proved, 2007-01-22*)
+declare [[ atp_problem_prefix = "Tarski__lubY_le_flubY" ]]  (*ALL THEOREMS*)
+lemma (in Tarski) lubY_le_flubY: "(lub Y cl, f (lub Y cl)) \<in> r"
+(*sledgehammer*) 
+apply (rule lub_least)
+apply (rule Y_subset_A)
+apply (rule f_in_funcset [THEN funcset_mem])
+apply (rule lubY_in_A)
+-- {* @{text "Y \<subseteq> P ==> f x = x"} *}
+apply (rule ballI)
+using [[ atp_problem_prefix = "Tarski__lubY_le_flubY_simpler" ]]  (*ALL THEOREMS*)
+(*sledgehammer *)
+apply (rule_tac t = "x" in fix_imp_eq [THEN subst])
+apply (erule Y_ss [simplified P_def, THEN subsetD])
+-- {* @{text "reduce (f x, f (lub Y cl)) \<in> r to (x, lub Y cl) \<in> r"} by monotonicity *}
+using [[ atp_problem_prefix = "Tarski__lubY_le_flubY_simplest" ]]  (*ALL THEOREMS*)
+(*sledgehammer*)
+apply (rule_tac f = "f" in monotoneE)
+apply (rule monotone_f)
+apply (simp add: Y_subset_A [THEN subsetD])
+apply (rule lubY_in_A)
+apply (simp add: lub_upper Y_subset_A)
+done
+
+(*first proved 2007-01-25 after relaxing relevance*)
+declare [[ atp_problem_prefix = "Tarski__intY1_subset" ]]  (*ALL THEOREMS*)
+lemma (in Tarski) intY1_subset: "intY1 \<subseteq> A"
+(*sledgehammer*) 
+apply (unfold intY1_def)
+apply (rule interval_subset)
+apply (rule lubY_in_A)
+apply (rule Top_in_lattice)
+done
+
+lemmas (in Tarski) intY1_elem = intY1_subset [THEN subsetD]
+
+(*never proved, 2007-01-22*)
+declare [[ atp_problem_prefix = "Tarski__intY1_f_closed" ]]  (*ALL THEOREMS*)
+lemma (in Tarski) intY1_f_closed: "x \<in> intY1 \<Longrightarrow> f x \<in> intY1"
+(*sledgehammer*) 
+apply (simp add: intY1_def  interval_def)
+apply (rule conjI)
+apply (rule transE)
+apply (rule lubY_le_flubY)
+-- {* @{text "(f (lub Y cl), f x) \<in> r"} *}
+using [[ atp_problem_prefix = "Tarski__intY1_f_closed_simpler" ]]  (*ALL THEOREMS*)
+(*sledgehammer [has been proved before now...]*)
+apply (rule_tac f=f in monotoneE)
+apply (rule monotone_f)
+apply (rule lubY_in_A)
+apply (simp add: intY1_def interval_def  intY1_elem)
+apply (simp add: intY1_def  interval_def)
+-- {* @{text "(f x, Top cl) \<in> r"} *} 
+apply (rule Top_prop)
+apply (rule f_in_funcset [THEN funcset_mem])
+apply (simp add: intY1_def interval_def  intY1_elem)
+done
+
+declare [[ atp_problem_prefix = "Tarski__intY1_func" ]]  (*ALL THEOREMS*)
+lemma (in Tarski) intY1_func: "(%x: intY1. f x) \<in> intY1 -> intY1"
+apply (rule restrict_in_funcset)
+apply (metis intY1_f_closed restrict_in_funcset)
+done
+
+declare [[ atp_problem_prefix = "Tarski__intY1_mono" ]]  (*ALL THEOREMS*)
+lemma (in Tarski) intY1_mono:
+     "monotone (%x: intY1. f x) intY1 (induced intY1 r)"
+(*sledgehammer *)
+apply (auto simp add: monotone_def induced_def intY1_f_closed)
+apply (blast intro: intY1_elem monotone_f [THEN monotoneE])
+done
+
+(*proof requires relaxing relevance: 2007-01-25*)
+declare [[ atp_problem_prefix = "Tarski__intY1_is_cl" ]]  (*ALL THEOREMS*)
+lemma (in Tarski) intY1_is_cl:
+    "(| pset = intY1, order = induced intY1 r |) \<in> CompleteLattice"
+(*sledgehammer*) 
+apply (unfold intY1_def)
+apply (rule interv_is_compl_latt)
+apply (rule lubY_in_A)
+apply (rule Top_in_lattice)
+apply (rule Top_intv_not_empty)
+apply (rule lubY_in_A)
+done
+
+(*never proved, 2007-01-22*)
+declare [[ atp_problem_prefix = "Tarski__v_in_P" ]]  (*ALL THEOREMS*)
+lemma (in Tarski) v_in_P: "v \<in> P"
+(*sledgehammer*) 
+apply (unfold P_def)
+apply (rule_tac A = "intY1" in fixf_subset)
+apply (rule intY1_subset)
+apply (simp add: CLF.glbH_is_fixp [OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro, OF _ intY1_is_cl, simplified]
+                 v_def CL_imp_PO intY1_is_cl CLF_set_def intY1_func intY1_mono)
+done
+
+declare [[ atp_problem_prefix = "Tarski__z_in_interval" ]]  (*ALL THEOREMS*)
+lemma (in Tarski) z_in_interval:
+     "[| z \<in> P; \<forall>y\<in>Y. (y, z) \<in> induced P r |] ==> z \<in> intY1"
+(*sledgehammer *)
+apply (unfold intY1_def P_def)
+apply (rule intervalI)
+prefer 2
+ apply (erule fix_subset [THEN subsetD, THEN Top_prop])
+apply (rule lub_least)
+apply (rule Y_subset_A)
+apply (fast elim!: fix_subset [THEN subsetD])
+apply (simp add: induced_def)
+done
+
+declare [[ atp_problem_prefix = "Tarski__fz_in_int_rel" ]]  (*ALL THEOREMS*)
+lemma (in Tarski) f'z_in_int_rel: "[| z \<in> P; \<forall>y\<in>Y. (y, z) \<in> induced P r |]
+      ==> ((%x: intY1. f x) z, z) \<in> induced intY1 r" 
+apply (metis P_def acc_def fix_imp_eq fix_subset indI reflE restrict_apply subset_eq z_in_interval)
+done
+
+(*never proved, 2007-01-22*)
+declare [[ atp_problem_prefix = "Tarski__tarski_full_lemma" ]]  (*ALL THEOREMS*)
+lemma (in Tarski) tarski_full_lemma:
+     "\<exists>L. isLub Y (| pset = P, order = induced P r |) L"
+apply (rule_tac x = "v" in exI)
+apply (simp add: isLub_def)
+-- {* @{text "v \<in> P"} *}
+apply (simp add: v_in_P)
+apply (rule conjI)
+(*sledgehammer*) 
+-- {* @{text v} is lub *}
+-- {* @{text "1. \<forall>y:Y. (y, v) \<in> induced P r"} *}
+apply (rule ballI)
+apply (simp add: induced_def subsetD v_in_P)
+apply (rule conjI)
+apply (erule Y_ss [THEN subsetD])
+apply (rule_tac b = "lub Y cl" in transE)
+apply (rule lub_upper)
+apply (rule Y_subset_A, assumption)
+apply (rule_tac b = "Top cl" in interval_imp_mem)
+apply (simp add: v_def)
+apply (fold intY1_def)
+apply (rule CL.glb_in_lattice [OF CL.intro, OF PO.intro CL_axioms.intro, OF _ intY1_is_cl, simplified])
+ apply (simp add: CL_imp_PO intY1_is_cl, force)
+-- {* @{text v} is LEAST ub *}
+apply clarify
+apply (rule indI)
+  prefer 3 apply assumption
+ prefer 2 apply (simp add: v_in_P)
+apply (unfold v_def)
+(*never proved, 2007-01-22*)
+using [[ atp_problem_prefix = "Tarski__tarski_full_lemma_simpler" ]]
+(*sledgehammer*) 
+apply (rule indE)
+apply (rule_tac [2] intY1_subset)
+(*never proved, 2007-01-22*)
+using [[ atp_problem_prefix = "Tarski__tarski_full_lemma_simplest" ]]
+(*sledgehammer*) 
+apply (rule CL.glb_lower [OF CL.intro, OF PO.intro CL_axioms.intro, OF _ intY1_is_cl, simplified])
+  apply (simp add: CL_imp_PO intY1_is_cl)
+ apply force
+apply (simp add: induced_def intY1_f_closed z_in_interval)
+apply (simp add: P_def fix_imp_eq [of _ f A] reflE
+                 fix_subset [of f A, THEN subsetD])
+done
+
+lemma CompleteLatticeI_simp:
+     "[| (| pset = A, order = r |) \<in> PartialOrder;
+         \<forall>S. S \<subseteq> A --> (\<exists>L. isLub S (| pset = A, order = r |)  L) |]
+    ==> (| pset = A, order = r |) \<in> CompleteLattice"
+by (simp add: CompleteLatticeI Rdual)
+
+
+(*never proved, 2007-01-22*)
+declare [[ atp_problem_prefix = "Tarski__Tarski_full" ]]
+  declare (in CLF) fixf_po[intro] P_def [simp] A_def [simp] r_def [simp]
+               Tarski.tarski_full_lemma [intro] cl_po [intro] cl_co [intro]
+               CompleteLatticeI_simp [intro]
+theorem (in CLF) Tarski_full:
+     "(| pset = P, order = induced P r|) \<in> CompleteLattice"
+(*sledgehammer*) 
+apply (rule CompleteLatticeI_simp)
+apply (rule fixf_po, clarify)
+(*never proved, 2007-01-22*)
+using [[ atp_problem_prefix = "Tarski__Tarski_full_simpler" ]]
+(*sledgehammer*) 
+apply (simp add: P_def A_def r_def)
+apply (blast intro!: Tarski.tarski_full_lemma [OF Tarski.intro, OF CLF.intro Tarski_axioms.intro,
+  OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro] cl_po cl_co f_cl)
+done
+  declare (in CLF) fixf_po[rule del] P_def [simp del] A_def [simp del] r_def [simp del]
+         Tarski.tarski_full_lemma [rule del] cl_po [rule del] cl_co [rule del]
+         CompleteLatticeI_simp [rule del]
+
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Metis_Examples/TransClosure.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,62 @@
+(*  Title:      HOL/MetisTest/TransClosure.thy
+    ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+
+Testing the metis method
+*)
+
+theory TransClosure
+imports Main
+begin
+
+types addr = nat
+
+datatype val
+  = Unit        -- "dummy result value of void expressions"
+  | Null        -- "null reference"
+  | Bool bool   -- "Boolean value"
+  | Intg int    -- "integer value" 
+  | Addr addr   -- "addresses of objects in the heap"
+
+consts R::"(addr \<times> addr) set"
+
+consts f:: "addr \<Rightarrow> val"
+
+declare [[ atp_problem_prefix = "TransClosure__test" ]]
+lemma "\<lbrakk> f c = Intg x; \<forall> y. f b = Intg y \<longrightarrow> y \<noteq> x; (a,b) \<in> R\<^sup>*; (b,c) \<in> R\<^sup>* \<rbrakk> 
+   \<Longrightarrow> \<exists> c. (b,c) \<in> R \<and> (a,c) \<in> R\<^sup>*"  
+by (metis Transitive_Closure.rtrancl_into_rtrancl converse_rtranclE trancl_reflcl)
+
+lemma "\<lbrakk> f c = Intg x; \<forall> y. f b = Intg y \<longrightarrow> y \<noteq> x; (a,b) \<in> R\<^sup>*; (b,c) \<in> R\<^sup>* \<rbrakk> 
+   \<Longrightarrow> \<exists> c. (b,c) \<in> R \<and> (a,c) \<in> R\<^sup>*"
+proof (neg_clausify)
+assume 0: "f c = Intg x"
+assume 1: "(a, b) \<in> R\<^sup>*"
+assume 2: "(b, c) \<in> R\<^sup>*"
+assume 3: "f b \<noteq> Intg x"
+assume 4: "\<And>A. (b, A) \<notin> R \<or> (a, A) \<notin> R\<^sup>*"
+have 5: "b = c \<or> b \<in> Domain R"
+  by (metis Not_Domain_rtrancl 2)
+have 6: "\<And>X1. (a, X1) \<in> R\<^sup>* \<or> (b, X1) \<notin> R"
+  by (metis Transitive_Closure.rtrancl_into_rtrancl 1)
+have 7: "\<And>X1. (b, X1) \<notin> R"
+  by (metis 6 4)
+have 8: "b \<notin> Domain R"
+  by (metis 7 DomainE)
+have 9: "c = b"
+  by (metis 5 8)
+have 10: "f b = Intg x"
+  by (metis 0 9)
+show "False"
+  by (metis 10 3)
+qed
+
+declare [[ atp_problem_prefix = "TransClosure__test_simpler" ]]
+lemma "\<lbrakk> f c = Intg x; \<forall> y. f b = Intg y \<longrightarrow> y \<noteq> x; (a,b) \<in> R\<^sup>*; (b,c) \<in> R\<^sup>* \<rbrakk> 
+   \<Longrightarrow> \<exists> c. (b,c) \<in> R \<and> (a,c) \<in> R\<^sup>*"
+apply (erule_tac x="b" in converse_rtranclE)
+apply (metis rel_pow_0_E rel_pow_0_I)
+apply (metis DomainE Domain_iff Transitive_Closure.rtrancl_into_rtrancl)
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Metis_Examples/set.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,283 @@
+(*  Title:      HOL/Metis_Examples/set.thy
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+
+Testing the metis method.
+*)
+
+theory set
+imports Main
+begin
+
+lemma "EX x X. ALL y. EX z Z. (~P(y,y) | P(x,x) | ~S(z,x)) &
+               (S(x,y) | ~S(y,z) | Q(Z,Z))  &
+               (Q(X,y) | ~Q(y,Z) | S(X,X))" 
+by metis
+(*??But metis can't prove the single-step version...*)
+
+
+
+lemma "P(n::nat) ==> ~P(0) ==> n ~= 0"
+by metis
+
+declare [[sledgehammer_modulus = 1]]
+
+
+(*multiple versions of this example*)
+lemma (*equal_union: *)
+   "(X = Y \<union> Z) =
+    (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
+proof (neg_clausify)
+fix x
+assume 0: "Y \<subseteq> X \<or> X = Y \<union> Z"
+assume 1: "Z \<subseteq> X \<or> X = Y \<union> Z"
+assume 2: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Y \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
+assume 3: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Z \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
+assume 4: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> \<not> X \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
+assume 5: "\<And>V. ((\<not> Y \<subseteq> V \<or> \<not> Z \<subseteq> V) \<or> X \<subseteq> V) \<or> X = Y \<union> Z"
+have 6: "sup Y Z = X \<or> Y \<subseteq> X"
+  by (metis 0)
+have 7: "sup Y Z = X \<or> Z \<subseteq> X"
+  by (metis 1)
+have 8: "\<And>X3. sup Y Z = X \<or> X \<subseteq> X3 \<or> \<not> Y \<subseteq> X3 \<or> \<not> Z \<subseteq> X3"
+  by (metis 5)
+have 9: "Y \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X"
+  by (metis 2)
+have 10: "Z \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X"
+  by (metis 3)
+have 11: "sup Y Z \<noteq> X \<or> \<not> X \<subseteq> x \<or> \<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X"
+  by (metis 4)
+have 12: "Z \<subseteq> X"
+  by (metis Un_upper2 7)
+have 13: "\<And>X3. sup Y Z = X \<or> X \<subseteq> sup X3 Z \<or> \<not> Y \<subseteq> sup X3 Z"
+  by (metis 8 Un_upper2)
+have 14: "Y \<subseteq> X"
+  by (metis Un_upper1 6)
+have 15: "Z \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X"
+  by (metis 10 12)
+have 16: "Y \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X"
+  by (metis 9 12)
+have 17: "sup Y Z \<noteq> X \<or> \<not> X \<subseteq> x \<or> \<not> Y \<subseteq> X"
+  by (metis 11 12)
+have 18: "sup Y Z \<noteq> X \<or> \<not> X \<subseteq> x"
+  by (metis 17 14)
+have 19: "Z \<subseteq> x \<or> sup Y Z \<noteq> X"
+  by (metis 15 14)
+have 20: "Y \<subseteq> x \<or> sup Y Z \<noteq> X"
+  by (metis 16 14)
+have 21: "sup Y Z = X \<or> X \<subseteq> sup Y Z"
+  by (metis 13 Un_upper1)
+have 22: "sup Y Z = X \<or> \<not> sup Y Z \<subseteq> X"
+  by (metis equalityI 21)
+have 23: "sup Y Z = X \<or> \<not> Z \<subseteq> X \<or> \<not> Y \<subseteq> X"
+  by (metis 22 Un_least)
+have 24: "sup Y Z = X \<or> \<not> Y \<subseteq> X"
+  by (metis 23 12)
+have 25: "sup Y Z = X"
+  by (metis 24 14)
+have 26: "\<And>X3. X \<subseteq> X3 \<or> \<not> Z \<subseteq> X3 \<or> \<not> Y \<subseteq> X3"
+  by (metis Un_least 25)
+have 27: "Y \<subseteq> x"
+  by (metis 20 25)
+have 28: "Z \<subseteq> x"
+  by (metis 19 25)
+have 29: "\<not> X \<subseteq> x"
+  by (metis 18 25)
+have 30: "X \<subseteq> x \<or> \<not> Y \<subseteq> x"
+  by (metis 26 28)
+have 31: "X \<subseteq> x"
+  by (metis 30 27)
+show "False"
+  by (metis 31 29)
+qed
+
+declare [[sledgehammer_modulus = 2]]
+
+lemma (*equal_union: *)
+   "(X = Y \<union> Z) =
+    (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
+proof (neg_clausify)
+fix x
+assume 0: "Y \<subseteq> X \<or> X = Y \<union> Z"
+assume 1: "Z \<subseteq> X \<or> X = Y \<union> Z"
+assume 2: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Y \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
+assume 3: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Z \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
+assume 4: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> \<not> X \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
+assume 5: "\<And>V. ((\<not> Y \<subseteq> V \<or> \<not> Z \<subseteq> V) \<or> X \<subseteq> V) \<or> X = Y \<union> Z"
+have 6: "sup Y Z = X \<or> Y \<subseteq> X"
+  by (metis 0)
+have 7: "Y \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X"
+  by (metis 2)
+have 8: "sup Y Z \<noteq> X \<or> \<not> X \<subseteq> x \<or> \<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X"
+  by (metis 4)
+have 9: "\<And>X3. sup Y Z = X \<or> X \<subseteq> sup X3 Z \<or> \<not> Y \<subseteq> sup X3 Z"
+  by (metis 5 Un_upper2)
+have 10: "Z \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X"
+  by (metis 3 Un_upper2)
+have 11: "sup Y Z \<noteq> X \<or> \<not> X \<subseteq> x \<or> \<not> Y \<subseteq> X"
+  by (metis 8 Un_upper2)
+have 12: "Z \<subseteq> x \<or> sup Y Z \<noteq> X"
+  by (metis 10 Un_upper1)
+have 13: "sup Y Z = X \<or> X \<subseteq> sup Y Z"
+  by (metis 9 Un_upper1)
+have 14: "sup Y Z = X \<or> \<not> Z \<subseteq> X \<or> \<not> Y \<subseteq> X"
+  by (metis equalityI 13 Un_least)
+have 15: "sup Y Z = X"
+  by (metis 14 1 6)
+have 16: "Y \<subseteq> x"
+  by (metis 7 Un_upper2 Un_upper1 15)
+have 17: "\<not> X \<subseteq> x"
+  by (metis 11 Un_upper1 15)
+have 18: "X \<subseteq> x"
+  by (metis Un_least 15 12 15 16)
+show "False"
+  by (metis 18 17)
+qed
+
+declare [[sledgehammer_modulus = 3]]
+
+lemma (*equal_union: *)
+   "(X = Y \<union> Z) =
+    (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
+proof (neg_clausify)
+fix x
+assume 0: "Y \<subseteq> X \<or> X = Y \<union> Z"
+assume 1: "Z \<subseteq> X \<or> X = Y \<union> Z"
+assume 2: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Y \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
+assume 3: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Z \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
+assume 4: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> \<not> X \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
+assume 5: "\<And>V. ((\<not> Y \<subseteq> V \<or> \<not> Z \<subseteq> V) \<or> X \<subseteq> V) \<or> X = Y \<union> Z"
+have 6: "Z \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X"
+  by (metis 3)
+have 7: "\<And>X3. sup Y Z = X \<or> X \<subseteq> sup X3 Z \<or> \<not> Y \<subseteq> sup X3 Z"
+  by (metis 5 Un_upper2)
+have 8: "Y \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X"
+  by (metis 2 Un_upper2)
+have 9: "Z \<subseteq> x \<or> sup Y Z \<noteq> X"
+  by (metis 6 Un_upper2 Un_upper1)
+have 10: "sup Y Z = X \<or> \<not> sup Y Z \<subseteq> X"
+  by (metis equalityI 7 Un_upper1)
+have 11: "sup Y Z = X"
+  by (metis 10 Un_least 1 0)
+have 12: "Z \<subseteq> x"
+  by (metis 9 11)
+have 13: "X \<subseteq> x"
+  by (metis Un_least 11 12 8 Un_upper1 11)
+show "False"
+  by (metis 13 4 Un_upper2 Un_upper1 11)
+qed
+
+(*Example included in TPHOLs paper*)
+
+declare [[sledgehammer_modulus = 4]]
+
+lemma (*equal_union: *)
+   "(X = Y \<union> Z) =
+    (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
+proof (neg_clausify)
+fix x
+assume 0: "Y \<subseteq> X \<or> X = Y \<union> Z"
+assume 1: "Z \<subseteq> X \<or> X = Y \<union> Z"
+assume 2: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Y \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
+assume 3: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Z \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
+assume 4: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> \<not> X \<subseteq> x) \<or> X \<noteq> Y \<union> Z"
+assume 5: "\<And>V. ((\<not> Y \<subseteq> V \<or> \<not> Z \<subseteq> V) \<or> X \<subseteq> V) \<or> X = Y \<union> Z"
+have 6: "sup Y Z \<noteq> X \<or> \<not> X \<subseteq> x \<or> \<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X"
+  by (metis 4)
+have 7: "Z \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X"
+  by (metis 3 Un_upper2)
+have 8: "Z \<subseteq> x \<or> sup Y Z \<noteq> X"
+  by (metis 7 Un_upper1)
+have 9: "sup Y Z = X \<or> \<not> Z \<subseteq> X \<or> \<not> Y \<subseteq> X"
+  by (metis equalityI 5 Un_upper2 Un_upper1 Un_least)
+have 10: "Y \<subseteq> x"
+  by (metis 2 Un_upper2 1 Un_upper1 0 9 Un_upper2 1 Un_upper1 0)
+have 11: "X \<subseteq> x"
+  by (metis Un_least 9 Un_upper2 1 Un_upper1 0 8 9 Un_upper2 1 Un_upper1 0 10)
+show "False"
+  by (metis 11 6 Un_upper2 1 Un_upper1 0 9 Un_upper2 1 Un_upper1 0)
+qed 
+
+declare [[ atp_problem_prefix = "set__equal_union" ]]
+lemma (*equal_union: *)
+   "(X = Y \<union> Z) =
+    (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" 
+(*One shot proof: hand-reduced. Metis can't do the full proof any more.*)
+by (metis Un_least Un_upper1 Un_upper2 set_eq_subset)
+
+
+declare [[ atp_problem_prefix = "set__equal_inter" ]]
+lemma "(X = Y \<inter> Z) =
+    (X \<subseteq> Y \<and> X \<subseteq> Z \<and> (\<forall>V. V \<subseteq> Y \<and> V \<subseteq> Z \<longrightarrow> V \<subseteq> X))"
+by (metis Int_greatest Int_lower1 Int_lower2 set_eq_subset)
+
+declare [[ atp_problem_prefix = "set__fixedpoint" ]]
+lemma fixedpoint:
+    "\<exists>!x. f (g x) = x \<Longrightarrow> \<exists>!y. g (f y) = y"
+by metis
+
+lemma (*fixedpoint:*)
+    "\<exists>!x. f (g x) = x \<Longrightarrow> \<exists>!y. g (f y) = y"
+proof (neg_clausify)
+fix x xa
+assume 0: "f (g x) = x"
+assume 1: "\<And>y. y = x \<or> f (g y) \<noteq> y"
+assume 2: "\<And>x. g (f (xa x)) = xa x \<or> g (f x) \<noteq> x"
+assume 3: "\<And>x. g (f x) \<noteq> x \<or> xa x \<noteq> x"
+have 4: "\<And>X1. g (f X1) \<noteq> X1 \<or> g x \<noteq> X1"
+  by (metis 3 1 2)
+show "False"
+  by (metis 4 0)
+qed
+
+declare [[ atp_problem_prefix = "set__singleton_example" ]]
+lemma (*singleton_example_2:*)
+     "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
+by (metis Set.subsetI Union_upper insertCI set_eq_subset)
+  --{*found by SPASS*}
+
+lemma (*singleton_example_2:*)
+     "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
+by (metis Set.subsetI Union_upper insert_iff set_eq_subset)
+
+lemma singleton_example_2:
+     "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
+proof (neg_clausify)
+assume 0: "\<And>x. \<not> S \<subseteq> {x}"
+assume 1: "\<And>x. x \<notin> S \<or> \<Union>S \<subseteq> x"
+have 2: "\<And>X3. X3 = \<Union>S \<or> \<not> X3 \<subseteq> \<Union>S \<or> X3 \<notin> S"
+  by (metis set_eq_subset 1)
+have 3: "\<And>X3. S \<subseteq> insert (\<Union>S) X3"
+  by (metis insert_iff Set.subsetI Union_upper 2 Set.subsetI)
+show "False"
+  by (metis 3 0)
+qed
+
+
+
+text {*
+  From W. W. Bledsoe and Guohui Feng, SET-VAR. JAR 11 (3), 1993, pages
+  293-314.
+*}
+
+declare [[ atp_problem_prefix = "set__Bledsoe_Fung" ]]
+(*Notes: 1, the numbering doesn't completely agree with the paper. 
+2, we must rename set variables to avoid type clashes.*)
+lemma "\<exists>B. (\<forall>x \<in> B. x \<le> (0::int))"
+      "D \<in> F \<Longrightarrow> \<exists>G. \<forall>A \<in> G. \<exists>B \<in> F. A \<subseteq> B"
+      "P a \<Longrightarrow> \<exists>A. (\<forall>x \<in> A. P x) \<and> (\<exists>y. y \<in> A)"
+      "a < b \<and> b < (c::int) \<Longrightarrow> \<exists>B. a \<notin> B \<and> b \<in> B \<and> c \<notin> B"
+      "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A"
+      "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A"
+      "\<exists>A. a \<notin> A"
+      "(\<forall>C. (0, 0) \<in> C \<and> (\<forall>x y. (x, y) \<in> C \<longrightarrow> (Suc x, Suc y) \<in> C) \<longrightarrow> (n, m) \<in> C) \<and> Q n \<longrightarrow> Q m" 
+apply (metis atMost_iff)
+apply (metis emptyE)
+apply (metis insert_iff singletonE)
+apply (metis insertCI singletonE zless_le)
+apply (metis Collect_def Collect_mem_eq)
+apply (metis Collect_def Collect_mem_eq)
+apply (metis DiffE)
+apply (metis pair_in_Id_conv) 
+done
+
+end
--- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -279,9 +279,18 @@
 
 
 fun get_atp thy args =
-  AList.lookup (op =) args proverK
-  |> the_default (hd (ATP_Manager.get_atps ()))  (* FIXME partial *)
-  |> (fn name => (name, the (ATP_Manager.get_prover thy name)))  (* FIXME partial *)
+  let
+    fun default_atp_name () = hd (ATP_Manager.get_atps ())
+      handle Empty => error "No ATP available."
+    fun get_prover name =
+      (case ATP_Manager.get_prover thy name of
+        SOME prover => (name, prover)
+      | NONE => error ("Bad ATP: " ^ quote name))
+  in
+    (case AList.lookup (op =) args proverK of
+      SOME name => get_prover name
+    | NONE => get_prover (default_atp_name ()))
+  end
 
 local
 
--- a/src/HOL/Modelcheck/MuckeSyn.thy	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Modelcheck/MuckeSyn.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -123,8 +123,8 @@
                      REPEAT (resolve_tac prems 1)]);
 
   val sig_move_thm = Thm.theory_of_thm move_thm;
-  val bCterm = cterm_of sig_move_thm (valOf (search_var "b" (concl_of move_thm)));
-  val aCterm = cterm_of sig_move_thm (valOf (search_var "a" (hd(prems_of move_thm)))); 
+  val bCterm = cterm_of sig_move_thm (the (search_var "b" (concl_of move_thm)));
+  val aCterm = cterm_of sig_move_thm (the (search_var "a" (hd(prems_of move_thm)))); 
 
 in
 
--- a/src/HOL/Nominal/nominal_datatype.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Nominal/nominal_datatype.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -50,7 +50,7 @@
 
 fun dt_cases (descr: descr) (_, args, constrs) =
   let
-    fun the_bname i = Long_Name.base_name (#1 (valOf (AList.lookup (op =) descr i)));
+    fun the_bname i = Long_Name.base_name (#1 (the (AList.lookup (op =) descr i)));
     val bnames = map the_bname (distinct op = (maps dt_recs args));
   in map (fn (c, _) => space_implode "_" (Long_Name.base_name c :: bnames)) constrs end;
 
@@ -58,7 +58,7 @@
 fun induct_cases descr =
   DatatypeProp.indexify_names (maps (dt_cases descr) (map #2 descr));
 
-fun exhaust_cases descr i = dt_cases descr (valOf (AList.lookup (op =) descr i));
+fun exhaust_cases descr i = dt_cases descr (the (AList.lookup (op =) descr i));
 
 in
 
@@ -233,7 +233,7 @@
     fun replace_types (Type ("Nominal.ABS", [T, U])) =
           Type ("fun", [T, Type ("Nominal.noption", [replace_types U])])
       | replace_types (Type (s, Ts)) =
-          Type (getOpt (AList.lookup op = ps s, s), map replace_types Ts)
+          Type (the_default s (AList.lookup op = ps s), map replace_types Ts)
       | replace_types T = T;
 
     val dts'' = map (fn (tvs, tname, mx, constrs) => (tvs, Binding.name (tname ^ "_Rep"), NoSyn,
--- a/src/HOL/Nominal/nominal_fresh_fun.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Nominal/nominal_fresh_fun.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -7,24 +7,10 @@
 
 (* First some functions that should be in the library *)
 
-(* A tactical which applies a list of int -> tactic to the          *) 
-(* corresponding subgoals present after the application of          *) 
-(* another tactic.                                                  *)
-(*                                                                  *)
-(*  T THENL [A,B,C] is equivalent to T THEN (C 3 THEN B 2 THEN A 1) *) 
-
-infix 1 THENL
-fun tac THENL tacs =
- tac THEN
-  (EVERY (map  (fn (tac,i) => tac i) (rev tacs ~~ (length tacs downto 1))))
-
 (* A tactic which only succeeds when the argument *)
 (* tactic solves completely the specified subgoal *)
 fun SOLVEI t = t THEN_ALL_NEW (fn i => no_tac);
 
-(* A version of TRY for int -> tactic *)
-fun TRY' tac i =  TRY (tac i);
-
 fun gen_res_inst_tac_term instf tyinst tinst elim th i st =
   let
     val thy = theory_of_thm st;
@@ -45,10 +31,10 @@
     compose_tac (elim, th', nprems_of th) i st
   end handle Subscript => Seq.empty;
 
-val res_inst_tac_term = 
+val res_inst_tac_term =
   gen_res_inst_tac_term (curry Thm.instantiate);
 
-val res_inst_tac_term' = 
+val res_inst_tac_term' =
   gen_res_inst_tac_term (K Drule.cterm_instantiate) [];
 
 fun cut_inst_tac_term' tinst th =
@@ -65,18 +51,18 @@
 val fresh_fun_app'   = @{thm "fresh_fun_app'"};
 val fresh_prod       = @{thm "fresh_prod"};
 
-(* A tactic to generate a name fresh for  all the free *) 
+(* A tactic to generate a name fresh for  all the free *)
 (* variables and parameters of the goal                *)
 
 fun generate_fresh_tac atom_name i thm =
- let 
+ let
    val thy = theory_of_thm thm;
 (* the parsing function returns a qualified name, we get back the base name *)
    val atom_basename = Long_Name.base_name atom_name;
    val goal = List.nth(prems_of thm, i-1);
    val ps = Logic.strip_params goal;
    val Ts = rev (map snd ps);
-   fun is_of_fs_name T = Sign.of_sort thy (T, Sign.intern_sort thy ["fs_"^atom_basename]); 
+   fun is_of_fs_name T = Sign.of_sort thy (T, Sign.intern_sort thy ["fs_"^atom_basename]);
 (* rebuild de bruijn indices *)
    val bvs = map_index (Bound o fst) ps;
 (* select variables of the right class *)
@@ -90,7 +76,7 @@
    val exists_fresh' = at_name_inst_thm RS at_exists_fresh';
 (* find the variable we want to instantiate *)
    val x = hd (OldTerm.term_vars (prop_of exists_fresh'));
- in 
+ in
    (cut_inst_tac_term' [(x,s)] exists_fresh' 1 THEN
    rtac fs_name_thm 1 THEN
    etac exE 1) thm
@@ -98,18 +84,18 @@
   end;
 
 fun get_inner_fresh_fun (Bound j) = NONE
-  | get_inner_fresh_fun (v as Free _) = NONE 
+  | get_inner_fresh_fun (v as Free _) = NONE
   | get_inner_fresh_fun (v as Var _)  = NONE
   | get_inner_fresh_fun (Const _) = NONE
-  | get_inner_fresh_fun (Abs (_, _, t)) = get_inner_fresh_fun t 
-  | get_inner_fresh_fun (Const ("Nominal.fresh_fun",Type("fun",[Type ("fun",[Type (T,_),_]),_])) $ u) 
-                           = SOME T 
-  | get_inner_fresh_fun (t $ u) = 
+  | get_inner_fresh_fun (Abs (_, _, t)) = get_inner_fresh_fun t
+  | get_inner_fresh_fun (Const ("Nominal.fresh_fun",Type("fun",[Type ("fun",[Type (T,_),_]),_])) $ u)
+                           = SOME T
+  | get_inner_fresh_fun (t $ u) =
      let val a = get_inner_fresh_fun u in
-     if a = NONE then get_inner_fresh_fun t else a 
+     if a = NONE then get_inner_fresh_fun t else a
      end;
 
-(* This tactic generates a fresh name of the atom type *) 
+(* This tactic generates a fresh name of the atom type *)
 (* given by the innermost fresh_fun                    *)
 
 fun generate_fresh_fun_tac i thm =
@@ -117,32 +103,32 @@
     val goal = List.nth(prems_of thm, i-1);
     val atom_name_opt = get_inner_fresh_fun goal;
   in
-  case atom_name_opt of 
+  case atom_name_opt of
     NONE => all_tac thm
-  | SOME atom_name  => generate_fresh_tac atom_name i thm               
+  | SOME atom_name  => generate_fresh_tac atom_name i thm
   end
 
-(* Two substitution tactics which looks for the innermost occurence in 
+(* Two substitution tactics which looks for the innermost occurence in
    one assumption or in the conclusion *)
 
-val search_fun     = curry (Seq.flat o (uncurry EqSubst.searchf_bt_unify_valid));
+val search_fun = curry (Seq.flat o uncurry EqSubst.searchf_bt_unify_valid);
 val search_fun_asm = EqSubst.skip_first_asm_occs_search EqSubst.searchf_bt_unify_valid;
 
-fun subst_inner_tac           ctx = EqSubst.eqsubst_tac' ctx search_fun;
-fun subst_inner_asm_tac_aux i ctx = EqSubst.eqsubst_asm_tac' ctx search_fun_asm i;
+fun subst_inner_tac ctxt = EqSubst.eqsubst_tac' ctxt search_fun;
+fun subst_inner_asm_tac_aux i ctxt = EqSubst.eqsubst_asm_tac' ctxt search_fun_asm i;
 
-(* A tactic to substitute in the first assumption 
+(* A tactic to substitute in the first assumption
    which contains an occurence. *)
 
-fun subst_inner_asm_tac ctx th =  
-   curry (curry (FIRST' (map uncurry (map uncurry (map subst_inner_asm_tac_aux 
-                                                             (1 upto Thm.nprems_of th)))))) ctx th;
+fun subst_inner_asm_tac ctxt th =
+  curry (curry (FIRST' (map uncurry (map uncurry (map subst_inner_asm_tac_aux
+            (1 upto Thm.nprems_of th)))))) ctxt th;
 
-fun fresh_fun_tac no_asm i thm = 
+fun fresh_fun_tac no_asm i thm =
   (* Find the variable we instantiate *)
   let
     val thy = theory_of_thm thm;
-    val ctx = Context.init_proof thy;
+    val ctxt = ProofContext.init thy;
     val ss = global_simpset_of thy;
     val abs_fresh = PureThy.get_thms thy "abs_fresh";
     val fresh_perm_app = PureThy.get_thms thy "fresh_perm_app";
@@ -151,45 +137,45 @@
     val x = hd (tl (OldTerm.term_vars (prop_of exI)));
     val goal = nth (prems_of thm) (i-1);
     val atom_name_opt = get_inner_fresh_fun goal;
-    val n = List.length (Logic.strip_params goal);
+    val n = length (Logic.strip_params goal);
     (* Here we rely on the fact that the variable introduced by generate_fresh_tac *)
     (* is the last one in the list, the inner one *)
   in
-  case atom_name_opt of 
+  case atom_name_opt of
     NONE => all_tac thm
-  | SOME atom_name  =>    
-  let 
+  | SOME atom_name =>
+  let
     val atom_basename = Long_Name.base_name atom_name;
     val pt_name_inst = get_dyn_thm thy ("pt_"^atom_basename^"_inst") atom_basename;
     val at_name_inst = get_dyn_thm thy ("at_"^atom_basename^"_inst") atom_basename;
     fun inst_fresh vars params i st =
    let val vars' = OldTerm.term_vars (prop_of st);
        val thy = theory_of_thm st;
-   in case vars' \\ vars of 
+   in case vars' \\ vars of
      [x] => Seq.single (Thm.instantiate ([],[(cterm_of thy x,cterm_of thy (list_abs (params,Bound 0)))]) st)
     | _ => error "fresh_fun_simp: Too many variables, please report."
   end
   in
   ((fn st =>
-  let 
+  let
     val vars = OldTerm.term_vars (prop_of st);
     val params = Logic.strip_params (nth (prems_of st) (i-1))
-    (* The tactics which solve the subgoals generated 
+    (* The tactics which solve the subgoals generated
        by the conditionnal rewrite rule. *)
-    val post_rewrite_tacs =  
+    val post_rewrite_tacs =
           [rtac pt_name_inst,
            rtac at_name_inst,
-           TRY' (SOLVEI (NominalPermeq.finite_guess_tac ss'')),
+           TRY o SOLVEI (NominalPermeq.finite_guess_tac ss''),
            inst_fresh vars params THEN'
-           (TRY' (SOLVEI (NominalPermeq.fresh_guess_tac ss''))) THEN'
-           (TRY' (SOLVEI (asm_full_simp_tac ss'')))] 
-  in 
+           (TRY o SOLVEI (NominalPermeq.fresh_guess_tac ss'')) THEN'
+           (TRY o SOLVEI (asm_full_simp_tac ss''))]
+  in
    ((if no_asm then no_tac else
-    (subst_inner_asm_tac ctx fresh_fun_app' i THENL post_rewrite_tacs)) 
+    (subst_inner_asm_tac ctxt fresh_fun_app' i THEN (RANGE post_rewrite_tacs i)))
     ORELSE
-    (subst_inner_tac     ctx fresh_fun_app' i THENL post_rewrite_tacs)) st
+    (subst_inner_tac ctxt fresh_fun_app' i THEN (RANGE post_rewrite_tacs i))) st
   end)) thm
-  
+
   end
   end
 
--- a/src/HOL/README.html	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/README.html	Wed Oct 21 08:16:25 2009 +0200
@@ -60,7 +60,7 @@
 <dt>IOA
 <dd>a simple theory of Input/Output Automata
 
-<dt>Isar_examples
+<dt>Isar_Examples
 <dd>several introductory examples using Isabelle/Isar
 
 <dt>Lambda
@@ -99,7 +99,7 @@
 <dt>Hahn_Banach
 <dd>the Hahn-Banach theorem for real vector spaces (in Isabelle/Isar)
 
-<dt>SET-Protocol
+<dt>SET_Protocol
 <dd>verification of the SET Protocol
 
 <dt>Subst
--- a/src/HOL/SET-Protocol/Cardholder_Registration.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1054 +0,0 @@
-(*  Title:      HOL/SET-Protocol/Cardholder_Registration.thy
-    Author:     Giampaolo Bella
-    Author:     Fabio Massacci
-    Author:     Lawrence C Paulson
-    Author:     Piero Tramontano
-*)
-
-header{*The SET Cardholder Registration Protocol*}
-
-theory Cardholder_Registration imports PublicSET begin
-
-text{*Note: nonces seem to consist of 20 bytes.  That includes both freshness
-challenges (Chall-EE, etc.) and important secrets (CardSecret, PANsecret)
-*}
-
-text{*Simplifications involving @{text analz_image_keys_simps} appear to
-have become much slower. The cause is unclear. However, there is a big blow-up
-and the rewriting is very sensitive to the set of rewrite rules given.*}
-
-subsection{*Predicate Formalizing the Encryption Association between Keys *}
-
-consts
-  KeyCryptKey :: "[key, key, event list] => bool"
-
-primrec
-
-KeyCryptKey_Nil:
-  "KeyCryptKey DK K [] = False"
-
-KeyCryptKey_Cons:
-      --{*Says is the only important case.
-        1st case: CR5, where KC3 encrypts KC2.
-        2nd case: any use of priEK C.
-        Revision 1.12 has a more complicated version with separate treatment of
-          the dependency of KC1, KC2 and KC3 on priEK (CA i.)  Not needed since
-          priEK C is never sent (and so can't be lost except at the start). *}
-  "KeyCryptKey DK K (ev # evs) =
-   (KeyCryptKey DK K evs |
-    (case ev of
-      Says A B Z =>
-       ((\<exists>N X Y. A \<noteq> Spy &
-                 DK \<in> symKeys &
-                 Z = {|Crypt DK {|Agent A, Nonce N, Key K, X|}, Y|}) |
-        (\<exists>C. DK = priEK C))
-    | Gets A' X => False
-    | Notes A' X => False))"
-
-
-subsection{*Predicate formalizing the association between keys and nonces *}
-
-consts
-  KeyCryptNonce :: "[key, key, event list] => bool"
-
-primrec
-
-KeyCryptNonce_Nil:
-  "KeyCryptNonce EK K [] = False"
-
-KeyCryptNonce_Cons:
-  --{*Says is the only important case.
-    1st case: CR3, where KC1 encrypts NC2 (distinct from CR5 due to EXH);
-    2nd case: CR5, where KC3 encrypts NC3;
-    3rd case: CR6, where KC2 encrypts NC3;
-    4th case: CR6, where KC2 encrypts NonceCCA;
-    5th case: any use of @{term "priEK C"} (including CardSecret).
-    NB the only Nonces we need to keep secret are CardSecret and NonceCCA.
-    But we can't prove @{text Nonce_compromise} unless the relation covers ALL
-        nonces that the protocol keeps secret.
-  *}
-  "KeyCryptNonce DK N (ev # evs) =
-   (KeyCryptNonce DK N evs |
-    (case ev of
-      Says A B Z =>
-       A \<noteq> Spy &
-       ((\<exists>X Y. DK \<in> symKeys &
-               Z = (EXHcrypt DK X {|Agent A, Nonce N|} Y)) |
-        (\<exists>X Y. DK \<in> symKeys &
-               Z = {|Crypt DK {|Agent A, Nonce N, X|}, Y|}) |
-        (\<exists>K i X Y.
-          K \<in> symKeys &
-          Z = Crypt K {|sign (priSK (CA i)) {|Agent B, Nonce N, X|}, Y|} &
-          (DK=K | KeyCryptKey DK K evs)) |
-        (\<exists>K C NC3 Y.
-          K \<in> symKeys &
-          Z = Crypt K
-                {|sign (priSK C) {|Agent B, Nonce NC3, Agent C, Nonce N|},
-                  Y|} &
-          (DK=K | KeyCryptKey DK K evs)) |
-        (\<exists>C. DK = priEK C))
-    | Gets A' X => False
-    | Notes A' X => False))"
-
-
-subsection{*Formal protocol definition *}
-
-inductive_set
-  set_cr :: "event list set"
-where
-
-  Nil:    --{*Initial trace is empty*}
-          "[] \<in> set_cr"
-
-| Fake:    --{*The spy MAY say anything he CAN say.*}
-           "[| evsf \<in> set_cr; X \<in> synth (analz (knows Spy evsf)) |]
-            ==> Says Spy B X  # evsf \<in> set_cr"
-
-| Reception: --{*If A sends a message X to B, then B might receive it*}
-             "[| evsr \<in> set_cr; Says A B X \<in> set evsr |]
-              ==> Gets B X  # evsr \<in> set_cr"
-
-| SET_CR1: --{*CardCInitReq: C initiates a run, sending a nonce to CCA*}
-             "[| evs1 \<in> set_cr;  C = Cardholder k;  Nonce NC1 \<notin> used evs1 |]
-              ==> Says C (CA i) {|Agent C, Nonce NC1|} # evs1 \<in> set_cr"
-
-| SET_CR2: --{*CardCInitRes: CA responds sending NC1 and its certificates*}
-             "[| evs2 \<in> set_cr;
-                 Gets (CA i) {|Agent C, Nonce NC1|} \<in> set evs2 |]
-              ==> Says (CA i) C
-                       {|sign (priSK (CA i)) {|Agent C, Nonce NC1|},
-                         cert (CA i) (pubEK (CA i)) onlyEnc (priSK RCA),
-                         cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|}
-                    # evs2 \<in> set_cr"
-
-| SET_CR3:
-   --{*RegFormReq: C sends his PAN and a new nonce to CA.
-   C verifies that
-    - nonce received is the same as that sent;
-    - certificates are signed by RCA;
-    - certificates are an encryption certificate (flag is onlyEnc) and a
-      signature certificate (flag is onlySig);
-    - certificates pertain to the CA that C contacted (this is done by
-      checking the signature).
-   C generates a fresh symmetric key KC1.
-   The point of encrypting @{term "{|Agent C, Nonce NC2, Hash (Pan(pan C))|}"}
-   is not clear. *}
-"[| evs3 \<in> set_cr;  C = Cardholder k;
-    Nonce NC2 \<notin> used evs3;
-    Key KC1 \<notin> used evs3; KC1 \<in> symKeys;
-    Gets C {|sign (invKey SKi) {|Agent X, Nonce NC1|},
-             cert (CA i) EKi onlyEnc (priSK RCA),
-             cert (CA i) SKi onlySig (priSK RCA)|}
-       \<in> set evs3;
-    Says C (CA i) {|Agent C, Nonce NC1|} \<in> set evs3|]
- ==> Says C (CA i) (EXHcrypt KC1 EKi {|Agent C, Nonce NC2|} (Pan(pan C)))
-       # Notes C {|Key KC1, Agent (CA i)|}
-       # evs3 \<in> set_cr"
-
-| SET_CR4:
-    --{*RegFormRes:
-    CA responds sending NC2 back with a new nonce NCA, after checking that
-     - the digital envelope is correctly encrypted by @{term "pubEK (CA i)"}
-     - the entire message is encrypted with the same key found inside the
-       envelope (here, KC1) *}
-"[| evs4 \<in> set_cr;
-    Nonce NCA \<notin> used evs4;  KC1 \<in> symKeys;
-    Gets (CA i) (EXHcrypt KC1 EKi {|Agent C, Nonce NC2|} (Pan(pan X)))
-       \<in> set evs4 |]
-  ==> Says (CA i) C
-          {|sign (priSK (CA i)) {|Agent C, Nonce NC2, Nonce NCA|},
-            cert (CA i) (pubEK (CA i)) onlyEnc (priSK RCA),
-            cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|}
-       # evs4 \<in> set_cr"
-
-| SET_CR5:
-   --{*CertReq: C sends his PAN, a new nonce, its proposed public signature key
-       and its half of the secret value to CA.
-       We now assume that C has a fixed key pair, and he submits (pubSK C).
-       The protocol does not require this key to be fresh.
-       The encryption below is actually EncX.*}
-"[| evs5 \<in> set_cr;  C = Cardholder k;
-    Nonce NC3 \<notin> used evs5;  Nonce CardSecret \<notin> used evs5; NC3\<noteq>CardSecret;
-    Key KC2 \<notin> used evs5; KC2 \<in> symKeys;
-    Key KC3 \<notin> used evs5; KC3 \<in> symKeys; KC2\<noteq>KC3;
-    Gets C {|sign (invKey SKi) {|Agent C, Nonce NC2, Nonce NCA|},
-             cert (CA i) EKi onlyEnc (priSK RCA),
-             cert (CA i) SKi onlySig (priSK RCA) |}
-        \<in> set evs5;
-    Says C (CA i) (EXHcrypt KC1 EKi {|Agent C, Nonce NC2|} (Pan(pan C)))
-         \<in> set evs5 |]
-==> Says C (CA i)
-         {|Crypt KC3
-             {|Agent C, Nonce NC3, Key KC2, Key (pubSK C),
-               Crypt (priSK C)
-                 (Hash {|Agent C, Nonce NC3, Key KC2,
-                         Key (pubSK C), Pan (pan C), Nonce CardSecret|})|},
-           Crypt EKi {|Key KC3, Pan (pan C), Nonce CardSecret|} |}
-    # Notes C {|Key KC2, Agent (CA i)|}
-    # Notes C {|Key KC3, Agent (CA i)|}
-    # evs5 \<in> set_cr"
-
-
-  --{* CertRes: CA responds sending NC3 back with its half of the secret value,
-   its signature certificate and the new cardholder signature
-   certificate.  CA checks to have never certified the key proposed by C.
-   NOTE: In Merchant Registration, the corresponding rule (4)
-   uses the "sign" primitive. The encryption below is actually @{term EncK}, 
-   which is just @{term "Crypt K (sign SK X)"}.
-*}
-
-| SET_CR6:
-"[| evs6 \<in> set_cr;
-    Nonce NonceCCA \<notin> used evs6;
-    KC2 \<in> symKeys;  KC3 \<in> symKeys;  cardSK \<notin> symKeys;
-    Notes (CA i) (Key cardSK) \<notin> set evs6;
-    Gets (CA i)
-      {|Crypt KC3 {|Agent C, Nonce NC3, Key KC2, Key cardSK,
-                    Crypt (invKey cardSK)
-                      (Hash {|Agent C, Nonce NC3, Key KC2,
-                              Key cardSK, Pan (pan C), Nonce CardSecret|})|},
-        Crypt (pubEK (CA i)) {|Key KC3, Pan (pan C), Nonce CardSecret|} |}
-      \<in> set evs6 |]
-==> Says (CA i) C
-         (Crypt KC2
-          {|sign (priSK (CA i))
-                 {|Agent C, Nonce NC3, Agent(CA i), Nonce NonceCCA|},
-            certC (pan C) cardSK (XOR(CardSecret,NonceCCA)) onlySig (priSK (CA i)),
-            cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|})
-      # Notes (CA i) (Key cardSK)
-      # evs6 \<in> set_cr"
-
-
-declare Says_imp_knows_Spy [THEN parts.Inj, dest]
-declare parts.Body [dest]
-declare analz_into_parts [dest]
-declare Fake_parts_insert_in_Un [dest]
-
-text{*A "possibility property": there are traces that reach the end.
-      An unconstrained proof with many subgoals.*}
-
-lemma Says_to_Gets:
-     "Says A B X # evs \<in> set_cr ==> Gets B X # Says A B X # evs \<in> set_cr"
-by (rule set_cr.Reception, auto)
-
-text{*The many nonces and keys generated, some simultaneously, force us to
-  introduce them explicitly as shown below.*}
-lemma possibility_CR6:
-     "[|NC1 < (NC2::nat);  NC2 < NC3;  NC3 < NCA ;
-        NCA < NonceCCA;  NonceCCA < CardSecret;
-        KC1 < (KC2::key);  KC2 < KC3;
-        KC1 \<in> symKeys;  Key KC1 \<notin> used [];
-        KC2 \<in> symKeys;  Key KC2 \<notin> used [];
-        KC3 \<in> symKeys;  Key KC3 \<notin> used [];
-        C = Cardholder k|]
-   ==> \<exists>evs \<in> set_cr.
-       Says (CA i) C
-            (Crypt KC2
-             {|sign (priSK (CA i))
-                    {|Agent C, Nonce NC3, Agent(CA i), Nonce NonceCCA|},
-               certC (pan C) (pubSK (Cardholder k)) (XOR(CardSecret,NonceCCA))
-                     onlySig (priSK (CA i)),
-               cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|})
-          \<in> set evs"
-apply (intro exI bexI)
-apply (rule_tac [2] 
-       set_cr.Nil 
-        [THEN set_cr.SET_CR1 [of concl: C i NC1], 
-         THEN Says_to_Gets, 
-         THEN set_cr.SET_CR2 [of concl: i C NC1], 
-         THEN Says_to_Gets,  
-         THEN set_cr.SET_CR3 [of concl: C i KC1 _ NC2], 
-         THEN Says_to_Gets,  
-         THEN set_cr.SET_CR4 [of concl: i C NC2 NCA], 
-         THEN Says_to_Gets,  
-         THEN set_cr.SET_CR5 [of concl: C i KC3 NC3 KC2 CardSecret],
-         THEN Says_to_Gets,  
-         THEN set_cr.SET_CR6 [of concl: i C KC2]])
-apply basic_possibility
-apply (simp_all (no_asm_simp) add: symKeys_neq_imp_neq)
-done
-
-text{*General facts about message reception*}
-lemma Gets_imp_Says:
-     "[| Gets B X \<in> set evs; evs \<in> set_cr |] ==> \<exists>A. Says A B X \<in> set evs"
-apply (erule rev_mp)
-apply (erule set_cr.induct, auto)
-done
-
-lemma Gets_imp_knows_Spy:
-     "[| Gets B X \<in> set evs; evs \<in> set_cr |]  ==> X \<in> knows Spy evs"
-by (blast dest!: Gets_imp_Says Says_imp_knows_Spy)
-declare Gets_imp_knows_Spy [THEN parts.Inj, dest]
-
-
-subsection{*Proofs on keys *}
-
-text{*Spy never sees an agent's private keys! (unless it's bad at start)*}
-
-lemma Spy_see_private_Key [simp]:
-     "evs \<in> set_cr
-      ==> (Key(invKey (publicKey b A)) \<in> parts(knows Spy evs)) = (A \<in> bad)"
-by (erule set_cr.induct, auto)
-
-lemma Spy_analz_private_Key [simp]:
-     "evs \<in> set_cr ==>
-     (Key(invKey (publicKey b A)) \<in> analz(knows Spy evs)) = (A \<in> bad)"
-by auto
-
-declare Spy_see_private_Key [THEN [2] rev_iffD1, dest!]
-declare Spy_analz_private_Key [THEN [2] rev_iffD1, dest!]
-
-
-subsection{*Begin Piero's Theorems on Certificates*}
-text{*Trivial in the current model, where certificates by RCA are secure *}
-
-lemma Crypt_valid_pubEK:
-     "[| Crypt (priSK RCA) {|Agent C, Key EKi, onlyEnc|}
-           \<in> parts (knows Spy evs);
-         evs \<in> set_cr |] ==> EKi = pubEK C"
-apply (erule rev_mp)
-apply (erule set_cr.induct, auto)
-done
-
-lemma certificate_valid_pubEK:
-    "[| cert C EKi onlyEnc (priSK RCA) \<in> parts (knows Spy evs);
-        evs \<in> set_cr |]
-     ==> EKi = pubEK C"
-apply (unfold cert_def signCert_def)
-apply (blast dest!: Crypt_valid_pubEK)
-done
-
-lemma Crypt_valid_pubSK:
-     "[| Crypt (priSK RCA) {|Agent C, Key SKi, onlySig|}
-           \<in> parts (knows Spy evs);
-         evs \<in> set_cr |] ==> SKi = pubSK C"
-apply (erule rev_mp)
-apply (erule set_cr.induct, auto)
-done
-
-lemma certificate_valid_pubSK:
-    "[| cert C SKi onlySig (priSK RCA) \<in> parts (knows Spy evs);
-        evs \<in> set_cr |] ==> SKi = pubSK C"
-apply (unfold cert_def signCert_def)
-apply (blast dest!: Crypt_valid_pubSK)
-done
-
-lemma Gets_certificate_valid:
-     "[| Gets A {| X, cert C EKi onlyEnc (priSK RCA),
-                      cert C SKi onlySig (priSK RCA)|} \<in> set evs;
-         evs \<in> set_cr |]
-      ==> EKi = pubEK C & SKi = pubSK C"
-by (blast dest: certificate_valid_pubEK certificate_valid_pubSK)
-
-text{*Nobody can have used non-existent keys!*}
-lemma new_keys_not_used:
-     "[|K \<in> symKeys; Key K \<notin> used evs; evs \<in> set_cr|]
-      ==> K \<notin> keysFor (parts (knows Spy evs))"
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule set_cr.induct)
-apply (frule_tac [8] Gets_certificate_valid)
-apply (frule_tac [6] Gets_certificate_valid, simp_all)
-apply (force dest!: usedI keysFor_parts_insert) --{*Fake*}
-apply (blast,auto)  --{*Others*}
-done
-
-
-subsection{*New versions: as above, but generalized to have the KK argument *}
-
-lemma gen_new_keys_not_used:
-     "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_cr |]
-      ==> Key K \<notin> used evs --> K \<in> symKeys -->
-          K \<notin> keysFor (parts (Key`KK Un knows Spy evs))"
-by (auto simp add: new_keys_not_used)
-
-lemma gen_new_keys_not_analzd:
-     "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_cr |]
-      ==> K \<notin> keysFor (analz (Key`KK Un knows Spy evs))"
-by (blast intro: keysFor_mono [THEN [2] rev_subsetD]
-          dest: gen_new_keys_not_used)
-
-lemma analz_Key_image_insert_eq:
-     "[|K \<in> symKeys; Key K \<notin> used evs; evs \<in> set_cr |]
-      ==> analz (Key ` (insert K KK) \<union> knows Spy evs) =
-          insert (Key K) (analz (Key ` KK \<union> knows Spy evs))"
-by (simp add: gen_new_keys_not_analzd)
-
-lemma Crypt_parts_imp_used:
-     "[|Crypt K X \<in> parts (knows Spy evs);
-        K \<in> symKeys; evs \<in> set_cr |] ==> Key K \<in> used evs"
-apply (rule ccontr)
-apply (force dest: new_keys_not_used Crypt_imp_invKey_keysFor)
-done
-
-lemma Crypt_analz_imp_used:
-     "[|Crypt K X \<in> analz (knows Spy evs);
-        K \<in> symKeys; evs \<in> set_cr |] ==> Key K \<in> used evs"
-by (blast intro: Crypt_parts_imp_used)
-
-
-(*<*) 
-subsection{*Messages signed by CA*}
-
-text{*Message @{text SET_CR2}: C can check CA's signature if he has received
-     CA's certificate.*}
-lemma CA_Says_2_lemma:
-     "[| Crypt (priSK (CA i)) (Hash{|Agent C, Nonce NC1|})
-           \<in> parts (knows Spy evs);
-         evs \<in> set_cr; (CA i) \<notin> bad |]
-     ==> \<exists>Y. Says (CA i) C {|sign (priSK (CA i)) {|Agent C, Nonce NC1|}, Y|}
-                 \<in> set evs"
-apply (erule rev_mp)
-apply (erule set_cr.induct, auto)
-done
-
-text{*Ever used?*}
-lemma CA_Says_2:
-     "[| Crypt (invKey SK) (Hash{|Agent C, Nonce NC1|})
-           \<in> parts (knows Spy evs);
-         cert (CA i) SK onlySig (priSK RCA) \<in> parts (knows Spy evs);
-         evs \<in> set_cr; (CA i) \<notin> bad |]
-      ==> \<exists>Y. Says (CA i) C {|sign (priSK (CA i)) {|Agent C, Nonce NC1|}, Y|}
-                  \<in> set evs"
-by (blast dest!: certificate_valid_pubSK intro!: CA_Says_2_lemma)
-
-
-text{*Message @{text SET_CR4}: C can check CA's signature if he has received
-      CA's certificate.*}
-lemma CA_Says_4_lemma:
-     "[| Crypt (priSK (CA i)) (Hash{|Agent C, Nonce NC2, Nonce NCA|})
-           \<in> parts (knows Spy evs);
-         evs \<in> set_cr; (CA i) \<notin> bad |]
-      ==> \<exists>Y. Says (CA i) C {|sign (priSK (CA i))
-                     {|Agent C, Nonce NC2, Nonce NCA|}, Y|} \<in> set evs"
-apply (erule rev_mp)
-apply (erule set_cr.induct, auto)
-done
-
-text{*NEVER USED*}
-lemma CA_Says_4:
-     "[| Crypt (invKey SK) (Hash{|Agent C, Nonce NC2, Nonce NCA|})
-           \<in> parts (knows Spy evs);
-         cert (CA i) SK onlySig (priSK RCA) \<in> parts (knows Spy evs);
-         evs \<in> set_cr; (CA i) \<notin> bad |]
-      ==> \<exists>Y. Says (CA i) C {|sign (priSK (CA i))
-                   {|Agent C, Nonce NC2, Nonce NCA|}, Y|} \<in> set evs"
-by (blast dest!: certificate_valid_pubSK intro!: CA_Says_4_lemma)
-
-
-text{*Message @{text SET_CR6}: C can check CA's signature if he has
-      received CA's certificate.*}
-lemma CA_Says_6_lemma:
-     "[| Crypt (priSK (CA i)) 
-               (Hash{|Agent C, Nonce NC3, Agent (CA i), Nonce NonceCCA|})
-           \<in> parts (knows Spy evs);
-         evs \<in> set_cr; (CA i) \<notin> bad |]
-      ==> \<exists>Y K. Says (CA i) C (Crypt K {|sign (priSK (CA i))
-      {|Agent C, Nonce NC3, Agent (CA i), Nonce NonceCCA|}, Y|}) \<in> set evs"
-apply (erule rev_mp)
-apply (erule set_cr.induct, auto)
-done
-
-text{*NEVER USED*}
-lemma CA_Says_6:
-     "[| Crypt (invKey SK) (Hash{|Agent C, Nonce NC3, Agent (CA i), Nonce NonceCCA|})
-           \<in> parts (knows Spy evs);
-         cert (CA i) SK onlySig (priSK RCA) \<in> parts (knows Spy evs);
-         evs \<in> set_cr; (CA i) \<notin> bad |]
-      ==> \<exists>Y K. Says (CA i) C (Crypt K {|sign (priSK (CA i))
-                    {|Agent C, Nonce NC3, Agent (CA i), Nonce NonceCCA|}, Y|}) \<in> set evs"
-by (blast dest!: certificate_valid_pubSK intro!: CA_Says_6_lemma)
-(*>*)
-
-
-subsection{*Useful lemmas *}
-
-text{*Rewriting rule for private encryption keys.  Analogous rewriting rules
-for other keys aren't needed.*}
-
-lemma parts_image_priEK:
-     "[|Key (priEK C) \<in> parts (Key`KK Un (knows Spy evs));
-        evs \<in> set_cr|] ==> priEK C \<in> KK | C \<in> bad"
-by auto
-
-text{*trivial proof because (priEK C) never appears even in (parts evs)*}
-lemma analz_image_priEK:
-     "evs \<in> set_cr ==>
-          (Key (priEK C) \<in> analz (Key`KK Un (knows Spy evs))) =
-          (priEK C \<in> KK | C \<in> bad)"
-by (blast dest!: parts_image_priEK intro: analz_mono [THEN [2] rev_subsetD])
-
-
-subsection{*Secrecy of Session Keys *}
-
-subsubsection{*Lemmas about the predicate KeyCryptKey *}
-
-text{*A fresh DK cannot be associated with any other
-  (with respect to a given trace). *}
-lemma DK_fresh_not_KeyCryptKey:
-     "[| Key DK \<notin> used evs; evs \<in> set_cr |] ==> ~ KeyCryptKey DK K evs"
-apply (erule rev_mp)
-apply (erule set_cr.induct)
-apply (simp_all (no_asm_simp))
-apply (blast dest: Crypt_analz_imp_used)+
-done
-
-text{*A fresh K cannot be associated with any other.  The assumption that
-  DK isn't a private encryption key may be an artifact of the particular
-  definition of KeyCryptKey.*}
-lemma K_fresh_not_KeyCryptKey:
-     "[|\<forall>C. DK \<noteq> priEK C; Key K \<notin> used evs|] ==> ~ KeyCryptKey DK K evs"
-apply (induct evs)
-apply (auto simp add: parts_insert2 split add: event.split)
-done
-
-
-text{*This holds because if (priEK (CA i)) appears in any traffic then it must
-  be known to the Spy, by @{term Spy_see_private_Key}*}
-lemma cardSK_neq_priEK:
-     "[|Key cardSK \<notin> analz (knows Spy evs);
-        Key cardSK : parts (knows Spy evs);
-        evs \<in> set_cr|] ==> cardSK \<noteq> priEK C"
-by blast
-
-lemma not_KeyCryptKey_cardSK [rule_format (no_asm)]:
-     "[|cardSK \<notin> symKeys;  \<forall>C. cardSK \<noteq> priEK C;  evs \<in> set_cr|] ==>
-      Key cardSK \<notin> analz (knows Spy evs) --> ~ KeyCryptKey cardSK K evs"
-by (erule set_cr.induct, analz_mono_contra, auto)
-
-text{*Lemma for message 5: pubSK C is never used to encrypt Keys.*}
-lemma pubSK_not_KeyCryptKey [simp]: "~ KeyCryptKey (pubSK C) K evs"
-apply (induct_tac "evs")
-apply (auto simp add: parts_insert2 split add: event.split)
-done
-
-text{*Lemma for message 6: either cardSK is compromised (when we don't care)
-  or else cardSK hasn't been used to encrypt K.  Previously we treated
-  message 5 in the same way, but the current model assumes that rule
-  @{text SET_CR5} is executed only by honest agents.*}
-lemma msg6_KeyCryptKey_disj:
-     "[|Gets B {|Crypt KC3 {|Agent C, Nonce N, Key KC2, Key cardSK, X|}, Y|}
-          \<in> set evs;
-        cardSK \<notin> symKeys;  evs \<in> set_cr|]
-      ==> Key cardSK \<in> analz (knows Spy evs) |
-          (\<forall>K. ~ KeyCryptKey cardSK K evs)"
-by (blast dest: not_KeyCryptKey_cardSK intro: cardSK_neq_priEK)
-
-text{*As usual: we express the property as a logical equivalence*}
-lemma Key_analz_image_Key_lemma:
-     "P --> (Key K \<in> analz (Key`KK Un H)) --> (K \<in> KK | Key K \<in> analz H)
-      ==>
-      P --> (Key K \<in> analz (Key`KK Un H)) = (K \<in> KK | Key K \<in> analz H)"
-by (blast intro: analz_mono [THEN [2] rev_subsetD])
-
-method_setup valid_certificate_tac = {*
-  Args.goal_spec >> (fn quant => K (SIMPLE_METHOD'' quant
-    (fn i =>
-      EVERY [ftac @{thm Gets_certificate_valid} i,
-             assume_tac i,
-             etac conjE i, REPEAT (hyp_subst_tac i)])))
-*} ""
-
-text{*The @{text "(no_asm)"} attribute is essential, since it retains
-  the quantifier and allows the simprule's condition to itself be simplified.*}
-lemma symKey_compromise [rule_format (no_asm)]:
-     "evs \<in> set_cr ==>
-      (\<forall>SK KK. SK \<in> symKeys \<longrightarrow> (\<forall>K \<in> KK. ~ KeyCryptKey K SK evs)   -->
-               (Key SK \<in> analz (Key`KK Un (knows Spy evs))) =
-               (SK \<in> KK | Key SK \<in> analz (knows Spy evs)))"
-apply (erule set_cr.induct)
-apply (rule_tac [!] allI) +
-apply (rule_tac [!] impI [THEN Key_analz_image_Key_lemma, THEN impI])+
-apply (valid_certificate_tac [8]) --{*for message 5*}
-apply (valid_certificate_tac [6]) --{*for message 5*}
-apply (erule_tac [9] msg6_KeyCryptKey_disj [THEN disjE])
-apply (simp_all
-         del: image_insert image_Un imp_disjL
-         add: analz_image_keys_simps analz_knows_absorb
-              analz_Key_image_insert_eq notin_image_iff
-              K_fresh_not_KeyCryptKey
-              DK_fresh_not_KeyCryptKey ball_conj_distrib
-              analz_image_priEK disj_simps)
-  --{*9 seconds on a 1.6GHz machine*}
-apply spy_analz
-apply blast  --{*3*}
-apply blast  --{*5*}
-done
-
-text{*The remaining quantifiers seem to be essential.
-  NO NEED to assume the cardholder's OK: bad cardholders don't do anything
-  wrong!!*}
-lemma symKey_secrecy [rule_format]:
-     "[|CA i \<notin> bad;  K \<in> symKeys;  evs \<in> set_cr|]
-      ==> \<forall>X c. Says (Cardholder c) (CA i) X \<in> set evs -->
-                Key K \<in> parts{X} -->
-                Cardholder c \<notin> bad -->
-                Key K \<notin> analz (knows Spy evs)"
-apply (erule set_cr.induct)
-apply (frule_tac [8] Gets_certificate_valid) --{*for message 5*}
-apply (frule_tac [6] Gets_certificate_valid) --{*for message 3*}
-apply (erule_tac [11] msg6_KeyCryptKey_disj [THEN disjE])
-apply (simp_all del: image_insert image_Un imp_disjL
-         add: symKey_compromise fresh_notin_analz_knows_Spy
-              analz_image_keys_simps analz_knows_absorb
-              analz_Key_image_insert_eq notin_image_iff
-              K_fresh_not_KeyCryptKey
-              DK_fresh_not_KeyCryptKey
-              analz_image_priEK)
-  --{*2.5 seconds on a 1.6GHz machine*}
-apply spy_analz  --{*Fake*}
-apply (auto intro: analz_into_parts [THEN usedI] in_parts_Says_imp_used)
-done
-
-
-subsection{*Primary Goals of Cardholder Registration *}
-
-text{*The cardholder's certificate really was created by the CA, provided the
-    CA is uncompromised *}
-
-text{*Lemma concerning the actual signed message digest*}
-lemma cert_valid_lemma:
-     "[|Crypt (priSK (CA i)) {|Hash {|Nonce N, Pan(pan C)|}, Key cardSK, N1|}
-          \<in> parts (knows Spy evs);
-        CA i \<notin> bad; evs \<in> set_cr|]
-  ==> \<exists>KC2 X Y. Says (CA i) C
-                     (Crypt KC2 
-                       {|X, certC (pan C) cardSK N onlySig (priSK (CA i)), Y|})
-                  \<in> set evs"
-apply (erule rev_mp)
-apply (erule set_cr.induct)
-apply (simp_all (no_asm_simp))
-apply auto
-done
-
-text{*Pre-packaged version for cardholder.  We don't try to confirm the values
-  of KC2, X and Y, since they are not important.*}
-lemma certificate_valid_cardSK:
-    "[|Gets C (Crypt KC2 {|X, certC (pan C) cardSK N onlySig (invKey SKi),
-                              cert (CA i) SKi onlySig (priSK RCA)|}) \<in> set evs;
-        CA i \<notin> bad; evs \<in> set_cr|]
-  ==> \<exists>KC2 X Y. Says (CA i) C
-                     (Crypt KC2 
-                       {|X, certC (pan C) cardSK N onlySig (priSK (CA i)), Y|})
-                   \<in> set evs"
-by (force dest!: Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Body]
-                    certificate_valid_pubSK cert_valid_lemma)
-
-
-lemma Hash_imp_parts [rule_format]:
-     "evs \<in> set_cr
-      ==> Hash{|X, Nonce N|} \<in> parts (knows Spy evs) -->
-          Nonce N \<in> parts (knows Spy evs)"
-apply (erule set_cr.induct, force)
-apply (simp_all (no_asm_simp))
-apply (blast intro: parts_mono [THEN [2] rev_subsetD])
-done
-
-lemma Hash_imp_parts2 [rule_format]:
-     "evs \<in> set_cr
-      ==> Hash{|X, Nonce M, Y, Nonce N|} \<in> parts (knows Spy evs) -->
-          Nonce M \<in> parts (knows Spy evs) & Nonce N \<in> parts (knows Spy evs)"
-apply (erule set_cr.induct, force)
-apply (simp_all (no_asm_simp))
-apply (blast intro: parts_mono [THEN [2] rev_subsetD])
-done
-
-
-subsection{*Secrecy of Nonces*}
-
-subsubsection{*Lemmas about the predicate KeyCryptNonce *}
-
-text{*A fresh DK cannot be associated with any other
-  (with respect to a given trace). *}
-lemma DK_fresh_not_KeyCryptNonce:
-     "[| DK \<in> symKeys; Key DK \<notin> used evs; evs \<in> set_cr |]
-      ==> ~ KeyCryptNonce DK K evs"
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule set_cr.induct)
-apply (simp_all (no_asm_simp))
-apply blast
-apply blast
-apply (auto simp add: DK_fresh_not_KeyCryptKey)
-done
-
-text{*A fresh N cannot be associated with any other
-      (with respect to a given trace). *}
-lemma N_fresh_not_KeyCryptNonce:
-     "\<forall>C. DK \<noteq> priEK C ==> Nonce N \<notin> used evs --> ~ KeyCryptNonce DK N evs"
-apply (induct_tac "evs")
-apply (case_tac [2] "a")
-apply (auto simp add: parts_insert2)
-done
-
-lemma not_KeyCryptNonce_cardSK [rule_format (no_asm)]:
-     "[|cardSK \<notin> symKeys;  \<forall>C. cardSK \<noteq> priEK C;  evs \<in> set_cr|] ==>
-      Key cardSK \<notin> analz (knows Spy evs) --> ~ KeyCryptNonce cardSK N evs"
-apply (erule set_cr.induct, analz_mono_contra, simp_all)
-apply (blast dest: not_KeyCryptKey_cardSK)  --{*6*}
-done
-
-subsubsection{*Lemmas for message 5 and 6:
-  either cardSK is compromised (when we don't care)
-  or else cardSK hasn't been used to encrypt K. *}
-
-text{*Lemma for message 5: pubSK C is never used to encrypt Nonces.*}
-lemma pubSK_not_KeyCryptNonce [simp]: "~ KeyCryptNonce (pubSK C) N evs"
-apply (induct_tac "evs")
-apply (auto simp add: parts_insert2 split add: event.split)
-done
-
-text{*Lemma for message 6: either cardSK is compromised (when we don't care)
-  or else cardSK hasn't been used to encrypt K.*}
-lemma msg6_KeyCryptNonce_disj:
-     "[|Gets B {|Crypt KC3 {|Agent C, Nonce N, Key KC2, Key cardSK, X|}, Y|}
-          \<in> set evs;
-        cardSK \<notin> symKeys;  evs \<in> set_cr|]
-      ==> Key cardSK \<in> analz (knows Spy evs) |
-          ((\<forall>K. ~ KeyCryptKey cardSK K evs) &
-           (\<forall>N. ~ KeyCryptNonce cardSK N evs))"
-by (blast dest: not_KeyCryptKey_cardSK not_KeyCryptNonce_cardSK
-          intro: cardSK_neq_priEK)
-
-
-text{*As usual: we express the property as a logical equivalence*}
-lemma Nonce_analz_image_Key_lemma:
-     "P --> (Nonce N \<in> analz (Key`KK Un H)) --> (Nonce N \<in> analz H)
-      ==> P --> (Nonce N \<in> analz (Key`KK Un H)) = (Nonce N \<in> analz H)"
-by (blast intro: analz_mono [THEN [2] rev_subsetD])
-
-
-text{*The @{text "(no_asm)"} attribute is essential, since it retains
-  the quantifier and allows the simprule's condition to itself be simplified.*}
-lemma Nonce_compromise [rule_format (no_asm)]:
-     "evs \<in> set_cr ==>
-      (\<forall>N KK. (\<forall>K \<in> KK. ~ KeyCryptNonce K N evs)   -->
-               (Nonce N \<in> analz (Key`KK Un (knows Spy evs))) =
-               (Nonce N \<in> analz (knows Spy evs)))"
-apply (erule set_cr.induct)
-apply (rule_tac [!] allI)+
-apply (rule_tac [!] impI [THEN Nonce_analz_image_Key_lemma])+
-apply (frule_tac [8] Gets_certificate_valid) --{*for message 5*}
-apply (frule_tac [6] Gets_certificate_valid) --{*for message 3*}
-apply (frule_tac [11] msg6_KeyCryptNonce_disj)
-apply (erule_tac [13] disjE)
-apply (simp_all del: image_insert image_Un
-         add: symKey_compromise
-              analz_image_keys_simps analz_knows_absorb
-              analz_Key_image_insert_eq notin_image_iff
-              N_fresh_not_KeyCryptNonce
-              DK_fresh_not_KeyCryptNonce K_fresh_not_KeyCryptKey
-              ball_conj_distrib analz_image_priEK)
-  --{*14 seconds on a 1.6GHz machine*}
-apply spy_analz  --{*Fake*}
-apply blast  --{*3*}
-apply blast  --{*5*}
-txt{*Message 6*}
-apply (metis symKey_compromise)
-  --{*cardSK compromised*}
-txt{*Simplify again--necessary because the previous simplification introduces
-  some logical connectives*} 
-apply (force simp del: image_insert image_Un imp_disjL
-          simp add: analz_image_keys_simps symKey_compromise)
-done
-
-
-subsection{*Secrecy of CardSecret: the Cardholder's secret*}
-
-lemma NC2_not_CardSecret:
-     "[|Crypt EKj {|Key K, Pan p, Hash {|Agent D, Nonce N|}|}
-          \<in> parts (knows Spy evs);
-        Key K \<notin> analz (knows Spy evs);
-        Nonce N \<notin> analz (knows Spy evs);
-       evs \<in> set_cr|]
-      ==> Crypt EKi {|Key K', Pan p', Nonce N|} \<notin> parts (knows Spy evs)"
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule set_cr.induct, analz_mono_contra, simp_all)
-apply (blast dest: Hash_imp_parts)+
-done
-
-lemma KC2_secure_lemma [rule_format]:
-     "[|U = Crypt KC3 {|Agent C, Nonce N, Key KC2, X|};
-        U \<in> parts (knows Spy evs);
-        evs \<in> set_cr|]
-  ==> Nonce N \<notin> analz (knows Spy evs) -->
-      (\<exists>k i W. Says (Cardholder k) (CA i) {|U,W|} \<in> set evs & 
-               Cardholder k \<notin> bad & CA i \<notin> bad)"
-apply (erule_tac P = "U \<in> ?H" in rev_mp)
-apply (erule set_cr.induct)
-apply (valid_certificate_tac [8])  --{*for message 5*}
-apply (simp_all del: image_insert image_Un imp_disjL
-         add: analz_image_keys_simps analz_knows_absorb
-              analz_knows_absorb2 notin_image_iff)
-  --{*4 seconds on a 1.6GHz machine*}
-apply (simp_all (no_asm_simp)) --{*leaves 4 subgoals*}
-apply (blast intro!: analz_insertI)+
-done
-
-lemma KC2_secrecy:
-     "[|Gets B {|Crypt K {|Agent C, Nonce N, Key KC2, X|}, Y|} \<in> set evs;
-        Nonce N \<notin> analz (knows Spy evs);  KC2 \<in> symKeys;
-        evs \<in> set_cr|]
-       ==> Key KC2 \<notin> analz (knows Spy evs)"
-by (force dest!: refl [THEN KC2_secure_lemma] symKey_secrecy)
-
-
-text{*Inductive version*}
-lemma CardSecret_secrecy_lemma [rule_format]:
-     "[|CA i \<notin> bad;  evs \<in> set_cr|]
-      ==> Key K \<notin> analz (knows Spy evs) -->
-          Crypt (pubEK (CA i)) {|Key K, Pan p, Nonce CardSecret|}
-             \<in> parts (knows Spy evs) -->
-          Nonce CardSecret \<notin> analz (knows Spy evs)"
-apply (erule set_cr.induct, analz_mono_contra)
-apply (valid_certificate_tac [8]) --{*for message 5*}
-apply (valid_certificate_tac [6]) --{*for message 5*}
-apply (frule_tac [9] msg6_KeyCryptNonce_disj [THEN disjE])
-apply (simp_all
-         del: image_insert image_Un imp_disjL
-         add: analz_image_keys_simps analz_knows_absorb
-              analz_Key_image_insert_eq notin_image_iff
-              EXHcrypt_def Crypt_notin_image_Key
-              N_fresh_not_KeyCryptNonce DK_fresh_not_KeyCryptNonce
-              ball_conj_distrib Nonce_compromise symKey_compromise
-              analz_image_priEK)
-  --{*2.5 seconds on a 1.6GHz machine*}
-apply spy_analz  --{*Fake*}
-apply (simp_all (no_asm_simp))
-apply blast  --{*1*}
-apply (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj])  --{*2*}
-apply blast  --{*3*}
-apply (blast dest: NC2_not_CardSecret Gets_imp_knows_Spy [THEN analz.Inj] analz_symKeys_Decrypt)  --{*4*}
-apply blast  --{*5*}
-apply (blast dest: KC2_secrecy)+  --{*Message 6: two cases*}
-done
-
-
-text{*Packaged version for cardholder*}
-lemma CardSecret_secrecy:
-     "[|Cardholder k \<notin> bad;  CA i \<notin> bad;
-        Says (Cardholder k) (CA i)
-           {|X, Crypt EKi {|Key KC3, Pan p, Nonce CardSecret|}|} \<in> set evs;
-        Gets A {|Z, cert (CA i) EKi onlyEnc (priSK RCA),
-                    cert (CA i) SKi onlySig (priSK RCA)|} \<in> set evs;
-        KC3 \<in> symKeys;  evs \<in> set_cr|]
-      ==> Nonce CardSecret \<notin> analz (knows Spy evs)"
-apply (frule Gets_certificate_valid, assumption)
-apply (subgoal_tac "Key KC3 \<notin> analz (knows Spy evs) ")
-apply (blast dest: CardSecret_secrecy_lemma)
-apply (rule symKey_secrecy)
-apply (auto simp add: parts_insert2)
-done
-
-
-subsection{*Secrecy of NonceCCA [the CA's secret] *}
-
-lemma NC2_not_NonceCCA:
-     "[|Hash {|Agent C', Nonce N', Agent C, Nonce N|}
-          \<in> parts (knows Spy evs);
-        Nonce N \<notin> analz (knows Spy evs);
-       evs \<in> set_cr|]
-      ==> Crypt KC1 {|{|Agent B, Nonce N|}, Hash p|} \<notin> parts (knows Spy evs)"
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule set_cr.induct, analz_mono_contra, simp_all)
-apply (blast dest: Hash_imp_parts2)+
-done
-
-
-text{*Inductive version*}
-lemma NonceCCA_secrecy_lemma [rule_format]:
-     "[|CA i \<notin> bad;  evs \<in> set_cr|]
-      ==> Key K \<notin> analz (knows Spy evs) -->
-          Crypt K
-            {|sign (priSK (CA i))
-                   {|Agent C, Nonce N, Agent(CA i), Nonce NonceCCA|},
-              X, Y|}
-             \<in> parts (knows Spy evs) -->
-          Nonce NonceCCA \<notin> analz (knows Spy evs)"
-apply (erule set_cr.induct, analz_mono_contra)
-apply (valid_certificate_tac [8]) --{*for message 5*}
-apply (valid_certificate_tac [6]) --{*for message 5*}
-apply (frule_tac [9] msg6_KeyCryptNonce_disj [THEN disjE])
-apply (simp_all
-         del: image_insert image_Un imp_disjL
-         add: analz_image_keys_simps analz_knows_absorb sign_def
-              analz_Key_image_insert_eq notin_image_iff
-              EXHcrypt_def Crypt_notin_image_Key
-              N_fresh_not_KeyCryptNonce DK_fresh_not_KeyCryptNonce
-              ball_conj_distrib Nonce_compromise symKey_compromise
-              analz_image_priEK)
-  --{*3 seconds on a 1.6GHz machine*}
-apply spy_analz  --{*Fake*}
-apply blast  --{*1*}
-apply (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj])  --{*2*}
-apply blast  --{*3*}
-apply (blast dest: NC2_not_NonceCCA)  --{*4*}
-apply blast  --{*5*}
-apply (blast dest: KC2_secrecy)+  --{*Message 6: two cases*}
-done
-
-
-text{*Packaged version for cardholder*}
-lemma NonceCCA_secrecy:
-     "[|Cardholder k \<notin> bad;  CA i \<notin> bad;
-        Gets (Cardholder k)
-           (Crypt KC2
-            {|sign (priSK (CA i)) {|Agent C, Nonce N, Agent(CA i), Nonce NonceCCA|},
-              X, Y|}) \<in> set evs;
-        Says (Cardholder k) (CA i)
-           {|Crypt KC3 {|Agent C, Nonce NC3, Key KC2, X'|}, Y'|} \<in> set evs;
-        Gets A {|Z, cert (CA i) EKi onlyEnc (priSK RCA),
-                    cert (CA i) SKi onlySig (priSK RCA)|} \<in> set evs;
-        KC2 \<in> symKeys;  evs \<in> set_cr|]
-      ==> Nonce NonceCCA \<notin> analz (knows Spy evs)"
-apply (frule Gets_certificate_valid, assumption)
-apply (subgoal_tac "Key KC2 \<notin> analz (knows Spy evs) ")
-apply (blast dest: NonceCCA_secrecy_lemma)
-apply (rule symKey_secrecy)
-apply (auto simp add: parts_insert2)
-done
-
-text{*We don't bother to prove guarantees for the CA.  He doesn't care about
-  the PANSecret: it isn't his credit card!*}
-
-
-subsection{*Rewriting Rule for PANs*}
-
-text{*Lemma for message 6: either cardSK isn't a CA's private encryption key,
-  or if it is then (because it appears in traffic) that CA is bad,
-  and so the Spy knows that key already.  Either way, we can simplify
-  the expression @{term "analz (insert (Key cardSK) X)"}.*}
-lemma msg6_cardSK_disj:
-     "[|Gets A {|Crypt K {|c, n, k', Key cardSK, X|}, Y|}
-          \<in> set evs;  evs \<in> set_cr |]
-      ==> cardSK \<notin> range(invKey o pubEK o CA) | Key cardSK \<in> knows Spy evs"
-by auto
-
-lemma analz_image_pan_lemma:
-     "(Pan P \<in> analz (Key`nE Un H)) --> (Pan P \<in> analz H)  ==>
-      (Pan P \<in> analz (Key`nE Un H)) =   (Pan P \<in> analz H)"
-by (blast intro: analz_mono [THEN [2] rev_subsetD])
-
-lemma analz_image_pan [rule_format]:
-     "evs \<in> set_cr ==>
-       \<forall>KK. KK <= - invKey ` pubEK ` range CA -->
-            (Pan P \<in> analz (Key`KK Un (knows Spy evs))) =
-            (Pan P \<in> analz (knows Spy evs))"
-apply (erule set_cr.induct)
-apply (rule_tac [!] allI impI)+
-apply (rule_tac [!] analz_image_pan_lemma)
-apply (valid_certificate_tac [8]) --{*for message 5*}
-apply (valid_certificate_tac [6]) --{*for message 5*}
-apply (erule_tac [9] msg6_cardSK_disj [THEN disjE])
-apply (simp_all
-         del: image_insert image_Un
-         add: analz_image_keys_simps disjoint_image_iff
-              notin_image_iff analz_image_priEK)
-  --{*6 seconds on a 1.6GHz machine*}
-apply spy_analz
-apply (simp add: insert_absorb)  --{*6*}
-done
-
-lemma analz_insert_pan:
-     "[| evs \<in> set_cr;  K \<notin> invKey ` pubEK ` range CA |] ==>
-          (Pan P \<in> analz (insert (Key K) (knows Spy evs))) =
-          (Pan P \<in> analz (knows Spy evs))"
-by (simp del: image_insert image_Un
-         add: analz_image_keys_simps analz_image_pan)
-
-
-text{*Confidentiality of the PAN\@.  Maybe we could combine the statements of
-  this theorem with @{term analz_image_pan}, requiring a single induction but
-  a much more difficult proof.*}
-lemma pan_confidentiality:
-     "[| Pan (pan C) \<in> analz(knows Spy evs); C \<noteq>Spy; evs :set_cr|]
-    ==> \<exists>i X K HN.
-        Says C (CA i) {|X, Crypt (pubEK (CA i)) {|Key K, Pan (pan C), HN|} |}
-           \<in> set evs
-      & (CA i) \<in> bad"
-apply (erule rev_mp)
-apply (erule set_cr.induct)
-apply (valid_certificate_tac [8]) --{*for message 5*}
-apply (valid_certificate_tac [6]) --{*for message 5*}
-apply (erule_tac [9] msg6_cardSK_disj [THEN disjE])
-apply (simp_all
-         del: image_insert image_Un
-         add: analz_image_keys_simps analz_insert_pan analz_image_pan
-              notin_image_iff analz_image_priEK)
-  --{*3.5 seconds on a 1.6GHz machine*}
-apply spy_analz  --{*fake*}
-apply blast  --{*3*}
-apply blast  --{*5*}
-apply (simp (no_asm_simp) add: insert_absorb)  --{*6*}
-done
-
-
-subsection{*Unicity*}
-
-lemma CR6_Says_imp_Notes:
-     "[|Says (CA i) C (Crypt KC2
-          {|sign (priSK (CA i)) {|Agent C, Nonce NC3, Agent (CA i), Nonce Y|},
-            certC (pan C) cardSK X onlySig (priSK (CA i)),
-            cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|})  \<in> set evs;
-        evs \<in> set_cr |]
-      ==> Notes (CA i) (Key cardSK) \<in> set evs"
-apply (erule rev_mp)
-apply (erule set_cr.induct)
-apply (simp_all (no_asm_simp))
-done
-
-text{*Unicity of cardSK: it uniquely identifies the other components.  
-      This holds because a CA accepts a cardSK at most once.*}
-lemma cardholder_key_unicity:
-     "[|Says (CA i) C (Crypt KC2
-          {|sign (priSK (CA i)) {|Agent C, Nonce NC3, Agent (CA i), Nonce Y|},
-            certC (pan C) cardSK X onlySig (priSK (CA i)),
-            cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|})
-          \<in> set evs;
-        Says (CA i) C' (Crypt KC2'
-          {|sign (priSK (CA i)) {|Agent C', Nonce NC3', Agent (CA i), Nonce Y'|},
-            certC (pan C') cardSK X' onlySig (priSK (CA i)),
-            cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|})
-          \<in> set evs;
-        evs \<in> set_cr |] ==> C=C' & NC3=NC3' & X=X' & KC2=KC2' & Y=Y'"
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule set_cr.induct)
-apply (simp_all (no_asm_simp))
-apply (blast dest!: CR6_Says_imp_Notes)
-done
-
-
-(*<*)
-text{*UNUSED unicity result*}
-lemma unique_KC1:
-     "[|Says C B {|Crypt KC1 X, Crypt EK {|Key KC1, Y|}|}
-          \<in> set evs;
-        Says C B' {|Crypt KC1 X', Crypt EK' {|Key KC1, Y'|}|}
-          \<in> set evs;
-        C \<notin> bad;  evs \<in> set_cr|] ==> B'=B & Y'=Y"
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule set_cr.induct, auto)
-done
-
-text{*UNUSED unicity result*}
-lemma unique_KC2:
-     "[|Says C B {|Crypt K {|Agent C, nn, Key KC2, X|}, Y|} \<in> set evs;
-        Says C B' {|Crypt K' {|Agent C, nn', Key KC2, X'|}, Y'|} \<in> set evs;
-        C \<notin> bad;  evs \<in> set_cr|] ==> B'=B & X'=X"
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule set_cr.induct, auto)
-done
-(*>*)
-
-
-text{*Cannot show cardSK to be secret because it isn't assumed to be fresh
-  it could be a previously compromised cardSK [e.g. involving a bad CA]*}
-
-
-end
--- a/src/HOL/SET-Protocol/EventSET.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,196 +0,0 @@
-(*  Title:      HOL/SET-Protocol/EventSET.thy
-    Author:     Giampaolo Bella
-    Author:     Fabio Massacci
-    Author:     Lawrence C Paulson
-*)
-
-header{*Theory of Events for SET*}
-
-theory EventSET imports MessageSET begin
-
-text{*The Root Certification Authority*}
-syntax        RCA :: agent
-translations "RCA" == "CA 0"
-
-
-text{*Message events*}
-datatype
-  event = Says  agent agent msg
-        | Gets  agent       msg
-        | Notes agent       msg
-
-
-text{*compromised agents: keys known, Notes visible*}
-consts bad :: "agent set"
-
-text{*Spy has access to his own key for spoof messages, but RCA is secure*}
-specification (bad)
-  Spy_in_bad     [iff]: "Spy \<in> bad"
-  RCA_not_bad [iff]: "RCA \<notin> bad"
-    by (rule exI [of _ "{Spy}"], simp)
-
-
-subsection{*Agents' Knowledge*}
-
-consts  (*Initial states of agents -- parameter of the construction*)
-  initState :: "agent => msg set"
-  knows  :: "[agent, event list] => msg set"
-
-(* Message reception does not extend spy's knowledge because of
-   reception invariant enforced by Reception rule in protocol definition*)
-primrec
-
-knows_Nil:
-  "knows A []       = initState A"
-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))"
-
-
-subsection{*Used Messages*}
-
-consts
-  (*Set of items that might be visible to somebody:
-    complement of the set of fresh items*)
-  used :: "event list => msg set"
-
-(* As above, message reception does extend used items *)
-primrec
-  used_Nil:  "used []         = (UN B. parts (initState B))"
-  used_Cons: "used (ev # evs) =
-                 (case ev of
-                    Says A B X => parts {X} Un (used evs)
-                  | Gets A X   => used evs
-                  | Notes A X  => parts {X} Un (used evs))"
-
-
-
-(* Inserted by default but later removed.  This declaration lets the file
-be re-loaded. Addsimps [knows_Cons, used_Nil, *)
-
-(** Simplifying   parts (insert X (knows Spy evs))
-      = parts {X} Un parts (knows Spy evs) -- since general case loops*)
-
-lemmas parts_insert_knows_A = parts_insert [of _ "knows A evs", standard]
-
-lemma knows_Spy_Says [simp]:
-     "knows Spy (Says A B X # evs) = insert X (knows Spy evs)"
-by auto
-
-text{*Letting the Spy see "bad" agents' notes avoids redundant case-splits
-      on whether @{term "A=Spy"} and whether @{term "A\<in>bad"}*}
-lemma knows_Spy_Notes [simp]:
-     "knows Spy (Notes A X # evs) =
-          (if A:bad then insert X (knows Spy evs) else knows Spy evs)"
-apply auto
-done
-
-lemma knows_Spy_Gets [simp]: "knows Spy (Gets A X # evs) = knows Spy evs"
-by auto
-
-lemma initState_subset_knows: "initState A <= knows A evs"
-apply (induct_tac "evs")
-apply (auto split: event.split) 
-done
-
-lemma knows_Spy_subset_knows_Spy_Says:
-     "knows Spy evs <= knows Spy (Says A B X # evs)"
-by auto
-
-lemma knows_Spy_subset_knows_Spy_Notes:
-     "knows Spy evs <= knows Spy (Notes A X # evs)"
-by auto
-
-lemma knows_Spy_subset_knows_Spy_Gets:
-     "knows Spy evs <= knows Spy (Gets A X # evs)"
-by auto
-
-(*Spy sees what is sent on the traffic*)
-lemma Says_imp_knows_Spy [rule_format]:
-     "Says A B X \<in> set evs --> X \<in> knows Spy evs"
-apply (induct_tac "evs")
-apply (auto split: event.split) 
-done
-
-(*Use with addSEs to derive contradictions from old Says events containing
-  items known to be fresh*)
-lemmas knows_Spy_partsEs =
-     Says_imp_knows_Spy [THEN parts.Inj, THEN revcut_rl, standard] 
-     parts.Body [THEN revcut_rl, standard]
-
-
-subsection{*The Function @{term used}*}
-
-lemma parts_knows_Spy_subset_used: "parts (knows Spy evs) <= used evs"
-apply (induct_tac "evs")
-apply (auto simp add: parts_insert_knows_A split: event.split) 
-done
-
-lemmas usedI = parts_knows_Spy_subset_used [THEN subsetD, intro]
-
-lemma initState_subset_used: "parts (initState B) <= used evs"
-apply (induct_tac "evs")
-apply (auto split: event.split) 
-done
-
-lemmas initState_into_used = initState_subset_used [THEN subsetD]
-
-lemma used_Says [simp]: "used (Says A B X # evs) = parts{X} Un used evs"
-by auto
-
-lemma used_Notes [simp]: "used (Notes A X # evs) = parts{X} Un used evs"
-by auto
-
-lemma used_Gets [simp]: "used (Gets A X # evs) = used evs"
-by auto
-
-
-lemma Notes_imp_parts_subset_used [rule_format]:
-     "Notes A X \<in> set evs --> parts {X} <= used evs"
-apply (induct_tac "evs")
-apply (induct_tac [2] "a", auto)
-done
-
-text{*NOTE REMOVAL--laws above are cleaner, as they don't involve "case"*}
-declare knows_Cons [simp del]
-        used_Nil [simp del] used_Cons [simp del]
-
-
-text{*For proving theorems of the form @{term "X \<notin> analz (knows Spy evs) --> P"}
-  New events added by induction to "evs" are discarded.  Provided 
-  this information isn't needed, the proof will be much shorter, since
-  it will omit complicated reasoning about @{term analz}.*}
-
-lemmas analz_mono_contra =
-       knows_Spy_subset_knows_Spy_Says [THEN analz_mono, THEN contra_subsetD]
-       knows_Spy_subset_knows_Spy_Notes [THEN analz_mono, THEN contra_subsetD]
-       knows_Spy_subset_knows_Spy_Gets [THEN analz_mono, THEN contra_subsetD]
-
-lemmas analz_impI = impI [where P = "Y \<notin> analz (knows Spy evs)", standard]
-
-ML
-{*
-val analz_mono_contra_tac = 
-  rtac @{thm analz_impI} THEN' 
-  REPEAT1 o (dresolve_tac @{thms analz_mono_contra})
-  THEN' mp_tac
-*}
-
-method_setup analz_mono_contra = {*
-    Scan.succeed (K (SIMPLE_METHOD (REPEAT_FIRST analz_mono_contra_tac))) *}
-    "for proving theorems of the form X \<notin> analz (knows Spy evs) --> P"
-
-end
--- a/src/HOL/SET-Protocol/Merchant_Registration.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,427 +0,0 @@
-(*  Title:      HOL/SET-Protocol/Merchant_Registration.thy
-    Author:     Giampaolo Bella
-    Author:     Fabio Massacci
-    Author:     Lawrence C Paulson
-*)
-
-header{*The SET Merchant Registration Protocol*}
-
-theory Merchant_Registration imports PublicSET begin
-
-text{*Copmpared with Cardholder Reigstration, @{text KeyCryptKey} is not
-  needed: no session key encrypts another.  Instead we
-  prove the "key compromise" theorems for sets KK that contain no private
-  encryption keys (@{term "priEK C"}). *}
-
-
-inductive_set
-  set_mr :: "event list set"
-where
-
-  Nil:    --{*Initial trace is empty*}
-           "[] \<in> set_mr"
-
-
-| Fake:    --{*The spy MAY say anything he CAN say.*}
-           "[| evsf \<in> set_mr; X \<in> synth (analz (knows Spy evsf)) |]
-            ==> Says Spy B X  # evsf \<in> set_mr"
-        
-
-| Reception: --{*If A sends a message X to B, then B might receive it*}
-             "[| evsr \<in> set_mr; Says A B X \<in> set evsr |]
-              ==> Gets B X  # evsr \<in> set_mr"
-
-
-| SET_MR1: --{*RegFormReq: M requires a registration form to a CA*}
-           "[| evs1 \<in> set_mr; M = Merchant k; Nonce NM1 \<notin> used evs1 |]
-            ==> Says M (CA i) {|Agent M, Nonce NM1|} # evs1 \<in> set_mr"
-
-
-| SET_MR2: --{*RegFormRes: CA replies with the registration form and the 
-               certificates for her keys*}
-  "[| evs2 \<in> set_mr; Nonce NCA \<notin> used evs2;
-      Gets (CA i) {|Agent M, Nonce NM1|} \<in> set evs2 |]
-   ==> Says (CA i) M {|sign (priSK (CA i)) {|Agent M, Nonce NM1, Nonce NCA|},
-                       cert (CA i) (pubEK (CA i)) onlyEnc (priSK RCA),
-                       cert (CA i) (pubSK (CA i)) onlySig (priSK RCA) |}
-         # evs2 \<in> set_mr"
-
-| SET_MR3:
-         --{*CertReq: M submits the key pair to be certified.  The Notes
-             event allows KM1 to be lost if M is compromised. Piero remarks
-             that the agent mentioned inside the signature is not verified to
-             correspond to M.  As in CR, each Merchant has fixed key pairs.  M
-             is only optionally required to send NCA back, so M doesn't do so
-             in the model*}
-  "[| evs3 \<in> set_mr; M = Merchant k; Nonce NM2 \<notin> used evs3;
-      Key KM1 \<notin> used evs3;  KM1 \<in> symKeys;
-      Gets M {|sign (invKey SKi) {|Agent X, Nonce NM1, Nonce NCA|},
-               cert (CA i) EKi onlyEnc (priSK RCA),
-               cert (CA i) SKi onlySig (priSK RCA) |}
-        \<in> set evs3;
-      Says M (CA i) {|Agent M, Nonce NM1|} \<in> set evs3 |]
-   ==> Says M (CA i)
-            {|Crypt KM1 (sign (priSK M) {|Agent M, Nonce NM2,
-                                          Key (pubSK M), Key (pubEK M)|}),
-              Crypt EKi (Key KM1)|}
-         # Notes M {|Key KM1, Agent (CA i)|}
-         # evs3 \<in> set_mr"
-
-| SET_MR4:
-         --{*CertRes: CA issues the certificates for merSK and merEK,
-             while checking never to have certified the m even
-             separately. NOTE: In Cardholder Registration the
-             corresponding rule (6) doesn't use the "sign" primitive. "The
-             CertRes shall be signed but not encrypted if the EE is a Merchant
-             or Payment Gateway."-- Programmer's Guide, page 191.*}
-    "[| evs4 \<in> set_mr; M = Merchant k;
-        merSK \<notin> symKeys;  merEK \<notin> symKeys;
-        Notes (CA i) (Key merSK) \<notin> set evs4;
-        Notes (CA i) (Key merEK) \<notin> set evs4;
-        Gets (CA i) {|Crypt KM1 (sign (invKey merSK)
-                                 {|Agent M, Nonce NM2, Key merSK, Key merEK|}),
-                      Crypt (pubEK (CA i)) (Key KM1) |}
-          \<in> set evs4 |]
-    ==> Says (CA i) M {|sign (priSK(CA i)) {|Agent M, Nonce NM2, Agent(CA i)|},
-                        cert  M      merSK    onlySig (priSK (CA i)),
-                        cert  M      merEK    onlyEnc (priSK (CA i)),
-                        cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|}
-          # Notes (CA i) (Key merSK)
-          # Notes (CA i) (Key merEK)
-          # evs4 \<in> set_mr"
-
-
-text{*Note possibility proofs are missing.*}
-
-declare Says_imp_knows_Spy [THEN parts.Inj, dest]
-declare parts.Body [dest]
-declare analz_into_parts [dest]
-declare Fake_parts_insert_in_Un [dest]
-
-text{*General facts about message reception*}
-lemma Gets_imp_Says:
-     "[| Gets B X \<in> set evs; evs \<in> set_mr |] ==> \<exists>A. Says A B X \<in> set evs"
-apply (erule rev_mp)
-apply (erule set_mr.induct, auto)
-done
-
-lemma Gets_imp_knows_Spy:
-     "[| Gets B X \<in> set evs; evs \<in> set_mr |]  ==> X \<in> knows Spy evs"
-by (blast dest!: Gets_imp_Says Says_imp_knows_Spy)
-
-
-declare Gets_imp_knows_Spy [THEN parts.Inj, dest]
-
-subsubsection{*Proofs on keys *}
-
-text{*Spy never sees an agent's private keys! (unless it's bad at start)*}
-lemma Spy_see_private_Key [simp]:
-     "evs \<in> set_mr
-      ==> (Key(invKey (publicKey b A)) \<in> parts(knows Spy evs)) = (A \<in> bad)"
-apply (erule set_mr.induct)
-apply (auto dest!: Gets_imp_knows_Spy [THEN parts.Inj])
-done
-
-lemma Spy_analz_private_Key [simp]:
-     "evs \<in> set_mr ==>
-     (Key(invKey (publicKey b A)) \<in> analz(knows Spy evs)) = (A \<in> bad)"
-by auto
-
-declare Spy_see_private_Key [THEN [2] rev_iffD1, dest!]
-declare Spy_analz_private_Key [THEN [2] rev_iffD1, dest!]
-
-(*This is to state that the signed keys received in step 4
-  are into parts - rather than installing sign_def each time.
-  Needed in Spy_see_priSK_RCA, Spy_see_priEK and in Spy_see_priSK
-Goal "[|Gets C \<lbrace>Crypt KM1
-                (sign K \<lbrace>Agent M, Nonce NM2, Key merSK, Key merEK\<rbrace>), X\<rbrace>
-          \<in> set evs;  evs \<in> set_mr |]
-    ==> Key merSK \<in> parts (knows Spy evs) \<and>
-        Key merEK \<in> parts (knows Spy evs)"
-by (fast_tac (claset() addss (simpset())) 1);
-qed "signed_keys_in_parts";
-???*)
-
-text{*Proofs on certificates -
-  they hold, as in CR, because RCA's keys are secure*}
-
-lemma Crypt_valid_pubEK:
-     "[| Crypt (priSK RCA) {|Agent (CA i), Key EKi, onlyEnc|}
-           \<in> parts (knows Spy evs);
-         evs \<in> set_mr |] ==> EKi = pubEK (CA i)"
-apply (erule rev_mp)
-apply (erule set_mr.induct, auto)
-done
-
-lemma certificate_valid_pubEK:
-    "[| cert (CA i) EKi onlyEnc (priSK RCA) \<in> parts (knows Spy evs);
-        evs \<in> set_mr |]
-     ==> EKi = pubEK (CA i)"
-apply (unfold cert_def signCert_def)
-apply (blast dest!: Crypt_valid_pubEK)
-done
-
-lemma Crypt_valid_pubSK:
-     "[| Crypt (priSK RCA) {|Agent (CA i), Key SKi, onlySig|}
-           \<in> parts (knows Spy evs);
-         evs \<in> set_mr |] ==> SKi = pubSK (CA i)"
-apply (erule rev_mp)
-apply (erule set_mr.induct, auto)
-done
-
-lemma certificate_valid_pubSK:
-    "[| cert (CA i) SKi onlySig (priSK RCA) \<in> parts (knows Spy evs);
-        evs \<in> set_mr |] ==> SKi = pubSK (CA i)"
-apply (unfold cert_def signCert_def)
-apply (blast dest!: Crypt_valid_pubSK)
-done
-
-lemma Gets_certificate_valid:
-     "[| Gets A {| X, cert (CA i) EKi onlyEnc (priSK RCA),
-                      cert (CA i) SKi onlySig (priSK RCA)|} \<in> set evs;
-         evs \<in> set_mr |]
-      ==> EKi = pubEK (CA i) & SKi = pubSK (CA i)"
-by (blast dest: certificate_valid_pubEK certificate_valid_pubSK)
-
-
-text{*Nobody can have used non-existent keys!*}
-lemma new_keys_not_used [rule_format,simp]:
-     "evs \<in> set_mr
-      ==> Key K \<notin> used evs --> K \<in> symKeys -->
-          K \<notin> keysFor (parts (knows Spy evs))"
-apply (erule set_mr.induct, simp_all)
-apply (force dest!: usedI keysFor_parts_insert)  --{*Fake*}
-apply force  --{*Message 2*}
-apply (blast dest: Gets_certificate_valid)  --{*Message 3*}
-apply force  --{*Message 4*}
-done
-
-
-subsubsection{*New Versions: As Above, but Generalized with the Kk Argument*}
-
-lemma gen_new_keys_not_used [rule_format]:
-     "evs \<in> set_mr
-      ==> Key K \<notin> used evs --> K \<in> symKeys -->
-          K \<notin> keysFor (parts (Key`KK Un knows Spy evs))"
-by auto
-
-lemma gen_new_keys_not_analzd:
-     "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_mr |]
-      ==> K \<notin> keysFor (analz (Key`KK Un knows Spy evs))"
-by (blast intro: keysFor_mono [THEN [2] rev_subsetD]
-          dest: gen_new_keys_not_used)
-
-lemma analz_Key_image_insert_eq:
-     "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_mr |]
-      ==> analz (Key ` (insert K KK) \<union> knows Spy evs) =
-          insert (Key K) (analz (Key ` KK \<union> knows Spy evs))"
-by (simp add: gen_new_keys_not_analzd)
-
-
-lemma Crypt_parts_imp_used:
-     "[|Crypt K X \<in> parts (knows Spy evs);
-        K \<in> symKeys; evs \<in> set_mr |] ==> Key K \<in> used evs"
-apply (rule ccontr)
-apply (force dest: new_keys_not_used Crypt_imp_invKey_keysFor)
-done
-
-lemma Crypt_analz_imp_used:
-     "[|Crypt K X \<in> analz (knows Spy evs);
-        K \<in> symKeys; evs \<in> set_mr |] ==> Key K \<in> used evs"
-by (blast intro: Crypt_parts_imp_used)
-
-text{*Rewriting rule for private encryption keys.  Analogous rewriting rules
-for other keys aren't needed.*}
-
-lemma parts_image_priEK:
-     "[|Key (priEK (CA i)) \<in> parts (Key`KK Un (knows Spy evs));
-        evs \<in> set_mr|] ==> priEK (CA i) \<in> KK | CA i \<in> bad"
-by auto
-
-text{*trivial proof because (priEK (CA i)) never appears even in (parts evs)*}
-lemma analz_image_priEK:
-     "evs \<in> set_mr ==>
-          (Key (priEK (CA i)) \<in> analz (Key`KK Un (knows Spy evs))) =
-          (priEK (CA i) \<in> KK | CA i \<in> bad)"
-by (blast dest!: parts_image_priEK intro: analz_mono [THEN [2] rev_subsetD])
-
-
-subsection{*Secrecy of Session Keys*}
-
-text{*This holds because if (priEK (CA i)) appears in any traffic then it must
-  be known to the Spy, by @{text Spy_see_private_Key}*}
-lemma merK_neq_priEK:
-     "[|Key merK \<notin> analz (knows Spy evs);
-        Key merK \<in> parts (knows Spy evs);
-        evs \<in> set_mr|] ==> merK \<noteq> priEK C"
-by blast
-
-text{*Lemma for message 4: either merK is compromised (when we don't care)
-  or else merK hasn't been used to encrypt K.*}
-lemma msg4_priEK_disj:
-     "[|Gets B {|Crypt KM1
-                       (sign K {|Agent M, Nonce NM2, Key merSK, Key merEK|}),
-                 Y|} \<in> set evs;
-        evs \<in> set_mr|]
-  ==> (Key merSK \<in> analz (knows Spy evs) | merSK \<notin> range(\<lambda>C. priEK C))
-   &  (Key merEK \<in> analz (knows Spy evs) | merEK \<notin> range(\<lambda>C. priEK C))"
-apply (unfold sign_def)
-apply (blast dest: merK_neq_priEK)
-done
-
-
-lemma Key_analz_image_Key_lemma:
-     "P --> (Key K \<in> analz (Key`KK Un H)) --> (K\<in>KK | Key K \<in> analz H)
-      ==>
-      P --> (Key K \<in> analz (Key`KK Un H)) = (K\<in>KK | Key K \<in> analz H)"
-by (blast intro: analz_mono [THEN [2] rev_subsetD])
-
-lemma symKey_compromise:
-     "evs \<in> set_mr ==>
-      (\<forall>SK KK. SK \<in> symKeys \<longrightarrow> (\<forall>K \<in> KK. K \<notin> range(\<lambda>C. priEK C)) -->
-               (Key SK \<in> analz (Key`KK Un (knows Spy evs))) =
-               (SK \<in> KK | Key SK \<in> analz (knows Spy evs)))"
-apply (erule set_mr.induct)
-apply (safe del: impI intro!: Key_analz_image_Key_lemma [THEN impI])
-apply (drule_tac [7] msg4_priEK_disj)
-apply (frule_tac [6] Gets_certificate_valid)
-apply (safe del: impI)
-apply (simp_all del: image_insert image_Un imp_disjL
-         add: analz_image_keys_simps abbrev_simps analz_knows_absorb
-              analz_knows_absorb2 analz_Key_image_insert_eq notin_image_iff
-              Spy_analz_private_Key analz_image_priEK)
-  --{*5 seconds on a 1.6GHz machine*}
-apply spy_analz  --{*Fake*}
-apply auto  --{*Message 3*}
-done
-
-lemma symKey_secrecy [rule_format]:
-     "[|CA i \<notin> bad; K \<in> symKeys;  evs \<in> set_mr|]
-      ==> \<forall>X m. Says (Merchant m) (CA i) X \<in> set evs -->
-                Key K \<in> parts{X} -->
-                Merchant m \<notin> bad -->
-                Key K \<notin> analz (knows Spy evs)"
-apply (erule set_mr.induct)
-apply (drule_tac [7] msg4_priEK_disj)
-apply (frule_tac [6] Gets_certificate_valid)
-apply (safe del: impI)
-apply (simp_all del: image_insert image_Un imp_disjL
-         add: analz_image_keys_simps abbrev_simps analz_knows_absorb
-              analz_knows_absorb2 analz_Key_image_insert_eq
-              symKey_compromise notin_image_iff Spy_analz_private_Key
-              analz_image_priEK)
-apply spy_analz  --{*Fake*}
-apply force  --{*Message 1*}
-apply (auto intro: analz_into_parts [THEN usedI] in_parts_Says_imp_used)  --{*Message 3*}
-done
-
-subsection{*Unicity *}
-
-lemma msg4_Says_imp_Notes:
- "[|Says (CA i) M {|sign (priSK (CA i)) {|Agent M, Nonce NM2, Agent (CA i)|},
-                    cert  M      merSK    onlySig (priSK (CA i)),
-                    cert  M      merEK    onlyEnc (priSK (CA i)),
-                    cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|} \<in> set evs;
-    evs \<in> set_mr |]
-  ==> Notes (CA i) (Key merSK) \<in> set evs
-   &  Notes (CA i) (Key merEK) \<in> set evs"
-apply (erule rev_mp)
-apply (erule set_mr.induct)
-apply (simp_all (no_asm_simp))
-done
-
-text{*Unicity of merSK wrt a given CA:
-  merSK uniquely identifies the other components, including merEK*}
-lemma merSK_unicity:
- "[|Says (CA i) M {|sign (priSK(CA i)) {|Agent M, Nonce NM2, Agent (CA i)|},
-                    cert  M      merSK    onlySig (priSK (CA i)),
-                    cert  M      merEK    onlyEnc (priSK (CA i)),
-                    cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|} \<in> set evs;
-    Says (CA i) M' {|sign (priSK(CA i)) {|Agent M', Nonce NM2', Agent (CA i)|},
-                    cert  M'      merSK    onlySig (priSK (CA i)),
-                    cert  M'      merEK'    onlyEnc (priSK (CA i)),
-                    cert (CA i) (pubSK(CA i)) onlySig (priSK RCA)|} \<in> set evs;
-    evs \<in> set_mr |] ==> M=M' & NM2=NM2' & merEK=merEK'"
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule set_mr.induct)
-apply (simp_all (no_asm_simp))
-apply (blast dest!: msg4_Says_imp_Notes)
-done
-
-text{*Unicity of merEK wrt a given CA:
-  merEK uniquely identifies the other components, including merSK*}
-lemma merEK_unicity:
- "[|Says (CA i) M {|sign (priSK(CA i)) {|Agent M, Nonce NM2, Agent (CA i)|},
-                    cert  M      merSK    onlySig (priSK (CA i)),
-                    cert  M      merEK    onlyEnc (priSK (CA i)),
-                    cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|} \<in> set evs;
-    Says (CA i) M' {|sign (priSK(CA i)) {|Agent M', Nonce NM2', Agent (CA i)|},
-                     cert  M'      merSK'    onlySig (priSK (CA i)),
-                     cert  M'      merEK    onlyEnc (priSK (CA i)),
-                     cert (CA i) (pubSK(CA i)) onlySig (priSK RCA)|} \<in> set evs;
-    evs \<in> set_mr |] 
-  ==> M=M' & NM2=NM2' & merSK=merSK'"
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule set_mr.induct)
-apply (simp_all (no_asm_simp))
-apply (blast dest!: msg4_Says_imp_Notes)
-done
-
-
-text{* -No interest on secrecy of nonces: they appear to be used
-    only for freshness.
-   -No interest on secrecy of merSK or merEK, as in CR.
-   -There's no equivalent of the PAN*}
-
-
-subsection{*Primary Goals of Merchant Registration *}
-
-subsubsection{*The merchant's certificates really were created by the CA,
-provided the CA is uncompromised *}
-
-text{*The assumption @{term "CA i \<noteq> RCA"} is required: step 2 uses 
-  certificates of the same form.*}
-lemma certificate_merSK_valid_lemma [intro]:
-     "[|Crypt (priSK (CA i)) {|Agent M, Key merSK, onlySig|}
-          \<in> parts (knows Spy evs);
-        CA i \<notin> bad;  CA i \<noteq> RCA;  evs \<in> set_mr|]
- ==> \<exists>X Y Z. Says (CA i) M
-                  {|X, cert M merSK onlySig (priSK (CA i)), Y, Z|} \<in> set evs"
-apply (erule rev_mp)
-apply (erule set_mr.induct)
-apply (simp_all (no_asm_simp))
-apply auto
-done
-
-lemma certificate_merSK_valid:
-     "[| cert M merSK onlySig (priSK (CA i)) \<in> parts (knows Spy evs);
-         CA i \<notin> bad;  CA i \<noteq> RCA;  evs \<in> set_mr|]
- ==> \<exists>X Y Z. Says (CA i) M
-                  {|X, cert M merSK onlySig (priSK (CA i)), Y, Z|} \<in> set evs"
-by auto
-
-lemma certificate_merEK_valid_lemma [intro]:
-     "[|Crypt (priSK (CA i)) {|Agent M, Key merEK, onlyEnc|}
-          \<in> parts (knows Spy evs);
-        CA i \<notin> bad;  CA i \<noteq> RCA;  evs \<in> set_mr|]
- ==> \<exists>X Y Z. Says (CA i) M
-                  {|X, Y, cert M merEK onlyEnc (priSK (CA i)), Z|} \<in> set evs"
-apply (erule rev_mp)
-apply (erule set_mr.induct)
-apply (simp_all (no_asm_simp))
-apply auto
-done
-
-lemma certificate_merEK_valid:
-     "[| cert M merEK onlyEnc (priSK (CA i)) \<in> parts (knows Spy evs);
-         CA i \<notin> bad;  CA i \<noteq> RCA;  evs \<in> set_mr|]
- ==> \<exists>X Y Z. Says (CA i) M
-                  {|X, Y, cert M merEK onlyEnc (priSK (CA i)), Z|} \<in> set evs"
-by auto
-
-text{*The two certificates - for merSK and for merEK - cannot be proved to
-  have originated together*}
-
-end
--- a/src/HOL/SET-Protocol/MessageSET.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,957 +0,0 @@
-(*  Title:      HOL/SET-Protocol/MessageSET.thy
-    Author:     Giampaolo Bella
-    Author:     Fabio Massacci
-    Author:     Lawrence C Paulson
-*)
-
-header{*The Message Theory, Modified for SET*}
-
-theory MessageSET
-imports Main Nat_Int_Bij
-begin
-
-subsection{*General Lemmas*}
-
-text{*Needed occasionally with @{text spy_analz_tac}, e.g. in
-     @{text analz_insert_Key_newK}*}
-
-lemma Un_absorb3 [simp] : "A \<union> (B \<union> A) = B \<union> A"
-by blast
-
-text{*Collapses redundant cases in the huge protocol proofs*}
-lemmas disj_simps = disj_comms disj_left_absorb disj_assoc 
-
-text{*Effective with assumptions like @{term "K \<notin> range pubK"} and 
-   @{term "K \<notin> invKey`range pubK"}*}
-lemma notin_image_iff: "(y \<notin> f`I) = (\<forall>i\<in>I. f i \<noteq> y)"
-by blast
-
-text{*Effective with the assumption @{term "KK \<subseteq> - (range(invKey o pubK))"} *}
-lemma disjoint_image_iff: "(A <= - (f`I)) = (\<forall>i\<in>I. f i \<notin> A)"
-by blast
-
-
-
-types
-  key = nat
-
-consts
-  all_symmetric :: bool        --{*true if all keys are symmetric*}
-  invKey        :: "key=>key"  --{*inverse of a symmetric key*}
-
-specification (invKey)
-  invKey [simp]: "invKey (invKey K) = K"
-  invKey_symmetric: "all_symmetric --> invKey = id"
-    by (rule exI [of _ id], auto)
-
-
-text{*The inverse of a symmetric key is itself; that of a public key
-      is the private key and vice versa*}
-
-constdefs
-  symKeys :: "key set"
-  "symKeys == {K. invKey K = K}"
-
-text{*Agents. We allow any number of certification authorities, cardholders
-            merchants, and payment gateways.*}
-datatype
-  agent = CA nat | Cardholder nat | Merchant nat | PG nat | Spy
-
-text{*Messages*}
-datatype
-     msg = Agent  agent     --{*Agent names*}
-         | Number nat       --{*Ordinary integers, timestamps, ...*}
-         | Nonce  nat       --{*Unguessable nonces*}
-         | Pan    nat       --{*Unguessable Primary Account Numbers (??)*}
-         | Key    key       --{*Crypto keys*}
-         | Hash   msg       --{*Hashing*}
-         | MPair  msg msg   --{*Compound messages*}
-         | Crypt  key msg   --{*Encryption, public- or shared-key*}
-
-
-(*Concrete syntax: messages appear as {|A,B,NA|}, etc...*)
-syntax
-  "@MTuple"      :: "['a, args] => 'a * 'b"       ("(2{|_,/ _|})")
-
-syntax (xsymbols)
-  "@MTuple"      :: "['a, args] => 'a * 'b"       ("(2\<lbrace>_,/ _\<rbrace>)")
-
-translations
-  "{|x, y, z|}"   == "{|x, {|y, z|}|}"
-  "{|x, y|}"      == "MPair x y"
-
-
-constdefs
-  nat_of_agent :: "agent => nat"
-   "nat_of_agent == agent_case (curry nat2_to_nat 0)
-                               (curry nat2_to_nat 1)
-                               (curry nat2_to_nat 2)
-                               (curry nat2_to_nat 3)
-                               (nat2_to_nat (4,0))"
-    --{*maps each agent to a unique natural number, for specifications*}
-
-text{*The function is indeed injective*}
-lemma inj_nat_of_agent: "inj nat_of_agent"
-by (simp add: nat_of_agent_def inj_on_def curry_def
-              nat2_to_nat_inj [THEN inj_eq]  split: agent.split) 
-
-
-constdefs
-  (*Keys useful to decrypt elements of a message set*)
-  keysFor :: "msg set => key set"
-  "keysFor H == invKey ` {K. \<exists>X. Crypt K X \<in> H}"
-
-subsubsection{*Inductive definition of all "parts" of a message.*}
-
-inductive_set
-  parts :: "msg set => msg set"
-  for H :: "msg set"
-  where
-    Inj [intro]:               "X \<in> H ==> X \<in> parts H"
-  | Fst:         "{|X,Y|}   \<in> parts H ==> X \<in> parts H"
-  | Snd:         "{|X,Y|}   \<in> parts H ==> Y \<in> parts H"
-  | Body:        "Crypt K X \<in> parts H ==> X \<in> parts H"
-
-
-(*Monotonicity*)
-lemma parts_mono: "G<=H ==> parts(G) <= parts(H)"
-apply auto
-apply (erule parts.induct)
-apply (auto dest: Fst Snd Body)
-done
-
-
-subsubsection{*Inverse of keys*}
-
-(*Equations hold because constructors are injective; cannot prove for all f*)
-lemma Key_image_eq [simp]: "(Key x \<in> Key`A) = (x\<in>A)"
-by auto
-
-lemma Nonce_Key_image_eq [simp]: "(Nonce x \<notin> Key`A)"
-by auto
-
-lemma Cardholder_image_eq [simp]: "(Cardholder x \<in> Cardholder`A) = (x \<in> A)"
-by auto
-
-lemma CA_image_eq [simp]: "(CA x \<in> CA`A) = (x \<in> A)"
-by auto
-
-lemma Pan_image_eq [simp]: "(Pan x \<in> Pan`A) = (x \<in> A)"
-by auto
-
-lemma Pan_Key_image_eq [simp]: "(Pan x \<notin> Key`A)"
-by auto
-
-lemma Nonce_Pan_image_eq [simp]: "(Nonce x \<notin> Pan`A)"
-by auto
-
-lemma invKey_eq [simp]: "(invKey K = invKey K') = (K=K')"
-apply safe
-apply (drule_tac f = invKey in arg_cong, simp)
-done
-
-
-subsection{*keysFor operator*}
-
-lemma keysFor_empty [simp]: "keysFor {} = {}"
-by (unfold keysFor_def, blast)
-
-lemma keysFor_Un [simp]: "keysFor (H \<union> H') = keysFor H \<union> keysFor H'"
-by (unfold keysFor_def, blast)
-
-lemma keysFor_UN [simp]: "keysFor (\<Union>i\<in>A. H i) = (\<Union>i\<in>A. keysFor (H i))"
-by (unfold keysFor_def, blast)
-
-(*Monotonicity*)
-lemma keysFor_mono: "G\<subseteq>H ==> keysFor(G) \<subseteq> keysFor(H)"
-by (unfold keysFor_def, blast)
-
-lemma keysFor_insert_Agent [simp]: "keysFor (insert (Agent A) H) = keysFor H"
-by (unfold keysFor_def, auto)
-
-lemma keysFor_insert_Nonce [simp]: "keysFor (insert (Nonce N) H) = keysFor H"
-by (unfold keysFor_def, auto)
-
-lemma keysFor_insert_Number [simp]: "keysFor (insert (Number N) H) = keysFor H"
-by (unfold keysFor_def, auto)
-
-lemma keysFor_insert_Key [simp]: "keysFor (insert (Key K) H) = keysFor H"
-by (unfold keysFor_def, auto)
-
-lemma keysFor_insert_Pan [simp]: "keysFor (insert (Pan A) H) = keysFor H"
-by (unfold keysFor_def, auto)
-
-lemma keysFor_insert_Hash [simp]: "keysFor (insert (Hash X) H) = keysFor H"
-by (unfold keysFor_def, auto)
-
-lemma keysFor_insert_MPair [simp]: "keysFor (insert {|X,Y|} H) = keysFor H"
-by (unfold keysFor_def, auto)
-
-lemma keysFor_insert_Crypt [simp]:
-    "keysFor (insert (Crypt K X) H) = insert (invKey K) (keysFor H)"
-by (unfold keysFor_def, auto)
-
-lemma keysFor_image_Key [simp]: "keysFor (Key`E) = {}"
-by (unfold keysFor_def, auto)
-
-lemma Crypt_imp_invKey_keysFor: "Crypt K X \<in> H ==> invKey K \<in> keysFor H"
-by (unfold keysFor_def, blast)
-
-
-subsection{*Inductive relation "parts"*}
-
-lemma MPair_parts:
-     "[| {|X,Y|} \<in> parts H;
-         [| X \<in> parts H; Y \<in> parts H |] ==> P |] ==> P"
-by (blast dest: parts.Fst parts.Snd)
-
-declare MPair_parts [elim!]  parts.Body [dest!]
-text{*NB These two rules are UNSAFE in the formal sense, as they discard the
-     compound message.  They work well on THIS FILE.
-  @{text MPair_parts} is left as SAFE because it speeds up proofs.
-  The Crypt rule is normally kept UNSAFE to avoid breaking up certificates.*}
-
-lemma parts_increasing: "H \<subseteq> parts(H)"
-by blast
-
-lemmas parts_insertI = subset_insertI [THEN parts_mono, THEN subsetD, standard]
-
-lemma parts_empty [simp]: "parts{} = {}"
-apply safe
-apply (erule parts.induct, blast+)
-done
-
-lemma parts_emptyE [elim!]: "X\<in> parts{} ==> P"
-by simp
-
-(*WARNING: loops if H = {Y}, therefore must not be repeated!*)
-lemma parts_singleton: "X\<in> parts H ==> \<exists>Y\<in>H. X\<in> parts {Y}"
-by (erule parts.induct, fast+)
-
-
-subsubsection{*Unions*}
-
-lemma parts_Un_subset1: "parts(G) \<union> parts(H) \<subseteq> parts(G \<union> H)"
-by (intro Un_least parts_mono Un_upper1 Un_upper2)
-
-lemma parts_Un_subset2: "parts(G \<union> H) \<subseteq> parts(G) \<union> parts(H)"
-apply (rule subsetI)
-apply (erule parts.induct, blast+)
-done
-
-lemma parts_Un [simp]: "parts(G \<union> H) = parts(G) \<union> parts(H)"
-by (intro equalityI parts_Un_subset1 parts_Un_subset2)
-
-lemma parts_insert: "parts (insert X H) = parts {X} \<union> parts H"
-apply (subst insert_is_Un [of _ H])
-apply (simp only: parts_Un)
-done
-
-(*TWO inserts to avoid looping.  This rewrite is better than nothing.
-  Not suitable for Addsimps: its behaviour can be strange.*)
-lemma parts_insert2:
-     "parts (insert X (insert Y H)) = parts {X} \<union> parts {Y} \<union> parts H"
-apply (simp add: Un_assoc)
-apply (simp add: parts_insert [symmetric])
-done
-
-lemma parts_UN_subset1: "(\<Union>x\<in>A. parts(H x)) \<subseteq> parts(\<Union>x\<in>A. H x)"
-by (intro UN_least parts_mono UN_upper)
-
-lemma parts_UN_subset2: "parts(\<Union>x\<in>A. H x) \<subseteq> (\<Union>x\<in>A. parts(H x))"
-apply (rule subsetI)
-apply (erule parts.induct, blast+)
-done
-
-lemma parts_UN [simp]: "parts(\<Union>x\<in>A. H x) = (\<Union>x\<in>A. parts(H x))"
-by (intro equalityI parts_UN_subset1 parts_UN_subset2)
-
-(*Added to simplify arguments to parts, analz and synth.
-  NOTE: the UN versions are no longer used!*)
-
-
-text{*This allows @{text blast} to simplify occurrences of
-  @{term "parts(G\<union>H)"} in the assumption.*}
-declare parts_Un [THEN equalityD1, THEN subsetD, THEN UnE, elim!]
-
-
-lemma parts_insert_subset: "insert X (parts H) \<subseteq> parts(insert X H)"
-by (blast intro: parts_mono [THEN [2] rev_subsetD])
-
-subsubsection{*Idempotence and transitivity*}
-
-lemma parts_partsD [dest!]: "X\<in> parts (parts H) ==> X\<in> parts H"
-by (erule parts.induct, blast+)
-
-lemma parts_idem [simp]: "parts (parts H) = parts H"
-by blast
-
-lemma parts_trans: "[| X\<in> parts G;  G \<subseteq> parts H |] ==> X\<in> parts H"
-by (drule parts_mono, blast)
-
-(*Cut*)
-lemma parts_cut:
-     "[| Y\<in> parts (insert X G);  X\<in> parts H |] ==> Y\<in> parts (G \<union> H)"
-by (erule parts_trans, auto)
-
-lemma parts_cut_eq [simp]: "X\<in> parts H ==> parts (insert X H) = parts H"
-by (force dest!: parts_cut intro: parts_insertI)
-
-
-subsubsection{*Rewrite rules for pulling out atomic messages*}
-
-lemmas parts_insert_eq_I = equalityI [OF subsetI parts_insert_subset]
-
-
-lemma parts_insert_Agent [simp]:
-     "parts (insert (Agent agt) H) = insert (Agent agt) (parts H)"
-apply (rule parts_insert_eq_I)
-apply (erule parts.induct, auto)
-done
-
-lemma parts_insert_Nonce [simp]:
-     "parts (insert (Nonce N) H) = insert (Nonce N) (parts H)"
-apply (rule parts_insert_eq_I)
-apply (erule parts.induct, auto)
-done
-
-lemma parts_insert_Number [simp]:
-     "parts (insert (Number N) H) = insert (Number N) (parts H)"
-apply (rule parts_insert_eq_I)
-apply (erule parts.induct, auto)
-done
-
-lemma parts_insert_Key [simp]:
-     "parts (insert (Key K) H) = insert (Key K) (parts H)"
-apply (rule parts_insert_eq_I)
-apply (erule parts.induct, auto)
-done
-
-lemma parts_insert_Pan [simp]:
-     "parts (insert (Pan A) H) = insert (Pan A) (parts H)"
-apply (rule parts_insert_eq_I)
-apply (erule parts.induct, auto)
-done
-
-lemma parts_insert_Hash [simp]:
-     "parts (insert (Hash X) H) = insert (Hash X) (parts H)"
-apply (rule parts_insert_eq_I)
-apply (erule parts.induct, auto)
-done
-
-lemma parts_insert_Crypt [simp]:
-     "parts (insert (Crypt K X) H) =
-          insert (Crypt K X) (parts (insert X H))"
-apply (rule equalityI)
-apply (rule subsetI)
-apply (erule parts.induct, auto)
-apply (erule parts.induct)
-apply (blast intro: parts.Body)+
-done
-
-lemma parts_insert_MPair [simp]:
-     "parts (insert {|X,Y|} H) =
-          insert {|X,Y|} (parts (insert X (insert Y H)))"
-apply (rule equalityI)
-apply (rule subsetI)
-apply (erule parts.induct, auto)
-apply (erule parts.induct)
-apply (blast intro: parts.Fst parts.Snd)+
-done
-
-lemma parts_image_Key [simp]: "parts (Key`N) = Key`N"
-apply auto
-apply (erule parts.induct, auto)
-done
-
-lemma parts_image_Pan [simp]: "parts (Pan`A) = Pan`A"
-apply auto
-apply (erule parts.induct, auto)
-done
-
-
-(*In any message, there is an upper bound N on its greatest nonce.*)
-lemma msg_Nonce_supply: "\<exists>N. \<forall>n. N\<le>n --> Nonce n \<notin> parts {msg}"
-apply (induct_tac "msg")
-apply (simp_all (no_asm_simp) add: exI parts_insert2)
-(*MPair case: blast_tac works out the necessary sum itself!*)
-prefer 2 apply (blast elim!: add_leE)
-(*Nonce case*)
-apply (rule_tac x = "N + Suc nat" in exI)
-apply (auto elim!: add_leE)
-done
-
-(* Ditto, for numbers.*)
-lemma msg_Number_supply: "\<exists>N. \<forall>n. N<=n --> Number n \<notin> parts {msg}"
-apply (induct_tac "msg")
-apply (simp_all (no_asm_simp) add: exI parts_insert2)
-prefer 2 apply (blast elim!: add_leE)
-apply (rule_tac x = "N + Suc nat" in exI, auto)
-done
-
-subsection{*Inductive relation "analz"*}
-
-text{*Inductive definition of "analz" -- what can be broken down from a set of
-    messages, including keys.  A form of downward closure.  Pairs can
-    be taken apart; messages decrypted with known keys.*}
-
-inductive_set
-  analz :: "msg set => msg set"
-  for H :: "msg set"
-  where
-    Inj [intro,simp] :    "X \<in> H ==> X \<in> analz H"
-  | Fst:     "{|X,Y|} \<in> analz H ==> X \<in> analz H"
-  | Snd:     "{|X,Y|} \<in> analz H ==> Y \<in> analz H"
-  | Decrypt [dest]:
-             "[|Crypt K X \<in> analz H; Key(invKey K): analz H|] ==> X \<in> analz H"
-
-
-(*Monotonicity; Lemma 1 of Lowe's paper*)
-lemma analz_mono: "G<=H ==> analz(G) <= analz(H)"
-apply auto
-apply (erule analz.induct)
-apply (auto dest: Fst Snd)
-done
-
-text{*Making it safe speeds up proofs*}
-lemma MPair_analz [elim!]:
-     "[| {|X,Y|} \<in> analz H;
-             [| X \<in> analz H; Y \<in> analz H |] ==> P
-          |] ==> P"
-by (blast dest: analz.Fst analz.Snd)
-
-lemma analz_increasing: "H \<subseteq> analz(H)"
-by blast
-
-lemma analz_subset_parts: "analz H \<subseteq> parts H"
-apply (rule subsetI)
-apply (erule analz.induct, blast+)
-done
-
-lemmas analz_into_parts = analz_subset_parts [THEN subsetD, standard]
-
-lemmas not_parts_not_analz = analz_subset_parts [THEN contra_subsetD, standard]
-
-
-lemma parts_analz [simp]: "parts (analz H) = parts H"
-apply (rule equalityI)
-apply (rule analz_subset_parts [THEN parts_mono, THEN subset_trans], simp)
-apply (blast intro: analz_increasing [THEN parts_mono, THEN subsetD])
-done
-
-lemma analz_parts [simp]: "analz (parts H) = parts H"
-apply auto
-apply (erule analz.induct, auto)
-done
-
-lemmas analz_insertI = subset_insertI [THEN analz_mono, THEN [2] rev_subsetD, standard]
-
-subsubsection{*General equational properties*}
-
-lemma analz_empty [simp]: "analz{} = {}"
-apply safe
-apply (erule analz.induct, blast+)
-done
-
-(*Converse fails: we can analz more from the union than from the
-  separate parts, as a key in one might decrypt a message in the other*)
-lemma analz_Un: "analz(G) \<union> analz(H) \<subseteq> analz(G \<union> H)"
-by (intro Un_least analz_mono Un_upper1 Un_upper2)
-
-lemma analz_insert: "insert X (analz H) \<subseteq> analz(insert X H)"
-by (blast intro: analz_mono [THEN [2] rev_subsetD])
-
-subsubsection{*Rewrite rules for pulling out atomic messages*}
-
-lemmas analz_insert_eq_I = equalityI [OF subsetI analz_insert]
-
-lemma analz_insert_Agent [simp]:
-     "analz (insert (Agent agt) H) = insert (Agent agt) (analz H)"
-apply (rule analz_insert_eq_I)
-apply (erule analz.induct, auto)
-done
-
-lemma analz_insert_Nonce [simp]:
-     "analz (insert (Nonce N) H) = insert (Nonce N) (analz H)"
-apply (rule analz_insert_eq_I)
-apply (erule analz.induct, auto)
-done
-
-lemma analz_insert_Number [simp]:
-     "analz (insert (Number N) H) = insert (Number N) (analz H)"
-apply (rule analz_insert_eq_I)
-apply (erule analz.induct, auto)
-done
-
-lemma analz_insert_Hash [simp]:
-     "analz (insert (Hash X) H) = insert (Hash X) (analz H)"
-apply (rule analz_insert_eq_I)
-apply (erule analz.induct, auto)
-done
-
-(*Can only pull out Keys if they are not needed to decrypt the rest*)
-lemma analz_insert_Key [simp]:
-    "K \<notin> keysFor (analz H) ==>
-          analz (insert (Key K) H) = insert (Key K) (analz H)"
-apply (unfold keysFor_def)
-apply (rule analz_insert_eq_I)
-apply (erule analz.induct, auto)
-done
-
-lemma analz_insert_MPair [simp]:
-     "analz (insert {|X,Y|} H) =
-          insert {|X,Y|} (analz (insert X (insert Y H)))"
-apply (rule equalityI)
-apply (rule subsetI)
-apply (erule analz.induct, auto)
-apply (erule analz.induct)
-apply (blast intro: analz.Fst analz.Snd)+
-done
-
-(*Can pull out enCrypted message if the Key is not known*)
-lemma analz_insert_Crypt:
-     "Key (invKey K) \<notin> analz H
-      ==> analz (insert (Crypt K X) H) = insert (Crypt K X) (analz H)"
-apply (rule analz_insert_eq_I)
-apply (erule analz.induct, auto)
-done
-
-lemma analz_insert_Pan [simp]:
-     "analz (insert (Pan A) H) = insert (Pan A) (analz H)"
-apply (rule analz_insert_eq_I)
-apply (erule analz.induct, auto)
-done
-
-lemma lemma1: "Key (invKey K) \<in> analz H ==>
-               analz (insert (Crypt K X) H) \<subseteq>
-               insert (Crypt K X) (analz (insert X H))"
-apply (rule subsetI)
-apply (erule_tac x = x in analz.induct, auto)
-done
-
-lemma lemma2: "Key (invKey K) \<in> analz H ==>
-               insert (Crypt K X) (analz (insert X H)) \<subseteq>
-               analz (insert (Crypt K X) H)"
-apply auto
-apply (erule_tac x = x in analz.induct, auto)
-apply (blast intro: analz_insertI analz.Decrypt)
-done
-
-lemma analz_insert_Decrypt:
-     "Key (invKey K) \<in> analz H ==>
-               analz (insert (Crypt K X) H) =
-               insert (Crypt K X) (analz (insert X H))"
-by (intro equalityI lemma1 lemma2)
-
-(*Case analysis: either the message is secure, or it is not!
-  Effective, but can cause subgoals to blow up!
-  Use with split_if;  apparently split_tac does not cope with patterns
-  such as "analz (insert (Crypt K X) H)" *)
-lemma analz_Crypt_if [simp]:
-     "analz (insert (Crypt K X) H) =
-          (if (Key (invKey K) \<in> analz H)
-           then insert (Crypt K X) (analz (insert X H))
-           else insert (Crypt K X) (analz H))"
-by (simp add: analz_insert_Crypt analz_insert_Decrypt)
-
-
-(*This rule supposes "for the sake of argument" that we have the key.*)
-lemma analz_insert_Crypt_subset:
-     "analz (insert (Crypt K X) H) \<subseteq>
-           insert (Crypt K X) (analz (insert X H))"
-apply (rule subsetI)
-apply (erule analz.induct, auto)
-done
-
-lemma analz_image_Key [simp]: "analz (Key`N) = Key`N"
-apply auto
-apply (erule analz.induct, auto)
-done
-
-lemma analz_image_Pan [simp]: "analz (Pan`A) = Pan`A"
-apply auto
-apply (erule analz.induct, auto)
-done
-
-
-subsubsection{*Idempotence and transitivity*}
-
-lemma analz_analzD [dest!]: "X\<in> analz (analz H) ==> X\<in> analz H"
-by (erule analz.induct, blast+)
-
-lemma analz_idem [simp]: "analz (analz H) = analz H"
-by blast
-
-lemma analz_trans: "[| X\<in> analz G;  G \<subseteq> analz H |] ==> X\<in> analz H"
-by (drule analz_mono, blast)
-
-(*Cut; Lemma 2 of Lowe*)
-lemma analz_cut: "[| Y\<in> analz (insert X H);  X\<in> analz H |] ==> Y\<in> analz H"
-by (erule analz_trans, blast)
-
-(*Cut can be proved easily by induction on
-   "Y: analz (insert X H) ==> X: analz H --> Y: analz H"
-*)
-
-(*This rewrite rule helps in the simplification of messages that involve
-  the forwarding of unknown components (X).  Without it, removing occurrences
-  of X can be very complicated. *)
-lemma analz_insert_eq: "X\<in> analz H ==> analz (insert X H) = analz H"
-by (blast intro: analz_cut analz_insertI)
-
-
-text{*A congruence rule for "analz"*}
-
-lemma analz_subset_cong:
-     "[| analz G \<subseteq> analz G'; analz H \<subseteq> analz H'
-               |] ==> analz (G \<union> H) \<subseteq> analz (G' \<union> H')"
-apply clarify
-apply (erule analz.induct)
-apply (best intro: analz_mono [THEN subsetD])+
-done
-
-lemma analz_cong:
-     "[| analz G = analz G'; analz H = analz H'
-               |] ==> analz (G \<union> H) = analz (G' \<union> H')"
-by (intro equalityI analz_subset_cong, simp_all)
-
-lemma analz_insert_cong:
-     "analz H = analz H' ==> analz(insert X H) = analz(insert X H')"
-by (force simp only: insert_def intro!: analz_cong)
-
-(*If there are no pairs or encryptions then analz does nothing*)
-lemma analz_trivial:
-     "[| \<forall>X Y. {|X,Y|} \<notin> H;  \<forall>X K. Crypt K X \<notin> H |] ==> analz H = H"
-apply safe
-apply (erule analz.induct, blast+)
-done
-
-(*These two are obsolete (with a single Spy) but cost little to prove...*)
-lemma analz_UN_analz_lemma:
-     "X\<in> analz (\<Union>i\<in>A. analz (H i)) ==> X\<in> analz (\<Union>i\<in>A. H i)"
-apply (erule analz.induct)
-apply (blast intro: analz_mono [THEN [2] rev_subsetD])+
-done
-
-lemma analz_UN_analz [simp]: "analz (\<Union>i\<in>A. analz (H i)) = analz (\<Union>i\<in>A. H i)"
-by (blast intro: analz_UN_analz_lemma analz_mono [THEN [2] rev_subsetD])
-
-
-subsection{*Inductive relation "synth"*}
-
-text{*Inductive definition of "synth" -- what can be built up from a set of
-    messages.  A form of upward closure.  Pairs can be built, messages
-    encrypted with known keys.  Agent names are public domain.
-    Numbers can be guessed, but Nonces cannot be.*}
-
-inductive_set
-  synth :: "msg set => msg set"
-  for H :: "msg set"
-  where
-    Inj    [intro]:   "X \<in> H ==> X \<in> synth H"
-  | Agent  [intro]:   "Agent agt \<in> synth H"
-  | Number [intro]:   "Number n  \<in> synth H"
-  | Hash   [intro]:   "X \<in> synth H ==> Hash X \<in> synth H"
-  | MPair  [intro]:   "[|X \<in> synth H;  Y \<in> synth H|] ==> {|X,Y|} \<in> synth H"
-  | Crypt  [intro]:   "[|X \<in> synth H;  Key(K) \<in> H|] ==> Crypt K X \<in> synth H"
-
-(*Monotonicity*)
-lemma synth_mono: "G<=H ==> synth(G) <= synth(H)"
-apply auto
-apply (erule synth.induct)
-apply (auto dest: Fst Snd Body)
-done
-
-(*NO Agent_synth, as any Agent name can be synthesized.  Ditto for Number*)
-inductive_cases Nonce_synth [elim!]: "Nonce n \<in> synth H"
-inductive_cases Key_synth   [elim!]: "Key K \<in> synth H"
-inductive_cases Hash_synth  [elim!]: "Hash X \<in> synth H"
-inductive_cases MPair_synth [elim!]: "{|X,Y|} \<in> synth H"
-inductive_cases Crypt_synth [elim!]: "Crypt K X \<in> synth H"
-inductive_cases Pan_synth   [elim!]: "Pan A \<in> synth H"
-
-
-lemma synth_increasing: "H \<subseteq> synth(H)"
-by blast
-
-subsubsection{*Unions*}
-
-(*Converse fails: we can synth more from the union than from the
-  separate parts, building a compound message using elements of each.*)
-lemma synth_Un: "synth(G) \<union> synth(H) \<subseteq> synth(G \<union> H)"
-by (intro Un_least synth_mono Un_upper1 Un_upper2)
-
-lemma synth_insert: "insert X (synth H) \<subseteq> synth(insert X H)"
-by (blast intro: synth_mono [THEN [2] rev_subsetD])
-
-subsubsection{*Idempotence and transitivity*}
-
-lemma synth_synthD [dest!]: "X\<in> synth (synth H) ==> X\<in> synth H"
-by (erule synth.induct, blast+)
-
-lemma synth_idem: "synth (synth H) = synth H"
-by blast
-
-lemma synth_trans: "[| X\<in> synth G;  G \<subseteq> synth H |] ==> X\<in> synth H"
-by (drule synth_mono, blast)
-
-(*Cut; Lemma 2 of Lowe*)
-lemma synth_cut: "[| Y\<in> synth (insert X H);  X\<in> synth H |] ==> Y\<in> synth H"
-by (erule synth_trans, blast)
-
-lemma Agent_synth [simp]: "Agent A \<in> synth H"
-by blast
-
-lemma Number_synth [simp]: "Number n \<in> synth H"
-by blast
-
-lemma Nonce_synth_eq [simp]: "(Nonce N \<in> synth H) = (Nonce N \<in> H)"
-by blast
-
-lemma Key_synth_eq [simp]: "(Key K \<in> synth H) = (Key K \<in> H)"
-by blast
-
-lemma Crypt_synth_eq [simp]: "Key K \<notin> H ==> (Crypt K X \<in> synth H) = (Crypt K X \<in> H)"
-by blast
-
-lemma Pan_synth_eq [simp]: "(Pan A \<in> synth H) = (Pan A \<in> H)"
-by blast
-
-lemma keysFor_synth [simp]:
-    "keysFor (synth H) = keysFor H \<union> invKey`{K. Key K \<in> H}"
-by (unfold keysFor_def, blast)
-
-
-subsubsection{*Combinations of parts, analz and synth*}
-
-lemma parts_synth [simp]: "parts (synth H) = parts H \<union> synth H"
-apply (rule equalityI)
-apply (rule subsetI)
-apply (erule parts.induct)
-apply (blast intro: synth_increasing [THEN parts_mono, THEN subsetD]
-                    parts.Fst parts.Snd parts.Body)+
-done
-
-lemma analz_analz_Un [simp]: "analz (analz G \<union> H) = analz (G \<union> H)"
-apply (intro equalityI analz_subset_cong)+
-apply simp_all
-done
-
-lemma analz_synth_Un [simp]: "analz (synth G \<union> H) = analz (G \<union> H) \<union> synth G"
-apply (rule equalityI)
-apply (rule subsetI)
-apply (erule analz.induct)
-prefer 5 apply (blast intro: analz_mono [THEN [2] rev_subsetD])
-apply (blast intro: analz.Fst analz.Snd analz.Decrypt)+
-done
-
-lemma analz_synth [simp]: "analz (synth H) = analz H \<union> synth H"
-apply (cut_tac H = "{}" in analz_synth_Un)
-apply (simp (no_asm_use))
-done
-
-
-subsubsection{*For reasoning about the Fake rule in traces*}
-
-lemma parts_insert_subset_Un: "X\<in> G ==> parts(insert X H) \<subseteq> parts G \<union> parts H"
-by (rule subset_trans [OF parts_mono parts_Un_subset2], blast)
-
-(*More specifically for Fake.  Very occasionally we could do with a version
-  of the form  parts{X} \<subseteq> synth (analz H) \<union> parts H *)
-lemma Fake_parts_insert: "X \<in> synth (analz H) ==>
-      parts (insert X H) \<subseteq> synth (analz H) \<union> parts H"
-apply (drule parts_insert_subset_Un)
-apply (simp (no_asm_use))
-apply blast
-done
-
-lemma Fake_parts_insert_in_Un:
-     "[|Z \<in> parts (insert X H);  X: synth (analz H)|] 
-      ==> Z \<in>  synth (analz H) \<union> parts H";
-by (blast dest: Fake_parts_insert [THEN subsetD, dest])
-
-(*H is sometimes (Key ` KK \<union> spies evs), so can't put G=H*)
-lemma Fake_analz_insert:
-     "X\<in> synth (analz G) ==>
-      analz (insert X H) \<subseteq> synth (analz G) \<union> analz (G \<union> H)"
-apply (rule subsetI)
-apply (subgoal_tac "x \<in> analz (synth (analz G) \<union> H) ")
-prefer 2 apply (blast intro: analz_mono [THEN [2] rev_subsetD] analz_mono [THEN synth_mono, THEN [2] rev_subsetD])
-apply (simp (no_asm_use))
-apply blast
-done
-
-lemma analz_conj_parts [simp]:
-     "(X \<in> analz H & X \<in> parts H) = (X \<in> analz H)"
-by (blast intro: analz_subset_parts [THEN subsetD])
-
-lemma analz_disj_parts [simp]:
-     "(X \<in> analz H | X \<in> parts H) = (X \<in> parts H)"
-by (blast intro: analz_subset_parts [THEN subsetD])
-
-(*Without this equation, other rules for synth and analz would yield
-  redundant cases*)
-lemma MPair_synth_analz [iff]:
-     "({|X,Y|} \<in> synth (analz H)) =
-      (X \<in> synth (analz H) & Y \<in> synth (analz H))"
-by blast
-
-lemma Crypt_synth_analz:
-     "[| Key K \<in> analz H;  Key (invKey K) \<in> analz H |]
-       ==> (Crypt K X \<in> synth (analz H)) = (X \<in> synth (analz H))"
-by blast
-
-
-lemma Hash_synth_analz [simp]:
-     "X \<notin> synth (analz H)
-      ==> (Hash{|X,Y|} \<in> synth (analz H)) = (Hash{|X,Y|} \<in> analz H)"
-by blast
-
-
-(*We do NOT want Crypt... messages broken up in protocols!!*)
-declare parts.Body [rule del]
-
-
-text{*Rewrites to push in Key and Crypt messages, so that other messages can
-    be pulled out using the @{text analz_insert} rules*}
-
-lemmas pushKeys [standard] =
-  insert_commute [of "Key K" "Agent C"]
-  insert_commute [of "Key K" "Nonce N"]
-  insert_commute [of "Key K" "Number N"]
-  insert_commute [of "Key K" "Pan PAN"]
-  insert_commute [of "Key K" "Hash X"]
-  insert_commute [of "Key K" "MPair X Y"]
-  insert_commute [of "Key K" "Crypt X K'"]
-
-lemmas pushCrypts [standard] =
-  insert_commute [of "Crypt X K" "Agent C"]
-  insert_commute [of "Crypt X K" "Nonce N"]
-  insert_commute [of "Crypt X K" "Number N"]
-  insert_commute [of "Crypt X K" "Pan PAN"]
-  insert_commute [of "Crypt X K" "Hash X'"]
-  insert_commute [of "Crypt X K" "MPair X' Y"]
-
-text{*Cannot be added with @{text "[simp]"} -- messages should not always be
-  re-ordered.*}
-lemmas pushes = pushKeys pushCrypts
-
-
-subsection{*Tactics useful for many protocol proofs*}
-(*<*)
-ML
-{*
-structure MessageSET =
-struct
-
-(*Prove base case (subgoal i) and simplify others.  A typical base case
-  concerns  Crypt K X \<notin> Key`shrK`bad  and cannot be proved by rewriting
-  alone.*)
-fun prove_simple_subgoals_tac (cs, ss) i =
-    force_tac (cs, ss addsimps [@{thm image_eq_UN}]) i THEN
-    ALLGOALS (asm_simp_tac ss)
-
-(*Analysis of Fake cases.  Also works for messages that forward unknown parts,
-  but this application is no longer necessary if analz_insert_eq is used.
-  Abstraction over i is ESSENTIAL: it delays the dereferencing of claset
-  DEPENDS UPON "X" REFERRING TO THE FRADULENT MESSAGE *)
-
-fun impOfSubs th = th RSN (2, @{thm rev_subsetD})
-
-(*Apply rules to break down assumptions of the form
-  Y \<in> parts(insert X H)  and  Y \<in> analz(insert X H)
-*)
-val Fake_insert_tac =
-    dresolve_tac [impOfSubs @{thm Fake_analz_insert},
-                  impOfSubs @{thm Fake_parts_insert}] THEN'
-    eresolve_tac [asm_rl, @{thm synth.Inj}];
-
-fun Fake_insert_simp_tac ss i =
-    REPEAT (Fake_insert_tac i) THEN asm_full_simp_tac ss i;
-
-fun atomic_spy_analz_tac (cs,ss) = SELECT_GOAL
-    (Fake_insert_simp_tac ss 1
-     THEN
-     IF_UNSOLVED (Blast.depth_tac
-                  (cs addIs [@{thm analz_insertI},
-                                   impOfSubs @{thm analz_subset_parts}]) 4 1))
-
-fun spy_analz_tac (cs,ss) i =
-  DETERM
-   (SELECT_GOAL
-     (EVERY
-      [  (*push in occurrences of X...*)
-       (REPEAT o CHANGED)
-           (res_inst_tac (Simplifier.the_context ss)
-             [(("x", 1), "X")] (insert_commute RS ssubst) 1),
-       (*...allowing further simplifications*)
-       simp_tac ss 1,
-       REPEAT (FIRSTGOAL (resolve_tac [allI,impI,notI,conjI,iffI])),
-       DEPTH_SOLVE (atomic_spy_analz_tac (cs,ss) 1)]) i)
-
-end
-*}
-(*>*)
-
-
-(*By default only o_apply is built-in.  But in the presence of eta-expansion
-  this means that some terms displayed as (f o g) will be rewritten, and others
-  will not!*)
-declare o_def [simp]
-
-
-lemma Crypt_notin_image_Key [simp]: "Crypt K X \<notin> Key ` A"
-by auto
-
-lemma Hash_notin_image_Key [simp] :"Hash X \<notin> Key ` A"
-by auto
-
-lemma synth_analz_mono: "G<=H ==> synth (analz(G)) <= synth (analz(H))"
-by (simp add: synth_mono analz_mono)
-
-lemma Fake_analz_eq [simp]:
-     "X \<in> synth(analz H) ==> synth (analz (insert X H)) = synth (analz H)"
-apply (drule Fake_analz_insert[of _ _ "H"])
-apply (simp add: synth_increasing[THEN Un_absorb2])
-apply (drule synth_mono)
-apply (simp add: synth_idem)
-apply (blast intro: synth_analz_mono [THEN [2] rev_subsetD])
-done
-
-text{*Two generalizations of @{text analz_insert_eq}*}
-lemma gen_analz_insert_eq [rule_format]:
-     "X \<in> analz H ==> ALL G. H \<subseteq> G --> analz (insert X G) = analz G";
-by (blast intro: analz_cut analz_insertI analz_mono [THEN [2] rev_subsetD])
-
-lemma synth_analz_insert_eq [rule_format]:
-     "X \<in> synth (analz H)
-      ==> ALL G. H \<subseteq> G --> (Key K \<in> analz (insert X G)) = (Key K \<in> analz G)";
-apply (erule synth.induct)
-apply (simp_all add: gen_analz_insert_eq subset_trans [OF _ subset_insertI])
-done
-
-lemma Fake_parts_sing:
-     "X \<in> synth (analz H) ==> parts{X} \<subseteq> synth (analz H) \<union> parts H";
-apply (rule subset_trans)
- apply (erule_tac [2] Fake_parts_insert)
-apply (simp add: parts_mono)
-done
-
-lemmas Fake_parts_sing_imp_Un = Fake_parts_sing [THEN [2] rev_subsetD]
-
-method_setup spy_analz = {*
-    Scan.succeed (fn ctxt =>
-        SIMPLE_METHOD' (MessageSET.spy_analz_tac (clasimpset_of ctxt))) *}
-    "for proving the Fake case when analz is involved"
-
-method_setup atomic_spy_analz = {*
-    Scan.succeed (fn ctxt =>
-        SIMPLE_METHOD' (MessageSET.atomic_spy_analz_tac (clasimpset_of ctxt))) *}
-    "for debugging spy_analz"
-
-method_setup Fake_insert_simp = {*
-    Scan.succeed (fn ctxt =>
-        SIMPLE_METHOD' (MessageSET.Fake_insert_simp_tac (simpset_of ctxt))) *}
-    "for debugging spy_analz"
-
-end
--- a/src/HOL/SET-Protocol/PublicSET.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,532 +0,0 @@
-(*  Title:      HOL/SET-Protocol/PublicSET.thy
-    Author:     Giampaolo Bella
-    Author:     Fabio Massacci
-    Author:     Lawrence C Paulson
-*)
-
-header{*The Public-Key Theory, Modified for SET*}
-
-theory PublicSET imports EventSET begin
-
-subsection{*Symmetric and Asymmetric Keys*}
-
-text{*definitions influenced by the wish to assign asymmetric keys 
-  - since the beginning - only to RCA and CAs, namely we need a partial 
-  function on type Agent*}
-
-
-text{*The SET specs mention two signature keys for CAs - we only have one*}
-
-consts
-  publicKey :: "[bool, agent] => key"
-    --{*the boolean is TRUE if a signing key*}
-
-syntax
-  pubEK :: "agent => key"
-  pubSK :: "agent => key"
-  priEK :: "agent => key"
-  priSK :: "agent => key"
-
-translations
-  "pubEK"  == "publicKey False"
-  "pubSK"  == "publicKey True"
-
-  (*BEWARE!! priEK, priSK DON'T WORK with inj, range, image, etc.*)
-  "priEK A"  == "invKey (pubEK A)"
-  "priSK A"  == "invKey (pubSK A)"
-
-text{*By freeness of agents, no two agents have the same key. Since
- @{term "True\<noteq>False"}, no agent has the same signing and encryption keys.*}
-
-specification (publicKey)
-  injective_publicKey:
-    "publicKey b A = publicKey c A' ==> b=c & A=A'"
-(*<*)
-   apply (rule exI [of _ "%b A. 2 * nat_of_agent A + (if b then 1 else 0)"]) 
-   apply (auto simp add: inj_on_def inj_nat_of_agent [THEN inj_eq] split: agent.split) 
-   apply (drule_tac f="%x. x mod 2" in arg_cong, simp add: mod_Suc)+
-(*or this, but presburger won't abstract out the function applications
-   apply presburger+
-*)
-   done                       
-(*>*)
-
-axioms
-  (*No private key equals any public key (essential to ensure that private
-    keys are private!) *)
-  privateKey_neq_publicKey [iff]:
-      "invKey (publicKey b A) \<noteq> publicKey b' A'"
-
-declare privateKey_neq_publicKey [THEN not_sym, iff]
-
-  
-subsection{*Initial Knowledge*}
-
-text{*This information is not necessary.  Each protocol distributes any needed
-certificates, and anyway our proofs require a formalization of the Spy's 
-knowledge only.  However, the initial knowledge is as follows:
-   All agents know RCA's public keys;
-   RCA and CAs know their own respective keys;
-   RCA (has already certified and therefore) knows all CAs public keys; 
-   Spy knows all keys of all bad agents.*}
-primrec    
-(*<*)
-  initState_CA:
-    "initState (CA i)  =
-       (if i=0 then Key ` ({priEK RCA, priSK RCA} Un
-                            pubEK ` (range CA) Un pubSK ` (range CA))
-        else {Key (priEK (CA i)), Key (priSK (CA i)),
-              Key (pubEK (CA i)), Key (pubSK (CA i)),
-              Key (pubEK RCA), Key (pubSK RCA)})"
-
-  initState_Cardholder:
-    "initState (Cardholder i)  =    
-       {Key (priEK (Cardholder i)), Key (priSK (Cardholder i)),
-        Key (pubEK (Cardholder i)), Key (pubSK (Cardholder i)),
-        Key (pubEK RCA), Key (pubSK RCA)}"
-
-  initState_Merchant:
-    "initState (Merchant i)  =    
-       {Key (priEK (Merchant i)), Key (priSK (Merchant i)),
-        Key (pubEK (Merchant i)), Key (pubSK (Merchant i)),
-        Key (pubEK RCA), Key (pubSK RCA)}"
-
-  initState_PG:
-    "initState (PG i)  = 
-       {Key (priEK (PG i)), Key (priSK (PG i)),
-        Key (pubEK (PG i)), Key (pubSK (PG i)),
-        Key (pubEK RCA), Key (pubSK RCA)}"
-(*>*)
-  initState_Spy:
-    "initState Spy = Key ` (invKey ` pubEK ` bad Un
-                            invKey ` pubSK ` bad Un
-                            range pubEK Un range pubSK)"
-
-
-text{*Injective mapping from agents to PANs: an agent can have only one card*}
-
-consts  pan :: "agent => nat"
-
-specification (pan)
-  inj_pan: "inj pan"
-  --{*No two agents have the same PAN*}
-(*<*)
-   apply (rule exI [of _ "nat_of_agent"]) 
-   apply (simp add: inj_on_def inj_nat_of_agent [THEN inj_eq]) 
-   done
-(*>*)
-
-declare inj_pan [THEN inj_eq, iff]
-
-consts
-  XOR :: "nat*nat => nat"  --{*no properties are assumed of exclusive-or*}
-
-
-subsection{*Signature Primitives*}
-
-constdefs 
-
- (* Signature = Message + signed Digest *)
-  sign :: "[key, msg]=>msg"
-    "sign K X == {|X, Crypt K (Hash X) |}"
-
- (* Signature Only = signed Digest Only *)
-  signOnly :: "[key, msg]=>msg"
-    "signOnly K X == Crypt K (Hash X)"
-
- (* Signature for Certificates = Message + signed Message *)
-  signCert :: "[key, msg]=>msg"
-    "signCert K X == {|X, Crypt K X |}"
-
- (* Certification Authority's Certificate.
-    Contains agent name, a key, a number specifying the key's target use,
-              a key to sign the entire certificate.
-
-    Should prove if signK=priSK RCA and C=CA i,
-                  then Ka=pubEK i or pubSK i depending on T  ??
- *)
-  cert :: "[agent, key, msg, key] => msg"
-    "cert A Ka T signK == signCert signK {|Agent A, Key Ka, T|}"
-
-
- (* Cardholder's Certificate.
-    Contains a PAN, the certified key Ka, the PANSecret PS,
-    a number specifying the target use for Ka, the signing key signK.
- *)
-  certC :: "[nat, key, nat, msg, key] => msg"
-    "certC PAN Ka PS T signK ==
-     signCert signK {|Hash {|Nonce PS, Pan PAN|}, Key Ka, T|}"
-
-  (*cert and certA have no repeated elements, so they could be translations,
-    but that's tricky and makes proofs slower*)
-
-syntax
-  "onlyEnc" :: msg      
-  "onlySig" :: msg
-  "authCode" :: msg
-
-translations
-  "onlyEnc"   == "Number 0"
-  "onlySig"  == "Number (Suc 0)"
-  "authCode" == "Number (Suc (Suc 0))"
-
-subsection{*Encryption Primitives*}
-
-constdefs
-
-  EXcrypt :: "[key,key,msg,msg] => msg"
-  --{*Extra Encryption*}
-    (*K: the symmetric key   EK: the public encryption key*)
-    "EXcrypt K EK M m ==
-       {|Crypt K {|M, Hash m|}, Crypt EK {|Key K, m|}|}"
-
-  EXHcrypt :: "[key,key,msg,msg] => msg"
-  --{*Extra Encryption with Hashing*}
-    (*K: the symmetric key   EK: the public encryption key*)
-    "EXHcrypt K EK M m ==
-       {|Crypt K {|M, Hash m|}, Crypt EK {|Key K, m, Hash M|}|}"
-
-  Enc :: "[key,key,key,msg] => msg"
-  --{*Simple Encapsulation with SIGNATURE*}
-    (*SK: the sender's signing key
-      K: the symmetric key
-      EK: the public encryption key*)
-    "Enc SK K EK M ==
-       {|Crypt K (sign SK M), Crypt EK (Key K)|}"
-
-  EncB :: "[key,key,key,msg,msg] => msg"
-  --{*Encapsulation with Baggage.  Keys as above, and baggage b.*}
-    "EncB SK K EK M b == 
-       {|Enc SK K EK {|M, Hash b|}, b|}"
-
-
-subsection{*Basic Properties of pubEK, pubSK, priEK and priSK *}
-
-lemma publicKey_eq_iff [iff]:
-     "(publicKey b A = publicKey b' A') = (b=b' & A=A')"
-by (blast dest: injective_publicKey)
-
-lemma privateKey_eq_iff [iff]:
-     "(invKey (publicKey b A) = invKey (publicKey b' A')) = (b=b' & A=A')"
-by auto
-
-lemma not_symKeys_publicKey [iff]: "publicKey b A \<notin> symKeys"
-by (simp add: symKeys_def)
-
-lemma not_symKeys_privateKey [iff]: "invKey (publicKey b A) \<notin> symKeys"
-by (simp add: symKeys_def)
-
-lemma symKeys_invKey_eq [simp]: "K \<in> symKeys ==> invKey K = K"
-by (simp add: symKeys_def)
-
-lemma symKeys_invKey_iff [simp]: "(invKey K \<in> symKeys) = (K \<in> symKeys)"
-by (unfold symKeys_def, auto)
-
-text{*Can be slow (or even loop) as a simprule*}
-lemma symKeys_neq_imp_neq: "(K \<in> symKeys) \<noteq> (K' \<in> symKeys) ==> K \<noteq> K'"
-by blast
-
-text{*These alternatives to @{text symKeys_neq_imp_neq} don't seem any better
-in practice.*}
-lemma publicKey_neq_symKey: "K \<in> symKeys ==> publicKey b A \<noteq> K"
-by blast
-
-lemma symKey_neq_publicKey: "K \<in> symKeys ==> K \<noteq> publicKey b A"
-by blast
-
-lemma privateKey_neq_symKey: "K \<in> symKeys ==> invKey (publicKey b A) \<noteq> K"
-by blast
-
-lemma symKey_neq_privateKey: "K \<in> symKeys ==> K \<noteq> invKey (publicKey b A)"
-by blast
-
-lemma analz_symKeys_Decrypt:
-     "[| Crypt K X \<in> analz H;  K \<in> symKeys;  Key K \<in> analz H |]  
-      ==> X \<in> analz H"
-by auto
-
-
-subsection{*"Image" Equations That Hold for Injective Functions *}
-
-lemma invKey_image_eq [iff]: "(invKey x \<in> invKey`A) = (x\<in>A)"
-by auto
-
-text{*holds because invKey is injective*}
-lemma publicKey_image_eq [iff]:
-     "(publicKey b A \<in> publicKey c ` AS) = (b=c & A\<in>AS)"
-by auto
-
-lemma privateKey_image_eq [iff]:
-     "(invKey (publicKey b A) \<in> invKey ` publicKey c ` AS) = (b=c & A\<in>AS)"
-by auto
-
-lemma privateKey_notin_image_publicKey [iff]:
-     "invKey (publicKey b A) \<notin> publicKey c ` AS"
-by auto
-
-lemma publicKey_notin_image_privateKey [iff]:
-     "publicKey b A \<notin> invKey ` publicKey c ` AS"
-by auto
-
-lemma keysFor_parts_initState [simp]: "keysFor (parts (initState C)) = {}"
-apply (simp add: keysFor_def)
-apply (induct_tac "C")
-apply (auto intro: range_eqI)
-done
-
-text{*for proving @{text new_keys_not_used}*}
-lemma keysFor_parts_insert:
-     "[| K \<in> keysFor (parts (insert X H));  X \<in> synth (analz H) |]  
-      ==> K \<in> keysFor (parts H) | Key (invKey K) \<in> parts H"
-by (force dest!: 
-         parts_insert_subset_Un [THEN keysFor_mono, THEN [2] rev_subsetD]
-         analz_subset_parts [THEN keysFor_mono, THEN [2] rev_subsetD] 
-            intro: analz_into_parts)
-
-lemma Crypt_imp_keysFor [intro]:
-     "[|K \<in> symKeys; Crypt K X \<in> H|] ==> K \<in> keysFor H"
-by (drule Crypt_imp_invKey_keysFor, simp)
-
-text{*Agents see their own private keys!*}
-lemma privateKey_in_initStateCA [iff]:
-     "Key (invKey (publicKey b A)) \<in> initState A"
-by (case_tac "A", auto)
-
-text{*Agents see their own public keys!*}
-lemma publicKey_in_initStateCA [iff]: "Key (publicKey b A) \<in> initState A"
-by (case_tac "A", auto)
-
-text{*RCA sees CAs' public keys! *}
-lemma pubK_CA_in_initState_RCA [iff]:
-     "Key (publicKey b (CA i)) \<in> initState RCA"
-by auto
-
-
-text{*Spy knows all public keys*}
-lemma knows_Spy_pubEK_i [iff]: "Key (publicKey b A) \<in> knows Spy evs"
-apply (induct_tac "evs")
-apply (simp_all add: imageI knows_Cons split add: event.split)
-done
-
-declare knows_Spy_pubEK_i [THEN analz.Inj, iff]
-                            (*needed????*)
-
-text{*Spy sees private keys of bad agents! [and obviously public keys too]*}
-lemma knows_Spy_bad_privateKey [intro!]:
-     "A \<in> bad ==> Key (invKey (publicKey b A)) \<in> knows Spy evs"
-by (rule initState_subset_knows [THEN subsetD], simp)
-
-
-subsection{*Fresh Nonces for Possibility Theorems*}
-
-lemma Nonce_notin_initState [iff]: "Nonce N \<notin> parts (initState B)"
-by (induct_tac "B", auto)
-
-lemma Nonce_notin_used_empty [simp]: "Nonce N \<notin> used []"
-by (simp add: used_Nil)
-
-text{*In any trace, there is an upper bound N on the greatest nonce in use.*}
-lemma Nonce_supply_lemma: "\<exists>N. \<forall>n. N<=n --> Nonce n \<notin> used evs"
-apply (induct_tac "evs")
-apply (rule_tac x = 0 in exI)
-apply (simp_all add: used_Cons split add: event.split, safe)
-apply (rule msg_Nonce_supply [THEN exE], blast elim!: add_leE)+
-done
-
-lemma Nonce_supply1: "\<exists>N. Nonce N \<notin> used evs"
-by (rule Nonce_supply_lemma [THEN exE], blast)
-
-lemma Nonce_supply: "Nonce (@ N. Nonce N \<notin> used evs) \<notin> used evs"
-apply (rule Nonce_supply_lemma [THEN exE])
-apply (rule someI, fast)
-done
-
-
-subsection{*Specialized Methods for Possibility Theorems*}
-
-ML
-{*
-structure PublicSET =
-struct
-
-(*Tactic for possibility theorems*)
-fun possibility_tac ctxt =
-    REPEAT (*omit used_Says so that Nonces start from different traces!*)
-    (ALLGOALS (simp_tac (simpset_of ctxt delsimps [@{thm used_Says}, @{thm used_Notes}]))
-     THEN
-     REPEAT_FIRST (eq_assume_tac ORELSE' 
-                   resolve_tac [refl, conjI, @{thm Nonce_supply}]))
-
-(*For harder protocols (such as SET_CR!), where we have to set up some
-  nonces and keys initially*)
-fun basic_possibility_tac ctxt =
-    REPEAT 
-    (ALLGOALS (asm_simp_tac (simpset_of ctxt setSolver safe_solver))
-     THEN
-     REPEAT_FIRST (resolve_tac [refl, conjI]))
-
-end
-*}
-
-method_setup possibility = {*
-    Scan.succeed (SIMPLE_METHOD o PublicSET.possibility_tac) *}
-    "for proving possibility theorems"
-
-method_setup basic_possibility = {*
-    Scan.succeed (SIMPLE_METHOD o PublicSET.basic_possibility_tac) *}
-    "for proving possibility theorems"
-
-
-subsection{*Specialized Rewriting for Theorems About @{term analz} and Image*}
-
-lemma insert_Key_singleton: "insert (Key K) H = Key ` {K} Un H"
-by blast
-
-lemma insert_Key_image:
-     "insert (Key K) (Key`KK Un C) = Key ` (insert K KK) Un C"
-by blast
-
-text{*Needed for @{text DK_fresh_not_KeyCryptKey}*}
-lemma publicKey_in_used [iff]: "Key (publicKey b A) \<in> used evs"
-by auto
-
-lemma privateKey_in_used [iff]: "Key (invKey (publicKey b A)) \<in> used evs"
-by (blast intro!: initState_into_used)
-
-text{*Reverse the normal simplification of "image" to build up (not break down)
-  the set of keys.  Based on @{text analz_image_freshK_ss}, but simpler.*}
-lemmas analz_image_keys_simps =
-       simp_thms mem_simps --{*these two allow its use with @{text "only:"}*}
-       image_insert [THEN sym] image_Un [THEN sym] 
-       rangeI symKeys_neq_imp_neq
-       insert_Key_singleton insert_Key_image Un_assoc [THEN sym]
-
-
-(*General lemmas proved by Larry*)
-
-subsection{*Controlled Unfolding of Abbreviations*}
-
-text{*A set is expanded only if a relation is applied to it*}
-lemma def_abbrev_simp_relation:
-     "A == B ==> (A \<in> X) = (B \<in> X) &  
-                 (u = A) = (u = B) &  
-                 (A = u) = (B = u)"
-by auto
-
-text{*A set is expanded only if one of the given functions is applied to it*}
-lemma def_abbrev_simp_function:
-     "A == B  
-      ==> parts (insert A X) = parts (insert B X) &  
-          analz (insert A X) = analz (insert B X) &  
-          keysFor (insert A X) = keysFor (insert B X)"
-by auto
-
-subsubsection{*Special Simplification Rules for @{term signCert}*}
-
-text{*Avoids duplicating X and its components!*}
-lemma parts_insert_signCert:
-     "parts (insert (signCert K X) H) =  
-      insert {|X, Crypt K X|} (parts (insert (Crypt K X) H))"
-by (simp add: signCert_def insert_commute [of X])
-
-text{*Avoids a case split! [X is always available]*}
-lemma analz_insert_signCert:
-     "analz (insert (signCert K X) H) =  
-      insert {|X, Crypt K X|} (insert (Crypt K X) (analz (insert X H)))"
-by (simp add: signCert_def insert_commute [of X])
-
-lemma keysFor_insert_signCert: "keysFor (insert (signCert K X) H) = keysFor H"
-by (simp add: signCert_def)
-
-text{*Controlled rewrite rules for @{term signCert}, just the definitions
-  of the others. Encryption primitives are just expanded, despite their huge
-  redundancy!*}
-lemmas abbrev_simps [simp] =
-    parts_insert_signCert analz_insert_signCert keysFor_insert_signCert
-    sign_def     [THEN def_abbrev_simp_relation]
-    sign_def     [THEN def_abbrev_simp_function]
-    signCert_def [THEN def_abbrev_simp_relation]
-    signCert_def [THEN def_abbrev_simp_function]
-    certC_def    [THEN def_abbrev_simp_relation]
-    certC_def    [THEN def_abbrev_simp_function]
-    cert_def     [THEN def_abbrev_simp_relation]
-    cert_def     [THEN def_abbrev_simp_function]
-    EXcrypt_def  [THEN def_abbrev_simp_relation]
-    EXcrypt_def  [THEN def_abbrev_simp_function]
-    EXHcrypt_def [THEN def_abbrev_simp_relation]
-    EXHcrypt_def [THEN def_abbrev_simp_function]
-    Enc_def      [THEN def_abbrev_simp_relation]
-    Enc_def      [THEN def_abbrev_simp_function]
-    EncB_def     [THEN def_abbrev_simp_relation]
-    EncB_def     [THEN def_abbrev_simp_function]
-
-
-subsubsection{*Elimination Rules for Controlled Rewriting *}
-
-lemma Enc_partsE: 
-     "!!R. [|Enc SK K EK M \<in> parts H;  
-             [|Crypt K (sign SK M) \<in> parts H;  
-               Crypt EK (Key K) \<in> parts H|] ==> R|]  
-           ==> R"
-
-by (unfold Enc_def, blast)
-
-lemma EncB_partsE: 
-     "!!R. [|EncB SK K EK M b \<in> parts H;  
-             [|Crypt K (sign SK {|M, Hash b|}) \<in> parts H;  
-               Crypt EK (Key K) \<in> parts H;  
-               b \<in> parts H|] ==> R|]  
-           ==> R"
-by (unfold EncB_def Enc_def, blast)
-
-lemma EXcrypt_partsE: 
-     "!!R. [|EXcrypt K EK M m \<in> parts H;  
-             [|Crypt K {|M, Hash m|} \<in> parts H;  
-               Crypt EK {|Key K, m|} \<in> parts H|] ==> R|]  
-           ==> R"
-by (unfold EXcrypt_def, blast)
-
-
-subsection{*Lemmas to Simplify Expressions Involving @{term analz} *}
-
-lemma analz_knows_absorb:
-     "Key K \<in> analz (knows Spy evs)  
-      ==> analz (Key ` (insert K H) \<union> knows Spy evs) =  
-          analz (Key ` H \<union> knows Spy evs)"
-by (simp add: analz_insert_eq Un_upper2 [THEN analz_mono, THEN subsetD])
-
-lemma analz_knows_absorb2:
-     "Key K \<in> analz (knows Spy evs)  
-      ==> analz (Key ` (insert X (insert K H)) \<union> knows Spy evs) =  
-          analz (Key ` (insert X H) \<union> knows Spy evs)"
-apply (subst insert_commute)
-apply (erule analz_knows_absorb)
-done
-
-lemma analz_insert_subset_eq:
-     "[|X \<in> analz (knows Spy evs);  knows Spy evs \<subseteq> H|]  
-      ==> analz (insert X H) = analz H"
-apply (rule analz_insert_eq)
-apply (blast intro: analz_mono [THEN [2] rev_subsetD])
-done
-
-lemmas analz_insert_simps = 
-         analz_insert_subset_eq Un_upper2
-         subset_insertI [THEN [2] subset_trans] 
-
-
-subsection{*Freshness Lemmas*}
-
-lemma in_parts_Says_imp_used:
-     "[|Key K \<in> parts {X}; Says A B X \<in> set evs|] ==> Key K \<in> used evs"
-by (blast intro: parts_trans dest!: Says_imp_knows_Spy [THEN parts.Inj])
-
-text{*A useful rewrite rule with @{term analz_image_keys_simps}*}
-lemma Crypt_notin_image_Key: "Crypt K X \<notin> Key ` KK"
-by auto
-
-lemma fresh_notin_analz_knows_Spy:
-     "Key K \<notin> used evs ==> Key K \<notin> analz (knows Spy evs)"
-by (auto dest: analz_into_parts)
-
-end
--- a/src/HOL/SET-Protocol/Purchase.thy	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1170 +0,0 @@
-(*  Title:      HOL/SET-Protocol/Purchase.thy
-    Author:     Giampaolo Bella
-    Author:     Fabio Massacci
-    Author:     Lawrence C Paulson
-*)
-
-header{*Purchase Phase of SET*}
-
-theory Purchase imports PublicSET begin
-
-text{*
-Note: nonces seem to consist of 20 bytes.  That includes both freshness
-challenges (Chall-EE, etc.) and important secrets (CardSecret, PANsecret)
-
-This version omits @{text LID_C} but retains @{text LID_M}. At first glance
-(Programmer's Guide page 267) it seems that both numbers are just introduced
-for the respective convenience of the Cardholder's and Merchant's
-system. However, omitting both of them would create a problem of
-identification: how can the Merchant's system know what transaction is it
-supposed to process?
-
-Further reading (Programmer's guide page 309) suggest that there is an outside
-bootstrapping message (SET initiation message) which is used by the Merchant
-and the Cardholder to agree on the actual transaction. This bootstrapping
-message is described in the SET External Interface Guide and ought to generate
-@{text LID_M}. According SET Extern Interface Guide, this number might be a
-cookie, an invoice number etc. The Programmer's Guide on page 310, states that
-in absence of @{text LID_M} the protocol must somehow ("outside SET") identify
-the transaction from OrderDesc, which is assumed to be a searchable text only
-field. Thus, it is assumed that the Merchant or the Cardholder somehow agreed
-out-of-bad on the value of @{text LID_M} (for instance a cookie in a web
-transaction etc.). This out-of-band agreement is expressed with a preliminary
-start action in which the merchant and the Cardholder agree on the appropriate
-values. Agreed values are stored with a suitable notes action.
-
-"XID is a transaction ID that is usually generated by the Merchant system,
-unless there is no PInitRes, in which case it is generated by the Cardholder
-system. It is a randomly generated 20 byte variable that is globally unique
-(statistically). Merchant and Cardholder systems shall use appropriate random
-number generators to ensure the global uniqueness of XID."
---Programmer's Guide, page 267.
-
-PI (Payment Instruction) is the most central and sensitive data structure in
-SET. It is used to pass the data required to authorize a payment card payment
-from the Cardholder to the Payment Gateway, which will use the data to
-initiate a payment card transaction through the traditional payment card
-financial network. The data is encrypted by the Cardholder and sent via the
-Merchant, such that the data is hidden from the Merchant unless the Acquirer
-passes the data back to the Merchant.
---Programmer's Guide, page 271.*}
-
-consts
-
-    CardSecret :: "nat => nat"
-     --{*Maps Cardholders to CardSecrets.
-         A CardSecret of 0 means no cerificate, must use unsigned format.*}
-
-    PANSecret :: "nat => nat"
-     --{*Maps Cardholders to PANSecrets.*}
-
-inductive_set
-  set_pur :: "event list set"
-where
-
-  Nil:   --{*Initial trace is empty*}
-         "[] \<in> set_pur"
-
-| Fake:  --{*The spy MAY say anything he CAN say.*}
-         "[| evsf \<in> set_pur;  X \<in> synth(analz(knows Spy evsf)) |]
-          ==> Says Spy B X  # evsf \<in> set_pur"
-
-
-| Reception: --{*If A sends a message X to B, then B might receive it*}
-             "[| evsr \<in> set_pur;  Says A B X \<in> set evsr |]
-              ==> Gets B X  # evsr \<in> set_pur"
-
-| Start: 
-      --{*Added start event which is out-of-band for SET: the Cardholder and
-          the merchant agree on the amounts and uses @{text LID_M} as an
-          identifier.
-          This is suggested by the External Interface Guide. The Programmer's
-          Guide, in absence of @{text LID_M}, states that the merchant uniquely
-          identifies the order out of some data contained in OrderDesc.*}
-   "[|evsStart \<in> set_pur;
-      Number LID_M \<notin> used evsStart;
-      C = Cardholder k; M = Merchant i; P = PG j;
-      Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
-      LID_M \<notin> range CardSecret;
-      LID_M \<notin> range PANSecret |]
-     ==> Notes C {|Number LID_M, Transaction|}
-       # Notes M {|Number LID_M, Agent P, Transaction|}
-       # evsStart \<in> set_pur"
-
-| PInitReq:
-     --{*Purchase initialization, page 72 of Formal Protocol Desc.*}
-   "[|evsPIReq \<in> set_pur;
-      Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
-      Nonce Chall_C \<notin> used evsPIReq;
-      Chall_C \<notin> range CardSecret; Chall_C \<notin> range PANSecret;
-      Notes C {|Number LID_M, Transaction |} \<in> set evsPIReq |]
-    ==> Says C M {|Number LID_M, Nonce Chall_C|} # evsPIReq \<in> set_pur"
-
-| PInitRes:
-     --{*Merchant replies with his own label XID and the encryption
-         key certificate of his chosen Payment Gateway. Page 74 of Formal
-         Protocol Desc. We use @{text LID_M} to identify Cardholder*}
-   "[|evsPIRes \<in> set_pur;
-      Gets M {|Number LID_M, Nonce Chall_C|} \<in> set evsPIRes;
-      Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
-      Notes M {|Number LID_M, Agent P, Transaction|} \<in> set evsPIRes;
-      Nonce Chall_M \<notin> used evsPIRes;
-      Chall_M \<notin> range CardSecret; Chall_M \<notin> range PANSecret;
-      Number XID \<notin> used evsPIRes;
-      XID \<notin> range CardSecret; XID \<notin> range PANSecret|]
-    ==> Says M C (sign (priSK M)
-                       {|Number LID_M, Number XID,
-                         Nonce Chall_C, Nonce Chall_M,
-                         cert P (pubEK P) onlyEnc (priSK RCA)|})
-          # evsPIRes \<in> set_pur"
-
-| PReqUns:
-      --{*UNSIGNED Purchase request (CardSecret = 0).
-        Page 79 of Formal Protocol Desc.
-        Merchant never sees the amount in clear. This holds of the real
-        protocol, where XID identifies the transaction. We omit
-        Hash{|Number XID, Nonce (CardSecret k)|} from PIHead because
-        the CardSecret is 0 and because AuthReq treated the unsigned case
-        very differently from the signed one anyway.*}
-   "!!Chall_C Chall_M OrderDesc P PurchAmt XID evsPReqU.
-    [|evsPReqU \<in> set_pur;
-      C = Cardholder k; CardSecret k = 0;
-      Key KC1 \<notin> used evsPReqU;  KC1 \<in> symKeys;
-      Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
-      HOD = Hash{|Number OrderDesc, Number PurchAmt|};
-      OIData = {|Number LID_M, Number XID, Nonce Chall_C, HOD,Nonce Chall_M|};
-      PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M|};
-      Gets C (sign (priSK M)
-                   {|Number LID_M, Number XID,
-                     Nonce Chall_C, Nonce Chall_M,
-                     cert P EKj onlyEnc (priSK RCA)|})
-        \<in> set evsPReqU;
-      Says C M {|Number LID_M, Nonce Chall_C|} \<in> set evsPReqU;
-      Notes C {|Number LID_M, Transaction|} \<in> set evsPReqU |]
-    ==> Says C M
-             {|EXHcrypt KC1 EKj {|PIHead, Hash OIData|} (Pan (pan C)),
-               OIData, Hash{|PIHead, Pan (pan C)|} |}
-          # Notes C {|Key KC1, Agent M|}
-          # evsPReqU \<in> set_pur"
-
-| PReqS:
-      --{*SIGNED Purchase request.  Page 77 of Formal Protocol Desc.
-          We could specify the equation
-          @{term "PIReqSigned = {| PIDualSigned, OIDualSigned |}"}, since the
-          Formal Desc. gives PIHead the same format in the unsigned case.
-          However, there's little point, as P treats the signed and 
-          unsigned cases differently.*}
-   "!!C Chall_C Chall_M EKj HOD KC2 LID_M M OIData
-      OIDualSigned OrderDesc P PANData PIData PIDualSigned
-      PIHead PurchAmt Transaction XID evsPReqS k.
-    [|evsPReqS \<in> set_pur;
-      C = Cardholder k;
-      CardSecret k \<noteq> 0;  Key KC2 \<notin> used evsPReqS;  KC2 \<in> symKeys;
-      Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
-      HOD = Hash{|Number OrderDesc, Number PurchAmt|};
-      OIData = {|Number LID_M, Number XID, Nonce Chall_C, HOD, Nonce Chall_M|};
-      PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M,
-                  Hash{|Number XID, Nonce (CardSecret k)|}|};
-      PANData = {|Pan (pan C), Nonce (PANSecret k)|};
-      PIData = {|PIHead, PANData|};
-      PIDualSigned = {|sign (priSK C) {|Hash PIData, Hash OIData|},
-                       EXcrypt KC2 EKj {|PIHead, Hash OIData|} PANData|};
-      OIDualSigned = {|OIData, Hash PIData|};
-      Gets C (sign (priSK M)
-                   {|Number LID_M, Number XID,
-                     Nonce Chall_C, Nonce Chall_M,
-                     cert P EKj onlyEnc (priSK RCA)|})
-        \<in> set evsPReqS;
-      Says C M {|Number LID_M, Nonce Chall_C|} \<in> set evsPReqS;
-      Notes C {|Number LID_M, Transaction|} \<in> set evsPReqS |]
-    ==> Says C M {|PIDualSigned, OIDualSigned|}
-          # Notes C {|Key KC2, Agent M|}
-          # evsPReqS \<in> set_pur"
-
-  --{*Authorization Request.  Page 92 of Formal Protocol Desc.
-    Sent in response to Purchase Request.*}
-| AuthReq:
-   "[| evsAReq \<in> set_pur;
-       Key KM \<notin> used evsAReq;  KM \<in> symKeys;
-       Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
-       HOD = Hash{|Number OrderDesc, Number PurchAmt|};
-       OIData = {|Number LID_M, Number XID, Nonce Chall_C, HOD,
-                  Nonce Chall_M|};
-       CardSecret k \<noteq> 0 -->
-         P_I = {|sign (priSK C) {|HPIData, Hash OIData|}, encPANData|};
-       Gets M {|P_I, OIData, HPIData|} \<in> set evsAReq;
-       Says M C (sign (priSK M) {|Number LID_M, Number XID,
-                                  Nonce Chall_C, Nonce Chall_M,
-                                  cert P EKj onlyEnc (priSK RCA)|})
-         \<in> set evsAReq;
-        Notes M {|Number LID_M, Agent P, Transaction|}
-           \<in> set evsAReq |]
-    ==> Says M P
-             (EncB (priSK M) KM (pubEK P)
-               {|Number LID_M, Number XID, Hash OIData, HOD|}   P_I)
-          # evsAReq \<in> set_pur"
-
-  --{*Authorization Response has two forms: for UNSIGNED and SIGNED PIs.
-    Page 99 of Formal Protocol Desc.
-    PI is a keyword (product!), so we call it @{text P_I}. The hashes HOD and
-    HOIData occur independently in @{text P_I} and in M's message.
-    The authCode in AuthRes represents the baggage of EncB, which in the
-    full protocol is [CapToken], [AcqCardMsg], [AuthToken]:
-    optional items for split shipments, recurring payments, etc.*}
-
-| AuthResUns:
-    --{*Authorization Response, UNSIGNED*}
-   "[| evsAResU \<in> set_pur;
-       C = Cardholder k; M = Merchant i;
-       Key KP \<notin> used evsAResU;  KP \<in> symKeys;
-       CardSecret k = 0;  KC1 \<in> symKeys;  KM \<in> symKeys;
-       PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M|};
-       P_I = EXHcrypt KC1 EKj {|PIHead, HOIData|} (Pan (pan C));
-       Gets P (EncB (priSK M) KM (pubEK P)
-               {|Number LID_M, Number XID, HOIData, HOD|} P_I)
-           \<in> set evsAResU |]
-   ==> Says P M
-            (EncB (priSK P) KP (pubEK M)
-              {|Number LID_M, Number XID, Number PurchAmt|}
-              authCode)
-       # evsAResU \<in> set_pur"
-
-| AuthResS:
-    --{*Authorization Response, SIGNED*}
-   "[| evsAResS \<in> set_pur;
-       C = Cardholder k;
-       Key KP \<notin> used evsAResS;  KP \<in> symKeys;
-       CardSecret k \<noteq> 0;  KC2 \<in> symKeys;  KM \<in> symKeys;
-       P_I = {|sign (priSK C) {|Hash PIData, HOIData|},
-               EXcrypt KC2 (pubEK P) {|PIHead, HOIData|} PANData|};
-       PANData = {|Pan (pan C), Nonce (PANSecret k)|};
-       PIData = {|PIHead, PANData|};
-       PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M,
-                  Hash{|Number XID, Nonce (CardSecret k)|}|};
-       Gets P (EncB (priSK M) KM (pubEK P)
-                {|Number LID_M, Number XID, HOIData, HOD|}
-               P_I)
-           \<in> set evsAResS |]
-   ==> Says P M
-            (EncB (priSK P) KP (pubEK M)
-              {|Number LID_M, Number XID, Number PurchAmt|}
-              authCode)
-       # evsAResS \<in> set_pur"
-
-| PRes:
-    --{*Purchase response.*}
-   "[| evsPRes \<in> set_pur;  KP \<in> symKeys;  M = Merchant i;
-       Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
-       Gets M (EncB (priSK P) KP (pubEK M)
-              {|Number LID_M, Number XID, Number PurchAmt|}
-              authCode)
-          \<in> set evsPRes;
-       Gets M {|Number LID_M, Nonce Chall_C|} \<in> set evsPRes;
-       Says M P
-            (EncB (priSK M) KM (pubEK P)
-              {|Number LID_M, Number XID, Hash OIData, HOD|} P_I)
-         \<in> set evsPRes;
-       Notes M {|Number LID_M, Agent P, Transaction|}
-          \<in> set evsPRes
-      |]
-   ==> Says M C
-         (sign (priSK M) {|Number LID_M, Number XID, Nonce Chall_C,
-                           Hash (Number PurchAmt)|})
-         # evsPRes \<in> set_pur"
-
-
-specification (CardSecret PANSecret)
-  inj_CardSecret:  "inj CardSecret"
-  inj_PANSecret:   "inj PANSecret"
-  CardSecret_neq_PANSecret: "CardSecret k \<noteq> PANSecret k'"
-    --{*No CardSecret equals any PANSecret*}
-  apply (rule_tac x="curry nat2_to_nat 0" in exI)
-  apply (rule_tac x="curry nat2_to_nat 1" in exI)
-  apply (simp add: nat2_to_nat_inj [THEN inj_eq] inj_on_def)
-  done
-
-declare Says_imp_knows_Spy [THEN parts.Inj, dest]
-declare parts.Body [dest]
-declare analz_into_parts [dest]
-declare Fake_parts_insert_in_Un [dest]
-
-declare CardSecret_neq_PANSecret [iff] 
-        CardSecret_neq_PANSecret [THEN not_sym, iff]
-declare inj_CardSecret [THEN inj_eq, iff] 
-        inj_PANSecret [THEN inj_eq, iff]
-
-
-subsection{*Possibility Properties*}
-
-lemma Says_to_Gets:
-     "Says A B X # evs \<in> set_pur ==> Gets B X # Says A B X # evs \<in> set_pur"
-by (rule set_pur.Reception, auto)
-
-text{*Possibility for UNSIGNED purchases. Note that we need to ensure
-that XID differs from OrderDesc and PurchAmt, since it is supposed to be
-a unique number!*}
-lemma possibility_Uns:
-    "[| CardSecret k = 0;
-        C = Cardholder k;  M = Merchant i;
-        Key KC \<notin> used []; Key KM \<notin> used []; Key KP \<notin> used []; 
-        KC \<in> symKeys; KM \<in> symKeys; KP \<in> symKeys; 
-        KC < KM; KM < KP;
-        Nonce Chall_C \<notin> used []; Chall_C \<notin> range CardSecret \<union> range PANSecret;
-        Nonce Chall_M \<notin> used []; Chall_M \<notin> range CardSecret \<union> range PANSecret;
-        Chall_C < Chall_M; 
-        Number LID_M \<notin> used []; LID_M \<notin> range CardSecret \<union> range PANSecret;
-        Number XID \<notin> used []; XID \<notin> range CardSecret \<union> range PANSecret;
-        LID_M < XID; XID < OrderDesc; OrderDesc < PurchAmt |] 
-   ==> \<exists>evs \<in> set_pur.
-          Says M C
-               (sign (priSK M)
-                    {|Number LID_M, Number XID, Nonce Chall_C, 
-                      Hash (Number PurchAmt)|})
-                  \<in> set evs" 
-apply (intro exI bexI)
-apply (rule_tac [2]
-        set_pur.Nil
-         [THEN set_pur.Start [of _ LID_M C k M i _ _ _ OrderDesc PurchAmt], 
-          THEN set_pur.PInitReq [of concl: C M LID_M Chall_C],
-          THEN Says_to_Gets, 
-          THEN set_pur.PInitRes [of concl: M C LID_M XID Chall_C Chall_M], 
-          THEN Says_to_Gets,
-          THEN set_pur.PReqUns [of concl: C M KC],
-          THEN Says_to_Gets, 
-          THEN set_pur.AuthReq [of concl: M "PG j" KM LID_M XID], 
-          THEN Says_to_Gets, 
-          THEN set_pur.AuthResUns [of concl: "PG j" M KP LID_M XID],
-          THEN Says_to_Gets, 
-          THEN set_pur.PRes]) 
-apply basic_possibility
-apply (simp_all add: used_Cons symKeys_neq_imp_neq) 
-done
-
-lemma possibility_S:
-    "[| CardSecret k \<noteq> 0;
-        C = Cardholder k;  M = Merchant i;
-        Key KC \<notin> used []; Key KM \<notin> used []; Key KP \<notin> used []; 
-        KC \<in> symKeys; KM \<in> symKeys; KP \<in> symKeys; 
-        KC < KM; KM < KP;
-        Nonce Chall_C \<notin> used []; Chall_C \<notin> range CardSecret \<union> range PANSecret;
-        Nonce Chall_M \<notin> used []; Chall_M \<notin> range CardSecret \<union> range PANSecret;
-        Chall_C < Chall_M; 
-        Number LID_M \<notin> used []; LID_M \<notin> range CardSecret \<union> range PANSecret;
-        Number XID \<notin> used []; XID \<notin> range CardSecret \<union> range PANSecret;
-        LID_M < XID; XID < OrderDesc; OrderDesc < PurchAmt |] 
-   ==>  \<exists>evs \<in> set_pur.
-            Says M C
-                 (sign (priSK M) {|Number LID_M, Number XID, Nonce Chall_C, 
-                                   Hash (Number PurchAmt)|})
-               \<in> set evs"
-apply (intro exI bexI)
-apply (rule_tac [2]
-        set_pur.Nil
-         [THEN set_pur.Start [of _ LID_M C k M i _ _ _ OrderDesc PurchAmt], 
-          THEN set_pur.PInitReq [of concl: C M LID_M Chall_C],
-          THEN Says_to_Gets, 
-          THEN set_pur.PInitRes [of concl: M C LID_M XID Chall_C Chall_M], 
-          THEN Says_to_Gets,
-          THEN set_pur.PReqS [of concl: C M _ _ KC],
-          THEN Says_to_Gets, 
-          THEN set_pur.AuthReq [of concl: M "PG j" KM LID_M XID], 
-          THEN Says_to_Gets, 
-          THEN set_pur.AuthResS [of concl: "PG j" M KP LID_M XID],
-          THEN Says_to_Gets, 
-          THEN set_pur.PRes]) 
-apply basic_possibility
-apply (auto simp add: used_Cons symKeys_neq_imp_neq) 
-done
-
-text{*General facts about message reception*}
-lemma Gets_imp_Says:
-     "[| Gets B X \<in> set evs; evs \<in> set_pur |]
-   ==> \<exists>A. Says A B X \<in> set evs"
-apply (erule rev_mp)
-apply (erule set_pur.induct, auto)
-done
-
-lemma Gets_imp_knows_Spy:
-     "[| Gets B X \<in> set evs; evs \<in> set_pur |]  ==> X \<in> knows Spy evs"
-by (blast dest!: Gets_imp_Says Says_imp_knows_Spy)
-
-declare Gets_imp_knows_Spy [THEN parts.Inj, dest]
-
-text{*Forwarding lemmas, to aid simplification*}
-
-lemma AuthReq_msg_in_parts_spies:
-     "[|Gets M {|P_I, OIData, HPIData|} \<in> set evs;
-        evs \<in> set_pur|] ==> P_I \<in> parts (knows Spy evs)"
-by auto
-
-lemma AuthReq_msg_in_analz_spies:
-     "[|Gets M {|P_I, OIData, HPIData|} \<in> set evs;
-        evs \<in> set_pur|] ==> P_I \<in> analz (knows Spy evs)"
-by (blast dest: Gets_imp_knows_Spy [THEN analz.Inj])
-
-
-subsection{*Proofs on Asymmetric Keys*}
-
-text{*Private Keys are Secret*}
-
-text{*Spy never sees an agent's private keys! (unless it's bad at start)*}
-lemma Spy_see_private_Key [simp]:
-     "evs \<in> set_pur
-      ==> (Key(invKey (publicKey b A)) \<in> parts(knows Spy evs)) = (A \<in> bad)"
-apply (erule set_pur.induct)
-apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
-apply auto
-done
-declare Spy_see_private_Key [THEN [2] rev_iffD1, dest!]
-
-lemma Spy_analz_private_Key [simp]:
-     "evs \<in> set_pur ==>
-     (Key(invKey (publicKey b A)) \<in> analz(knows Spy evs)) = (A \<in> bad)"
-by auto
-declare Spy_analz_private_Key [THEN [2] rev_iffD1, dest!]
-
-text{*rewriting rule for priEK's*}
-lemma parts_image_priEK:
-     "[|Key (priEK C) \<in> parts (Key`KK Un (knows Spy evs));
-        evs \<in> set_pur|] ==> priEK C \<in> KK | C \<in> bad"
-by auto
-
-text{*trivial proof because @{term"priEK C"} never appears even in
-  @{term "parts evs"}. *}
-lemma analz_image_priEK:
-     "evs \<in> set_pur ==>
-          (Key (priEK C) \<in> analz (Key`KK Un (knows Spy evs))) =
-          (priEK C \<in> KK | C \<in> bad)"
-by (blast dest!: parts_image_priEK intro: analz_mono [THEN [2] rev_subsetD])
-
-
-subsection{*Public Keys in Certificates are Correct*}
-
-lemma Crypt_valid_pubEK [dest!]:
-     "[| Crypt (priSK RCA) {|Agent C, Key EKi, onlyEnc|}
-           \<in> parts (knows Spy evs);
-         evs \<in> set_pur |] ==> EKi = pubEK C"
-by (erule rev_mp, erule set_pur.induct, auto)
-
-lemma Crypt_valid_pubSK [dest!]:
-     "[| Crypt (priSK RCA) {|Agent C, Key SKi, onlySig|}
-           \<in> parts (knows Spy evs);
-         evs \<in> set_pur |] ==> SKi = pubSK C"
-by (erule rev_mp, erule set_pur.induct, auto)
-
-lemma certificate_valid_pubEK:
-    "[| cert C EKi onlyEnc (priSK RCA) \<in> parts (knows Spy evs);
-        evs \<in> set_pur |]
-     ==> EKi = pubEK C"
-by (unfold cert_def signCert_def, auto)
-
-lemma certificate_valid_pubSK:
-    "[| cert C SKi onlySig (priSK RCA) \<in> parts (knows Spy evs);
-        evs \<in> set_pur |] ==> SKi = pubSK C"
-by (unfold cert_def signCert_def, auto)
-
-lemma Says_certificate_valid [simp]:
-     "[| Says A B (sign SK {|lid, xid, cc, cm,
-                           cert C EK onlyEnc (priSK RCA)|}) \<in> set evs;
-         evs \<in> set_pur |]
-      ==> EK = pubEK C"
-by (unfold sign_def, auto)
-
-lemma Gets_certificate_valid [simp]:
-     "[| Gets A (sign SK {|lid, xid, cc, cm,
-                           cert C EK onlyEnc (priSK RCA)|}) \<in> set evs;
-         evs \<in> set_pur |]
-      ==> EK = pubEK C"
-by (frule Gets_imp_Says, auto)
-
-method_setup valid_certificate_tac = {*
-  Args.goal_spec >> (fn quant =>
-    K (SIMPLE_METHOD'' quant (fn i =>
-      EVERY [ftac @{thm Gets_certificate_valid} i,
-             assume_tac i, REPEAT (hyp_subst_tac i)])))
-*} ""
-
-
-subsection{*Proofs on Symmetric Keys*}
-
-text{*Nobody can have used non-existent keys!*}
-lemma new_keys_not_used [rule_format,simp]:
-     "evs \<in> set_pur
-      ==> Key K \<notin> used evs --> K \<in> symKeys -->
-          K \<notin> keysFor (parts (knows Spy evs))"
-apply (erule set_pur.induct)
-apply (valid_certificate_tac [8]) --{*PReqS*}
-apply (valid_certificate_tac [7]) --{*PReqUns*}
-apply auto
-apply (force dest!: usedI keysFor_parts_insert) --{*Fake*}
-done
-
-lemma new_keys_not_analzd:
-     "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_pur |]
-      ==> K \<notin> keysFor (analz (knows Spy evs))"
-by (blast intro: keysFor_mono [THEN [2] rev_subsetD] dest: new_keys_not_used)
-
-lemma Crypt_parts_imp_used:
-     "[|Crypt K X \<in> parts (knows Spy evs);
-        K \<in> symKeys; evs \<in> set_pur |] ==> Key K \<in> used evs"
-apply (rule ccontr)
-apply (force dest: new_keys_not_used Crypt_imp_invKey_keysFor)
-done
-
-lemma Crypt_analz_imp_used:
-     "[|Crypt K X \<in> analz (knows Spy evs);
-        K \<in> symKeys; evs \<in> set_pur |] ==> Key K \<in> used evs"
-by (blast intro: Crypt_parts_imp_used)
-
-text{*New versions: as above, but generalized to have the KK argument*}
-
-lemma gen_new_keys_not_used:
-     "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_pur |]
-      ==> Key K \<notin> used evs --> K \<in> symKeys -->
-          K \<notin> keysFor (parts (Key`KK Un knows Spy evs))"
-by auto
-
-lemma gen_new_keys_not_analzd:
-     "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_pur |]
-      ==> K \<notin> keysFor (analz (Key`KK Un knows Spy evs))"
-by (blast intro: keysFor_mono [THEN subsetD] dest: gen_new_keys_not_used)
-
-lemma analz_Key_image_insert_eq:
-     "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_pur |]
-      ==> analz (Key ` (insert K KK) \<union> knows Spy evs) =
-          insert (Key K) (analz (Key ` KK \<union> knows Spy evs))"
-by (simp add: gen_new_keys_not_analzd)
-
-
-subsection{*Secrecy of Symmetric Keys*}
-
-lemma Key_analz_image_Key_lemma:
-     "P --> (Key K \<in> analz (Key`KK Un H)) --> (K\<in>KK | Key K \<in> analz H)
-      ==>
-      P --> (Key K \<in> analz (Key`KK Un H)) = (K\<in>KK | Key K \<in> analz H)"
-by (blast intro: analz_mono [THEN [2] rev_subsetD])
-
-
-lemma symKey_compromise:
-     "evs \<in> set_pur \<Longrightarrow>
-      (\<forall>SK KK. SK \<in> symKeys \<longrightarrow>
-        (\<forall>K \<in> KK. K \<notin> range(\<lambda>C. priEK C)) \<longrightarrow>
-               (Key SK \<in> analz (Key`KK \<union> (knows Spy evs))) =
-               (SK \<in> KK \<or> Key SK \<in> analz (knows Spy evs)))"
-apply (erule set_pur.induct)
-apply (rule_tac [!] allI)+
-apply (rule_tac [!] impI [THEN Key_analz_image_Key_lemma, THEN impI])+
-apply (frule_tac [9] AuthReq_msg_in_analz_spies) --{*AReq*}
-apply (valid_certificate_tac [8]) --{*PReqS*}
-apply (valid_certificate_tac [7]) --{*PReqUns*}
-apply (simp_all
-         del: image_insert image_Un imp_disjL
-         add: analz_image_keys_simps disj_simps
-              analz_Key_image_insert_eq notin_image_iff
-              analz_insert_simps analz_image_priEK)
-  --{*8 seconds on a 1.6GHz machine*}
-apply spy_analz --{*Fake*}
-apply (blast elim!: ballE)+ --{*PReq: unsigned and signed*}
-done
-
-
-
-subsection{*Secrecy of Nonces*}
-
-text{*As usual: we express the property as a logical equivalence*}
-lemma Nonce_analz_image_Key_lemma:
-     "P --> (Nonce N \<in> analz (Key`KK Un H)) --> (Nonce N \<in> analz H)
-      ==> P --> (Nonce N \<in> analz (Key`KK Un H)) = (Nonce N \<in> analz H)"
-by (blast intro: analz_mono [THEN [2] rev_subsetD])
-
-text{*The @{text "(no_asm)"} attribute is essential, since it retains
-  the quantifier and allows the simprule's condition to itself be simplified.*}
-lemma Nonce_compromise [rule_format (no_asm)]:
-     "evs \<in> set_pur ==>
-      (\<forall>N KK. (\<forall>K \<in> KK. K \<notin> range(\<lambda>C. priEK C))   -->
-              (Nonce N \<in> analz (Key`KK \<union> (knows Spy evs))) =
-              (Nonce N \<in> analz (knows Spy evs)))"
-apply (erule set_pur.induct)
-apply (rule_tac [!] allI)+
-apply (rule_tac [!] impI [THEN Nonce_analz_image_Key_lemma])+
-apply (frule_tac [9] AuthReq_msg_in_analz_spies) --{*AReq*}
-apply (valid_certificate_tac [8]) --{*PReqS*}
-apply (valid_certificate_tac [7]) --{*PReqUns*}
-apply (simp_all
-         del: image_insert image_Un imp_disjL
-         add: analz_image_keys_simps disj_simps symKey_compromise
-              analz_Key_image_insert_eq notin_image_iff
-              analz_insert_simps analz_image_priEK)
-  --{*8 seconds on a 1.6GHz machine*}
-apply spy_analz --{*Fake*}
-apply (blast elim!: ballE) --{*PReqS*}
-done
-
-lemma PANSecret_notin_spies:
-     "[|Nonce (PANSecret k) \<in> analz (knows Spy evs);  evs \<in> set_pur|]
-      ==> 
-       (\<exists>V W X Y KC2 M. \<exists>P \<in> bad.
-          Says (Cardholder k) M
-               {|{|W, EXcrypt KC2 (pubEK P) X {|Y, Nonce (PANSecret k)|}|},
-                 V|}  \<in>  set evs)"
-apply (erule rev_mp)
-apply (erule set_pur.induct)
-apply (frule_tac [9] AuthReq_msg_in_analz_spies)
-apply (valid_certificate_tac [8]) --{*PReqS*}
-apply (simp_all
-         del: image_insert image_Un imp_disjL
-         add: analz_image_keys_simps disj_simps
-              symKey_compromise pushes sign_def Nonce_compromise
-              analz_Key_image_insert_eq notin_image_iff
-              analz_insert_simps analz_image_priEK)
-  --{*2.5 seconds on a 1.6GHz machine*}
-apply spy_analz
-apply (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj])
-apply (blast dest: Says_imp_knows_Spy [THEN analz.Inj] 
-                   Gets_imp_knows_Spy [THEN analz.Inj])
-apply (blast dest: Gets_imp_knows_Spy [THEN analz.Inj]) --{*PReqS*}
-apply (blast dest: Says_imp_knows_Spy [THEN analz.Inj] 
-                   Gets_imp_knows_Spy [THEN analz.Inj]) --{*PRes*}
-done
-
-text{*This theorem is a bit silly, in that many CardSecrets are 0!
-  But then we don't care.  NOT USED*}
-lemma CardSecret_notin_spies:
-     "evs \<in> set_pur ==> Nonce (CardSecret i) \<notin> parts (knows Spy evs)"
-by (erule set_pur.induct, auto)
-
-
-subsection{*Confidentiality of PAN*}
-
-lemma analz_image_pan_lemma:
-     "(Pan P \<in> analz (Key`nE Un H)) --> (Pan P \<in> analz H)  ==>
-      (Pan P \<in> analz (Key`nE Un H)) =   (Pan P \<in> analz H)"
-by (blast intro: analz_mono [THEN [2] rev_subsetD])
-
-text{*The @{text "(no_asm)"} attribute is essential, since it retains
-  the quantifier and allows the simprule's condition to itself be simplified.*}
-lemma analz_image_pan [rule_format (no_asm)]:
-     "evs \<in> set_pur ==>
-       \<forall>KK. (\<forall>K \<in> KK. K \<notin> range(\<lambda>C. priEK C)) -->
-            (Pan P \<in> analz (Key`KK Un (knows Spy evs))) =
-            (Pan P \<in> analz (knows Spy evs))"
-apply (erule set_pur.induct)
-apply (rule_tac [!] allI impI)+
-apply (rule_tac [!] analz_image_pan_lemma)+
-apply (frule_tac [9] AuthReq_msg_in_analz_spies) --{*AReq*}
-apply (valid_certificate_tac [8]) --{*PReqS*}
-apply (valid_certificate_tac [7]) --{*PReqUns*}
-apply (simp_all
-         del: image_insert image_Un imp_disjL
-         add: analz_image_keys_simps
-              symKey_compromise pushes sign_def
-              analz_Key_image_insert_eq notin_image_iff
-              analz_insert_simps analz_image_priEK)
-  --{*7 seconds on a 1.6GHz machine*}
-apply spy_analz --{*Fake*}
-apply auto
-done
-
-lemma analz_insert_pan:
-     "[| evs \<in> set_pur;  K \<notin> range(\<lambda>C. priEK C) |] ==>
-          (Pan P \<in> analz (insert (Key K) (knows Spy evs))) =
-          (Pan P \<in> analz (knows Spy evs))"
-by (simp del: image_insert image_Un
-         add: analz_image_keys_simps analz_image_pan)
-
-text{*Confidentiality of the PAN, unsigned case.*}
-theorem pan_confidentiality_unsigned:
-     "[| Pan (pan C) \<in> analz(knows Spy evs);  C = Cardholder k;
-         CardSecret k = 0;  evs \<in> set_pur|]
-    ==> \<exists>P M KC1 K X Y.
-     Says C M {|EXHcrypt KC1 (pubEK P) X (Pan (pan C)), Y|}
-          \<in> set evs  &
-     P \<in> bad"
-apply (erule rev_mp)
-apply (erule set_pur.induct)
-apply (frule_tac [9] AuthReq_msg_in_analz_spies) --{*AReq*}
-apply (valid_certificate_tac [8]) --{*PReqS*}
-apply (valid_certificate_tac [7]) --{*PReqUns*}
-apply (simp_all
-         del: image_insert image_Un imp_disjL
-         add: analz_image_keys_simps analz_insert_pan analz_image_pan
-              notin_image_iff
-              analz_insert_simps analz_image_priEK)
-  --{*3 seconds on a 1.6GHz machine*}
-apply spy_analz --{*Fake*}
-apply blast --{*PReqUns: unsigned*}
-apply force --{*PReqS: signed*}
-done
-
-text{*Confidentiality of the PAN, signed case.*}
-theorem pan_confidentiality_signed:
- "[|Pan (pan C) \<in> analz(knows Spy evs);  C = Cardholder k;
-    CardSecret k \<noteq> 0;  evs \<in> set_pur|]
-  ==> \<exists>P M KC2 PIDualSign_1 PIDualSign_2 other OIDualSign.
-      Says C M {|{|PIDualSign_1, 
-                   EXcrypt KC2 (pubEK P) PIDualSign_2 {|Pan (pan C), other|}|}, 
-       OIDualSign|} \<in> set evs  &  P \<in> bad"
-apply (erule rev_mp)
-apply (erule set_pur.induct)
-apply (frule_tac [9] AuthReq_msg_in_analz_spies) --{*AReq*}
-apply (valid_certificate_tac [8]) --{*PReqS*}
-apply (valid_certificate_tac [7]) --{*PReqUns*}
-apply (simp_all
-         del: image_insert image_Un imp_disjL
-         add: analz_image_keys_simps analz_insert_pan analz_image_pan
-              notin_image_iff
-              analz_insert_simps analz_image_priEK)
-  --{*3 seconds on a 1.6GHz machine*}
-apply spy_analz --{*Fake*}
-apply force --{*PReqUns: unsigned*}
-apply blast --{*PReqS: signed*}
-done
-
-text{*General goal: that C, M and PG agree on those details of the transaction
-     that they are allowed to know about.  PG knows about price and account
-     details.  M knows about the order description and price.  C knows both.*}
-
-
-subsection{*Proofs Common to Signed and Unsigned Versions*}
-
-lemma M_Notes_PG:
-     "[|Notes M {|Number LID_M, Agent P, Agent M, Agent C, etc|} \<in> set evs;
-        evs \<in> set_pur|] ==> \<exists>j. P = PG j"
-by (erule rev_mp, erule set_pur.induct, simp_all)
-
-text{*If we trust M, then @{term LID_M} determines his choice of P
-      (Payment Gateway)*}
-lemma goodM_gives_correct_PG:
-     "[| MsgPInitRes = 
-            {|Number LID_M, xid, cc, cm, cert P EKj onlyEnc (priSK RCA)|};
-         Crypt (priSK M) (Hash MsgPInitRes) \<in> parts (knows Spy evs);
-         evs \<in> set_pur; M \<notin> bad |]
-      ==> \<exists>j trans.
-            P = PG j &
-            Notes M {|Number LID_M, Agent P, trans|} \<in> set evs"
-apply clarify
-apply (erule rev_mp)
-apply (erule set_pur.induct)
-apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
-apply simp_all
-apply (blast intro: M_Notes_PG)+
-done
-
-lemma C_gets_correct_PG:
-     "[| Gets A (sign (priSK M) {|Number LID_M, xid, cc, cm,
-                              cert P EKj onlyEnc (priSK RCA)|}) \<in> set evs;
-         evs \<in> set_pur;  M \<notin> bad|]
-      ==> \<exists>j trans.
-            P = PG j &
-            Notes M {|Number LID_M, Agent P, trans|} \<in> set evs &
-            EKj = pubEK P"
-by (rule refl [THEN goodM_gives_correct_PG, THEN exE], auto)
-
-text{*When C receives PInitRes, he learns M's choice of P*}
-lemma C_verifies_PInitRes:
- "[| MsgPInitRes = {|Number LID_M, Number XID, Nonce Chall_C, Nonce Chall_M,
-           cert P EKj onlyEnc (priSK RCA)|};
-     Crypt (priSK M) (Hash MsgPInitRes) \<in> parts (knows Spy evs);
-     evs \<in> set_pur;  M \<notin> bad|]
-  ==> \<exists>j trans.
-         Notes M {|Number LID_M, Agent P, trans|} \<in> set evs &
-         P = PG j &
-         EKj = pubEK P"
-apply clarify
-apply (erule rev_mp)
-apply (erule set_pur.induct)
-apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
-apply simp_all
-apply (blast intro: M_Notes_PG)+
-done
-
-text{*Corollary of previous one*}
-lemma Says_C_PInitRes:
-     "[|Says A C (sign (priSK M)
-                      {|Number LID_M, Number XID,
-                        Nonce Chall_C, Nonce Chall_M,
-                        cert P EKj onlyEnc (priSK RCA)|})
-           \<in> set evs;  M \<notin> bad;  evs \<in> set_pur|]
-      ==> \<exists>j trans.
-           Notes M {|Number LID_M, Agent P, trans|} \<in> set evs &
-           P = PG j &
-           EKj = pubEK (PG j)"
-apply (frule Says_certificate_valid)
-apply (auto simp add: sign_def)
-apply (blast dest: refl [THEN goodM_gives_correct_PG])
-apply (blast dest: refl [THEN C_verifies_PInitRes])
-done
-
-text{*When P receives an AuthReq, he knows that the signed part originated 
-      with M. PIRes also has a signed message from M....*}
-lemma P_verifies_AuthReq:
-     "[| AuthReqData = {|Number LID_M, Number XID, HOIData, HOD|};
-         Crypt (priSK M) (Hash {|AuthReqData, Hash P_I|})
-           \<in> parts (knows Spy evs);
-         evs \<in> set_pur;  M \<notin> bad|]
-      ==> \<exists>j trans KM OIData HPIData.
-            Notes M {|Number LID_M, Agent (PG j), trans|} \<in> set evs &
-            Gets M {|P_I, OIData, HPIData|} \<in> set evs &
-            Says M (PG j) (EncB (priSK M) KM (pubEK (PG j)) AuthReqData P_I)
-              \<in> set evs"
-apply clarify
-apply (erule rev_mp)
-apply (erule set_pur.induct, simp_all)
-apply (frule_tac [4] M_Notes_PG, auto)
-done
-
-text{*When M receives AuthRes, he knows that P signed it, including
-  the identifying tags and the purchase amount, which he can verify.
-  (Although the spec has SIGNED and UNSIGNED forms of AuthRes, they
-   send the same message to M.)  The conclusion is weak: M is existentially
-  quantified! That is because Authorization Response does not refer to M, while
-  the digital envelope weakens the link between @{term MsgAuthRes} and
-  @{term"priSK M"}.  Changing the precondition to refer to 
-  @{term "Crypt K (sign SK M)"} requires assuming @{term K} to be secure, since
-  otherwise the Spy could create that message.*}
-theorem M_verifies_AuthRes:
-  "[| MsgAuthRes = {|{|Number LID_M, Number XID, Number PurchAmt|}, 
-                     Hash authCode|};
-      Crypt (priSK (PG j)) (Hash MsgAuthRes) \<in> parts (knows Spy evs);
-      PG j \<notin> bad;  evs \<in> set_pur|]
-   ==> \<exists>M KM KP HOIData HOD P_I.
-        Gets (PG j)
-           (EncB (priSK M) KM (pubEK (PG j))
-                    {|Number LID_M, Number XID, HOIData, HOD|}
-                    P_I) \<in> set evs &
-        Says (PG j) M
-             (EncB (priSK (PG j)) KP (pubEK M)
-              {|Number LID_M, Number XID, Number PurchAmt|}
-              authCode) \<in> set evs"
-apply clarify
-apply (erule rev_mp)
-apply (erule set_pur.induct)
-apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
-apply simp_all
-apply blast+
-done
-
-
-subsection{*Proofs for Unsigned Purchases*}
-
-text{*What we can derive from the ASSUMPTION that C issued a purchase request.
-   In the unsigned case, we must trust "C": there's no authentication.*}
-lemma C_determines_EKj:
-     "[| Says C M {|EXHcrypt KC1 EKj {|PIHead, Hash OIData|} (Pan (pan C)),
-                    OIData, Hash{|PIHead, Pan (pan C)|} |} \<in> set evs;
-         PIHead = {|Number LID_M, Trans_details|};
-         evs \<in> set_pur;  C = Cardholder k;  M \<notin> bad|]
-  ==> \<exists>trans j.
-               Notes M {|Number LID_M, Agent (PG j), trans |} \<in> set evs &
-               EKj = pubEK (PG j)"
-apply clarify
-apply (erule rev_mp)
-apply (erule set_pur.induct, simp_all)
-apply (valid_certificate_tac [2]) --{*PReqUns*}
-apply auto
-apply (blast dest: Gets_imp_Says Says_C_PInitRes)
-done
-
-
-text{*Unicity of @{term LID_M} between Merchant and Cardholder notes*}
-lemma unique_LID_M:
-     "[|Notes (Merchant i) {|Number LID_M, Agent P, Trans|} \<in> set evs;
-        Notes C {|Number LID_M, Agent M, Agent C, Number OD,
-             Number PA|} \<in> set evs;
-        evs \<in> set_pur|]
-      ==> M = Merchant i & Trans = {|Agent M, Agent C, Number OD, Number PA|}"
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule set_pur.induct, simp_all)
-apply (force dest!: Notes_imp_parts_subset_used)
-done
-
-text{*Unicity of @{term LID_M}, for two Merchant Notes events*}
-lemma unique_LID_M2:
-     "[|Notes M {|Number LID_M, Trans|} \<in> set evs;
-        Notes M {|Number LID_M, Trans'|} \<in> set evs;
-        evs \<in> set_pur|] ==> Trans' = Trans"
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule set_pur.induct, simp_all)
-apply (force dest!: Notes_imp_parts_subset_used)
-done
-
-text{*Lemma needed below: for the case that
-  if PRes is present, then @{term LID_M} has been used.*}
-lemma signed_imp_used:
-     "[| Crypt (priSK M) (Hash X) \<in> parts (knows Spy evs);
-         M \<notin> bad;  evs \<in> set_pur|] ==> parts {X} \<subseteq> used evs"
-apply (erule rev_mp)
-apply (erule set_pur.induct)
-apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
-apply simp_all
-apply safe
-apply blast+
-done
-
-text{*Similar, with nested Hash*}
-lemma signed_Hash_imp_used:
-     "[| Crypt (priSK C) (Hash {|H, Hash X|}) \<in> parts (knows Spy evs);
-         C \<notin> bad;  evs \<in> set_pur|] ==> parts {X} \<subseteq> used evs"
-apply (erule rev_mp)
-apply (erule set_pur.induct)
-apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
-apply simp_all
-apply safe
-apply blast+
-done
-
-text{*Lemma needed below: for the case that
-  if PRes is present, then @{text LID_M} has been used.*}
-lemma PRes_imp_LID_used:
-     "[| Crypt (priSK M) (Hash {|N, X|}) \<in> parts (knows Spy evs);
-         M \<notin> bad;  evs \<in> set_pur|] ==> N \<in> used evs"
-by (drule signed_imp_used, auto)
-
-text{*When C receives PRes, he knows that M and P agreed to the purchase details.
-  He also knows that P is the same PG as before*}
-lemma C_verifies_PRes_lemma:
-     "[| Crypt (priSK M) (Hash MsgPRes) \<in> parts (knows Spy evs);
-         Notes C {|Number LID_M, Trans |} \<in> set evs;
-         Trans = {| Agent M, Agent C, Number OrderDesc, Number PurchAmt |};
-         MsgPRes = {|Number LID_M, Number XID, Nonce Chall_C,
-                Hash (Number PurchAmt)|};
-         evs \<in> set_pur;  M \<notin> bad|]
-  ==> \<exists>j KP.
-        Notes M {|Number LID_M, Agent (PG j), Trans |}
-          \<in> set evs &
-        Gets M (EncB (priSK (PG j)) KP (pubEK M)
-                {|Number LID_M, Number XID, Number PurchAmt|}
-                authCode)
-          \<in> set evs &
-        Says M C (sign (priSK M) MsgPRes) \<in> set evs"
-apply clarify
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule set_pur.induct)
-apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
-apply simp_all
-apply blast
-apply blast
-apply (blast dest: PRes_imp_LID_used)
-apply (frule M_Notes_PG, auto)
-apply (blast dest: unique_LID_M)
-done
-
-text{*When the Cardholder receives Purchase Response from an uncompromised
-Merchant, he knows that M sent it. He also knows that M received a message signed
-by a Payment Gateway chosen by M to authorize the purchase.*}
-theorem C_verifies_PRes:
-     "[| MsgPRes = {|Number LID_M, Number XID, Nonce Chall_C,
-                     Hash (Number PurchAmt)|};
-         Gets C (sign (priSK M) MsgPRes) \<in> set evs;
-         Notes C {|Number LID_M, Agent M, Agent C, Number OrderDesc,
-                   Number PurchAmt|} \<in> set evs;
-         evs \<in> set_pur;  M \<notin> bad|]
-  ==> \<exists>P KP trans.
-        Notes M {|Number LID_M,Agent P, trans|} \<in> set evs &
-        Gets M (EncB (priSK P) KP (pubEK M)
-                {|Number LID_M, Number XID, Number PurchAmt|}
-                authCode)  \<in>  set evs &
-        Says M C (sign (priSK M) MsgPRes) \<in> set evs"
-apply (rule C_verifies_PRes_lemma [THEN exE])
-apply (auto simp add: sign_def)
-done
-
-subsection{*Proofs for Signed Purchases*}
-
-text{*Some Useful Lemmas: the cardholder knows what he is doing*}
-
-lemma Crypt_imp_Says_Cardholder:
-     "[| Crypt K {|{|{|Number LID_M, others|}, Hash OIData|}, Hash PANData|}
-           \<in> parts (knows Spy evs);
-         PANData = {|Pan (pan (Cardholder k)), Nonce (PANSecret k)|};
-         Key K \<notin> analz (knows Spy evs);
-         evs \<in> set_pur|]
-  ==> \<exists>M shash EK HPIData.
-       Says (Cardholder k) M {|{|shash,
-          Crypt K
-            {|{|{|Number LID_M, others|}, Hash OIData|}, Hash PANData|},
-           Crypt EK {|Key K, PANData|}|},
-          OIData, HPIData|} \<in> set evs"
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule set_pur.induct, analz_mono_contra)
-apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
-apply simp_all
-apply auto
-done
-
-lemma Says_PReqS_imp_trans_details_C:
-     "[| MsgPReqS = {|{|shash,
-                 Crypt K
-                  {|{|{|Number LID_M, PIrest|}, Hash OIData|}, hashpd|},
-            cryptek|}, data|};
-         Says (Cardholder k) M MsgPReqS \<in> set evs;
-         evs \<in> set_pur |]
-   ==> \<exists>trans.
-           Notes (Cardholder k) 
-                 {|Number LID_M, Agent M, Agent (Cardholder k), trans|}
-            \<in> set evs"
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule set_pur.induct)
-apply (simp_all (no_asm_simp))
-apply auto
-done
-
-text{*Can't happen: only Merchants create this type of Note*}
-lemma Notes_Cardholder_self_False:
-     "[|Notes (Cardholder k)
-          {|Number n, Agent P, Agent (Cardholder k), Agent C, etc|} \<in> set evs;
-        evs \<in> set_pur|] ==> False"
-by (erule rev_mp, erule set_pur.induct, auto)
-
-text{*When M sees a dual signature, he knows that it originated with C.
-  Using XID he knows it was intended for him.
-  This guarantee isn't useful to P, who never gets OIData.*}
-theorem M_verifies_Signed_PReq:
- "[| MsgDualSign = {|HPIData, Hash OIData|};
-     OIData = {|Number LID_M, etc|};
-     Crypt (priSK C) (Hash MsgDualSign) \<in> parts (knows Spy evs);
-     Notes M {|Number LID_M, Agent P, extras|} \<in> set evs;
-     M = Merchant i;  C = Cardholder k;  C \<notin> bad;  evs \<in> set_pur|]
-  ==> \<exists>PIData PICrypt.
-        HPIData = Hash PIData &
-        Says C M {|{|sign (priSK C) MsgDualSign, PICrypt|}, OIData, Hash PIData|}
-          \<in> set evs"
-apply clarify
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule set_pur.induct)
-apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
-apply simp_all
-apply blast
-apply (metis subsetD insert_subset parts.Fst parts_increasing signed_Hash_imp_used)
-apply (metis unique_LID_M)
-apply (blast dest!: Notes_Cardholder_self_False)
-done
-
-text{*When P sees a dual signature, he knows that it originated with C.
-  and was intended for M. This guarantee isn't useful to M, who never gets
-  PIData. I don't see how to link @{term "PG j"} and @{text LID_M} without
-  assuming @{term "M \<notin> bad"}.*}
-theorem P_verifies_Signed_PReq:
-     "[| MsgDualSign = {|Hash PIData, HOIData|};
-         PIData = {|PIHead, PANData|};
-         PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M,
-                    TransStain|};
-         Crypt (priSK C) (Hash MsgDualSign) \<in> parts (knows Spy evs);
-         evs \<in> set_pur;  C \<notin> bad;  M \<notin> bad|]
-    ==> \<exists>OIData OrderDesc K j trans.
-          HOD = Hash{|Number OrderDesc, Number PurchAmt|} &
-          HOIData = Hash OIData &
-          Notes M {|Number LID_M, Agent (PG j), trans|} \<in> set evs &
-          Says C M {|{|sign (priSK C) MsgDualSign,
-                     EXcrypt K (pubEK (PG j))
-                                {|PIHead, Hash OIData|} PANData|},
-                     OIData, Hash PIData|}
-            \<in> set evs"
-apply clarify
-apply (erule rev_mp)
-apply (erule set_pur.induct, simp_all)
-apply (auto dest!: C_gets_correct_PG)
-done
-
-lemma C_determines_EKj_signed:
-     "[| Says C M {|{|sign (priSK C) text,
-                      EXcrypt K EKj {|PIHead, X|} Y|}, Z|} \<in> set evs;
-         PIHead = {|Number LID_M, Number XID, W|};
-         C = Cardholder k;  evs \<in> set_pur;  M \<notin> bad|]
-  ==> \<exists> trans j.
-         Notes M {|Number LID_M, Agent (PG j), trans|} \<in> set evs &
-         EKj = pubEK (PG j)"
-apply clarify
-apply (erule rev_mp)
-apply (erule set_pur.induct, simp_all, auto)
-apply (blast dest: C_gets_correct_PG)
-done
-
-lemma M_Says_AuthReq:
-     "[| AuthReqData = {|Number LID_M, Number XID, HOIData, HOD|};
-         sign (priSK M) {|AuthReqData, Hash P_I|} \<in> parts (knows Spy evs);
-         evs \<in> set_pur;  M \<notin> bad|]
-   ==> \<exists>j trans KM.
-           Notes M {|Number LID_M, Agent (PG j), trans |} \<in> set evs &
-             Says M (PG j)
-               (EncB (priSK M) KM (pubEK (PG j)) AuthReqData P_I)
-              \<in> set evs"
-apply (rule refl [THEN P_verifies_AuthReq, THEN exE])
-apply (auto simp add: sign_def)
-done
-
-text{*A variant of @{text M_verifies_Signed_PReq} with explicit PI information.
-  Even here we cannot be certain about what C sent to M, since a bad
-  PG could have replaced the two key fields.  (NOT USED)*}
-lemma Signed_PReq_imp_Says_Cardholder:
-     "[| MsgDualSign = {|Hash PIData, Hash OIData|};
-         OIData = {|Number LID_M, Number XID, Nonce Chall_C, HOD, etc|};
-         PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M,
-                    TransStain|};
-         PIData = {|PIHead, PANData|};
-         Crypt (priSK C) (Hash MsgDualSign) \<in> parts (knows Spy evs);
-         M = Merchant i;  C = Cardholder k;  C \<notin> bad;  evs \<in> set_pur|]
-      ==> \<exists>KC EKj.
-            Says C M {|{|sign (priSK C) MsgDualSign,
-                       EXcrypt KC EKj {|PIHead, Hash OIData|} PANData|},
-                       OIData, Hash PIData|}
-              \<in> set evs"
-apply clarify
-apply (erule rev_mp)
-apply (erule rev_mp)
-apply (erule set_pur.induct, simp_all, auto)
-done
-
-text{*When P receives an AuthReq and a dual signature, he knows that C and M
-  agree on the essential details.  PurchAmt however is never sent by M to
-  P; instead C and M both send 
-     @{term "HOD = Hash{|Number OrderDesc, Number PurchAmt|}"}
-  and P compares the two copies of HOD.
-
-  Agreement can't be proved for some things, including the symmetric keys
-  used in the digital envelopes.  On the other hand, M knows the true identity
-  of PG (namely j'), and sends AReq there; he can't, however, check that
-  the EXcrypt involves the correct PG's key.
-*}
-theorem P_sees_CM_agreement:
-     "[| AuthReqData = {|Number LID_M, Number XID, HOIData, HOD|};
-         KC \<in> symKeys;
-         Gets (PG j) (EncB (priSK M) KM (pubEK (PG j)) AuthReqData P_I)
-           \<in> set evs;
-         C = Cardholder k;
-         PI_sign = sign (priSK C) {|Hash PIData, HOIData|};
-         P_I = {|PI_sign,
-                 EXcrypt KC (pubEK (PG j)) {|PIHead, HOIData|} PANData|};
-         PANData = {|Pan (pan C), Nonce (PANSecret k)|};
-         PIData = {|PIHead, PANData|};
-         PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M,
-                    TransStain|};
-         evs \<in> set_pur;  C \<notin> bad;  M \<notin> bad|]
-  ==> \<exists>OIData OrderDesc KM' trans j' KC' KC'' P_I' P_I''.
-           HOD = Hash{|Number OrderDesc, Number PurchAmt|} &
-           HOIData = Hash OIData &
-           Notes M {|Number LID_M, Agent (PG j'), trans|} \<in> set evs &
-           Says C M {|P_I', OIData, Hash PIData|} \<in> set evs &
-           Says M (PG j') (EncB (priSK M) KM' (pubEK (PG j'))
-                           AuthReqData P_I'')  \<in>  set evs &
-           P_I' = {|PI_sign,
-             EXcrypt KC' (pubEK (PG j')) {|PIHead, Hash OIData|} PANData|} &
-           P_I'' = {|PI_sign,
-             EXcrypt KC'' (pubEK (PG j)) {|PIHead, Hash OIData|} PANData|}"
-apply clarify
-apply (rule exE)
-apply (rule P_verifies_Signed_PReq [OF refl refl refl])
-apply (simp (no_asm_use) add: sign_def EncB_def, blast)
-apply (assumption+, clarify, simp)
-apply (drule Gets_imp_knows_Spy [THEN parts.Inj], assumption)
-apply (blast elim: EncB_partsE dest: refl [THEN M_Says_AuthReq] unique_LID_M2)
-done
-
-end
--- a/src/HOL/SET-Protocol/ROOT.ML	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,10 +0,0 @@
-(*  Title:      HOL/SET-Protocol/ROOT.ML
-    ID:         $Id$
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   2003  University of Cambridge
-
-Root file for the SET protocol proofs.
-*)
-
-no_document use_thy "Nat_Int_Bij";
-use_thys ["Cardholder_Registration", "Merchant_Registration", "Purchase"];
--- a/src/HOL/SET-Protocol/document/root.tex	Wed Oct 21 08:14:38 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,27 +0,0 @@
-\documentclass[10pt,a4paper,twoside]{article}
-\usepackage{graphicx}
-\usepackage{latexsym,theorem}
-\usepackage{isabelle,isabellesym}
-\usepackage{pdfsetup}\urlstyle{rm}
-
-\begin{document}
-
-\pagestyle{headings}
-\pagenumbering{arabic}
-
-\title{Verification of The SET Protocol}
-\author{Giampaolo Bella, Fabio Massacci, Lawrence C. Paulson et al.}
-\maketitle
-
-\tableofcontents
-
-\begin{center}
-  \includegraphics[scale=0.5]{session_graph}  
-\end{center}
-
-\newpage
-
-\parindent 0pt\parskip 0.5ex
-
-\input{session}
-\end{document}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SET_Protocol/Cardholder_Registration.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,1056 @@
+(*  Title:      HOL/SET_Protocol/Cardholder_Registration.thy
+    Author:     Giampaolo Bella
+    Author:     Fabio Massacci
+    Author:     Lawrence C Paulson
+    Author:     Piero Tramontano
+*)
+
+header{*The SET Cardholder Registration Protocol*}
+
+theory Cardholder_Registration
+imports Public_SET
+begin
+
+text{*Note: nonces seem to consist of 20 bytes.  That includes both freshness
+challenges (Chall-EE, etc.) and important secrets (CardSecret, PANsecret)
+*}
+
+text{*Simplifications involving @{text analz_image_keys_simps} appear to
+have become much slower. The cause is unclear. However, there is a big blow-up
+and the rewriting is very sensitive to the set of rewrite rules given.*}
+
+subsection{*Predicate Formalizing the Encryption Association between Keys *}
+
+consts
+  KeyCryptKey :: "[key, key, event list] => bool"
+
+primrec
+
+KeyCryptKey_Nil:
+  "KeyCryptKey DK K [] = False"
+
+KeyCryptKey_Cons:
+      --{*Says is the only important case.
+        1st case: CR5, where KC3 encrypts KC2.
+        2nd case: any use of priEK C.
+        Revision 1.12 has a more complicated version with separate treatment of
+          the dependency of KC1, KC2 and KC3 on priEK (CA i.)  Not needed since
+          priEK C is never sent (and so can't be lost except at the start). *}
+  "KeyCryptKey DK K (ev # evs) =
+   (KeyCryptKey DK K evs |
+    (case ev of
+      Says A B Z =>
+       ((\<exists>N X Y. A \<noteq> Spy &
+                 DK \<in> symKeys &
+                 Z = {|Crypt DK {|Agent A, Nonce N, Key K, X|}, Y|}) |
+        (\<exists>C. DK = priEK C))
+    | Gets A' X => False
+    | Notes A' X => False))"
+
+
+subsection{*Predicate formalizing the association between keys and nonces *}
+
+consts
+  KeyCryptNonce :: "[key, key, event list] => bool"
+
+primrec
+
+KeyCryptNonce_Nil:
+  "KeyCryptNonce EK K [] = False"
+
+KeyCryptNonce_Cons:
+  --{*Says is the only important case.
+    1st case: CR3, where KC1 encrypts NC2 (distinct from CR5 due to EXH);
+    2nd case: CR5, where KC3 encrypts NC3;
+    3rd case: CR6, where KC2 encrypts NC3;
+    4th case: CR6, where KC2 encrypts NonceCCA;
+    5th case: any use of @{term "priEK C"} (including CardSecret).
+    NB the only Nonces we need to keep secret are CardSecret and NonceCCA.
+    But we can't prove @{text Nonce_compromise} unless the relation covers ALL
+        nonces that the protocol keeps secret.
+  *}
+  "KeyCryptNonce DK N (ev # evs) =
+   (KeyCryptNonce DK N evs |
+    (case ev of
+      Says A B Z =>
+       A \<noteq> Spy &
+       ((\<exists>X Y. DK \<in> symKeys &
+               Z = (EXHcrypt DK X {|Agent A, Nonce N|} Y)) |
+        (\<exists>X Y. DK \<in> symKeys &
+               Z = {|Crypt DK {|Agent A, Nonce N, X|}, Y|}) |
+        (\<exists>K i X Y.
+          K \<in> symKeys &
+          Z = Crypt K {|sign (priSK (CA i)) {|Agent B, Nonce N, X|}, Y|} &
+          (DK=K | KeyCryptKey DK K evs)) |
+        (\<exists>K C NC3 Y.
+          K \<in> symKeys &
+          Z = Crypt K
+                {|sign (priSK C) {|Agent B, Nonce NC3, Agent C, Nonce N|},
+                  Y|} &
+          (DK=K | KeyCryptKey DK K evs)) |
+        (\<exists>C. DK = priEK C))
+    | Gets A' X => False
+    | Notes A' X => False))"
+
+
+subsection{*Formal protocol definition *}
+
+inductive_set
+  set_cr :: "event list set"
+where
+
+  Nil:    --{*Initial trace is empty*}
+          "[] \<in> set_cr"
+
+| Fake:    --{*The spy MAY say anything he CAN say.*}
+           "[| evsf \<in> set_cr; X \<in> synth (analz (knows Spy evsf)) |]
+            ==> Says Spy B X  # evsf \<in> set_cr"
+
+| Reception: --{*If A sends a message X to B, then B might receive it*}
+             "[| evsr \<in> set_cr; Says A B X \<in> set evsr |]
+              ==> Gets B X  # evsr \<in> set_cr"
+
+| SET_CR1: --{*CardCInitReq: C initiates a run, sending a nonce to CCA*}
+             "[| evs1 \<in> set_cr;  C = Cardholder k;  Nonce NC1 \<notin> used evs1 |]
+              ==> Says C (CA i) {|Agent C, Nonce NC1|} # evs1 \<in> set_cr"
+
+| SET_CR2: --{*CardCInitRes: CA responds sending NC1 and its certificates*}
+             "[| evs2 \<in> set_cr;
+                 Gets (CA i) {|Agent C, Nonce NC1|} \<in> set evs2 |]
+              ==> Says (CA i) C
+                       {|sign (priSK (CA i)) {|Agent C, Nonce NC1|},
+                         cert (CA i) (pubEK (CA i)) onlyEnc (priSK RCA),
+                         cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|}
+                    # evs2 \<in> set_cr"
+
+| SET_CR3:
+   --{*RegFormReq: C sends his PAN and a new nonce to CA.
+   C verifies that
+    - nonce received is the same as that sent;
+    - certificates are signed by RCA;
+    - certificates are an encryption certificate (flag is onlyEnc) and a
+      signature certificate (flag is onlySig);
+    - certificates pertain to the CA that C contacted (this is done by
+      checking the signature).
+   C generates a fresh symmetric key KC1.
+   The point of encrypting @{term "{|Agent C, Nonce NC2, Hash (Pan(pan C))|}"}
+   is not clear. *}
+"[| evs3 \<in> set_cr;  C = Cardholder k;
+    Nonce NC2 \<notin> used evs3;
+    Key KC1 \<notin> used evs3; KC1 \<in> symKeys;
+    Gets C {|sign (invKey SKi) {|Agent X, Nonce NC1|},
+             cert (CA i) EKi onlyEnc (priSK RCA),
+             cert (CA i) SKi onlySig (priSK RCA)|}
+       \<in> set evs3;
+    Says C (CA i) {|Agent C, Nonce NC1|} \<in> set evs3|]
+ ==> Says C (CA i) (EXHcrypt KC1 EKi {|Agent C, Nonce NC2|} (Pan(pan C)))
+       # Notes C {|Key KC1, Agent (CA i)|}
+       # evs3 \<in> set_cr"
+
+| SET_CR4:
+    --{*RegFormRes:
+    CA responds sending NC2 back with a new nonce NCA, after checking that
+     - the digital envelope is correctly encrypted by @{term "pubEK (CA i)"}
+     - the entire message is encrypted with the same key found inside the
+       envelope (here, KC1) *}
+"[| evs4 \<in> set_cr;
+    Nonce NCA \<notin> used evs4;  KC1 \<in> symKeys;
+    Gets (CA i) (EXHcrypt KC1 EKi {|Agent C, Nonce NC2|} (Pan(pan X)))
+       \<in> set evs4 |]
+  ==> Says (CA i) C
+          {|sign (priSK (CA i)) {|Agent C, Nonce NC2, Nonce NCA|},
+            cert (CA i) (pubEK (CA i)) onlyEnc (priSK RCA),
+            cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|}
+       # evs4 \<in> set_cr"
+
+| SET_CR5:
+   --{*CertReq: C sends his PAN, a new nonce, its proposed public signature key
+       and its half of the secret value to CA.
+       We now assume that C has a fixed key pair, and he submits (pubSK C).
+       The protocol does not require this key to be fresh.
+       The encryption below is actually EncX.*}
+"[| evs5 \<in> set_cr;  C = Cardholder k;
+    Nonce NC3 \<notin> used evs5;  Nonce CardSecret \<notin> used evs5; NC3\<noteq>CardSecret;
+    Key KC2 \<notin> used evs5; KC2 \<in> symKeys;
+    Key KC3 \<notin> used evs5; KC3 \<in> symKeys; KC2\<noteq>KC3;
+    Gets C {|sign (invKey SKi) {|Agent C, Nonce NC2, Nonce NCA|},
+             cert (CA i) EKi onlyEnc (priSK RCA),
+             cert (CA i) SKi onlySig (priSK RCA) |}
+        \<in> set evs5;
+    Says C (CA i) (EXHcrypt KC1 EKi {|Agent C, Nonce NC2|} (Pan(pan C)))
+         \<in> set evs5 |]
+==> Says C (CA i)
+         {|Crypt KC3
+             {|Agent C, Nonce NC3, Key KC2, Key (pubSK C),
+               Crypt (priSK C)
+                 (Hash {|Agent C, Nonce NC3, Key KC2,
+                         Key (pubSK C), Pan (pan C), Nonce CardSecret|})|},
+           Crypt EKi {|Key KC3, Pan (pan C), Nonce CardSecret|} |}
+    # Notes C {|Key KC2, Agent (CA i)|}
+    # Notes C {|Key KC3, Agent (CA i)|}
+    # evs5 \<in> set_cr"
+
+
+  --{* CertRes: CA responds sending NC3 back with its half of the secret value,
+   its signature certificate and the new cardholder signature
+   certificate.  CA checks to have never certified the key proposed by C.
+   NOTE: In Merchant Registration, the corresponding rule (4)
+   uses the "sign" primitive. The encryption below is actually @{term EncK}, 
+   which is just @{term "Crypt K (sign SK X)"}.
+*}
+
+| SET_CR6:
+"[| evs6 \<in> set_cr;
+    Nonce NonceCCA \<notin> used evs6;
+    KC2 \<in> symKeys;  KC3 \<in> symKeys;  cardSK \<notin> symKeys;
+    Notes (CA i) (Key cardSK) \<notin> set evs6;
+    Gets (CA i)
+      {|Crypt KC3 {|Agent C, Nonce NC3, Key KC2, Key cardSK,
+                    Crypt (invKey cardSK)
+                      (Hash {|Agent C, Nonce NC3, Key KC2,
+                              Key cardSK, Pan (pan C), Nonce CardSecret|})|},
+        Crypt (pubEK (CA i)) {|Key KC3, Pan (pan C), Nonce CardSecret|} |}
+      \<in> set evs6 |]
+==> Says (CA i) C
+         (Crypt KC2
+          {|sign (priSK (CA i))
+                 {|Agent C, Nonce NC3, Agent(CA i), Nonce NonceCCA|},
+            certC (pan C) cardSK (XOR(CardSecret,NonceCCA)) onlySig (priSK (CA i)),
+            cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|})
+      # Notes (CA i) (Key cardSK)
+      # evs6 \<in> set_cr"
+
+
+declare Says_imp_knows_Spy [THEN parts.Inj, dest]
+declare parts.Body [dest]
+declare analz_into_parts [dest]
+declare Fake_parts_insert_in_Un [dest]
+
+text{*A "possibility property": there are traces that reach the end.
+      An unconstrained proof with many subgoals.*}
+
+lemma Says_to_Gets:
+     "Says A B X # evs \<in> set_cr ==> Gets B X # Says A B X # evs \<in> set_cr"
+by (rule set_cr.Reception, auto)
+
+text{*The many nonces and keys generated, some simultaneously, force us to
+  introduce them explicitly as shown below.*}
+lemma possibility_CR6:
+     "[|NC1 < (NC2::nat);  NC2 < NC3;  NC3 < NCA ;
+        NCA < NonceCCA;  NonceCCA < CardSecret;
+        KC1 < (KC2::key);  KC2 < KC3;
+        KC1 \<in> symKeys;  Key KC1 \<notin> used [];
+        KC2 \<in> symKeys;  Key KC2 \<notin> used [];
+        KC3 \<in> symKeys;  Key KC3 \<notin> used [];
+        C = Cardholder k|]
+   ==> \<exists>evs \<in> set_cr.
+       Says (CA i) C
+            (Crypt KC2
+             {|sign (priSK (CA i))
+                    {|Agent C, Nonce NC3, Agent(CA i), Nonce NonceCCA|},
+               certC (pan C) (pubSK (Cardholder k)) (XOR(CardSecret,NonceCCA))
+                     onlySig (priSK (CA i)),
+               cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|})
+          \<in> set evs"
+apply (intro exI bexI)
+apply (rule_tac [2] 
+       set_cr.Nil 
+        [THEN set_cr.SET_CR1 [of concl: C i NC1], 
+         THEN Says_to_Gets, 
+         THEN set_cr.SET_CR2 [of concl: i C NC1], 
+         THEN Says_to_Gets,  
+         THEN set_cr.SET_CR3 [of concl: C i KC1 _ NC2], 
+         THEN Says_to_Gets,  
+         THEN set_cr.SET_CR4 [of concl: i C NC2 NCA], 
+         THEN Says_to_Gets,  
+         THEN set_cr.SET_CR5 [of concl: C i KC3 NC3 KC2 CardSecret],
+         THEN Says_to_Gets,  
+         THEN set_cr.SET_CR6 [of concl: i C KC2]])
+apply basic_possibility
+apply (simp_all (no_asm_simp) add: symKeys_neq_imp_neq)
+done
+
+text{*General facts about message reception*}
+lemma Gets_imp_Says:
+     "[| Gets B X \<in> set evs; evs \<in> set_cr |] ==> \<exists>A. Says A B X \<in> set evs"
+apply (erule rev_mp)
+apply (erule set_cr.induct, auto)
+done
+
+lemma Gets_imp_knows_Spy:
+     "[| Gets B X \<in> set evs; evs \<in> set_cr |]  ==> X \<in> knows Spy evs"
+by (blast dest!: Gets_imp_Says Says_imp_knows_Spy)
+declare Gets_imp_knows_Spy [THEN parts.Inj, dest]
+
+
+subsection{*Proofs on keys *}
+
+text{*Spy never sees an agent's private keys! (unless it's bad at start)*}
+
+lemma Spy_see_private_Key [simp]:
+     "evs \<in> set_cr
+      ==> (Key(invKey (publicKey b A)) \<in> parts(knows Spy evs)) = (A \<in> bad)"
+by (erule set_cr.induct, auto)
+
+lemma Spy_analz_private_Key [simp]:
+     "evs \<in> set_cr ==>
+     (Key(invKey (publicKey b A)) \<in> analz(knows Spy evs)) = (A \<in> bad)"
+by auto
+
+declare Spy_see_private_Key [THEN [2] rev_iffD1, dest!]
+declare Spy_analz_private_Key [THEN [2] rev_iffD1, dest!]
+
+
+subsection{*Begin Piero's Theorems on Certificates*}
+text{*Trivial in the current model, where certificates by RCA are secure *}
+
+lemma Crypt_valid_pubEK:
+     "[| Crypt (priSK RCA) {|Agent C, Key EKi, onlyEnc|}
+           \<in> parts (knows Spy evs);
+         evs \<in> set_cr |] ==> EKi = pubEK C"
+apply (erule rev_mp)
+apply (erule set_cr.induct, auto)
+done
+
+lemma certificate_valid_pubEK:
+    "[| cert C EKi onlyEnc (priSK RCA) \<in> parts (knows Spy evs);
+        evs \<in> set_cr |]
+     ==> EKi = pubEK C"
+apply (unfold cert_def signCert_def)
+apply (blast dest!: Crypt_valid_pubEK)
+done
+
+lemma Crypt_valid_pubSK:
+     "[| Crypt (priSK RCA) {|Agent C, Key SKi, onlySig|}
+           \<in> parts (knows Spy evs);
+         evs \<in> set_cr |] ==> SKi = pubSK C"
+apply (erule rev_mp)
+apply (erule set_cr.induct, auto)
+done
+
+lemma certificate_valid_pubSK:
+    "[| cert C SKi onlySig (priSK RCA) \<in> parts (knows Spy evs);
+        evs \<in> set_cr |] ==> SKi = pubSK C"
+apply (unfold cert_def signCert_def)
+apply (blast dest!: Crypt_valid_pubSK)
+done
+
+lemma Gets_certificate_valid:
+     "[| Gets A {| X, cert C EKi onlyEnc (priSK RCA),
+                      cert C SKi onlySig (priSK RCA)|} \<in> set evs;
+         evs \<in> set_cr |]
+      ==> EKi = pubEK C & SKi = pubSK C"
+by (blast dest: certificate_valid_pubEK certificate_valid_pubSK)
+
+text{*Nobody can have used non-existent keys!*}
+lemma new_keys_not_used:
+     "[|K \<in> symKeys; Key K \<notin> used evs; evs \<in> set_cr|]
+      ==> K \<notin> keysFor (parts (knows Spy evs))"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule set_cr.induct)
+apply (frule_tac [8] Gets_certificate_valid)
+apply (frule_tac [6] Gets_certificate_valid, simp_all)
+apply (force dest!: usedI keysFor_parts_insert) --{*Fake*}
+apply (blast,auto)  --{*Others*}
+done
+
+
+subsection{*New versions: as above, but generalized to have the KK argument *}
+
+lemma gen_new_keys_not_used:
+     "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_cr |]
+      ==> Key K \<notin> used evs --> K \<in> symKeys -->
+          K \<notin> keysFor (parts (Key`KK Un knows Spy evs))"
+by (auto simp add: new_keys_not_used)
+
+lemma gen_new_keys_not_analzd:
+     "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_cr |]
+      ==> K \<notin> keysFor (analz (Key`KK Un knows Spy evs))"
+by (blast intro: keysFor_mono [THEN [2] rev_subsetD]
+          dest: gen_new_keys_not_used)
+
+lemma analz_Key_image_insert_eq:
+     "[|K \<in> symKeys; Key K \<notin> used evs; evs \<in> set_cr |]
+      ==> analz (Key ` (insert K KK) \<union> knows Spy evs) =
+          insert (Key K) (analz (Key ` KK \<union> knows Spy evs))"
+by (simp add: gen_new_keys_not_analzd)
+
+lemma Crypt_parts_imp_used:
+     "[|Crypt K X \<in> parts (knows Spy evs);
+        K \<in> symKeys; evs \<in> set_cr |] ==> Key K \<in> used evs"
+apply (rule ccontr)
+apply (force dest: new_keys_not_used Crypt_imp_invKey_keysFor)
+done
+
+lemma Crypt_analz_imp_used:
+     "[|Crypt K X \<in> analz (knows Spy evs);
+        K \<in> symKeys; evs \<in> set_cr |] ==> Key K \<in> used evs"
+by (blast intro: Crypt_parts_imp_used)
+
+
+(*<*) 
+subsection{*Messages signed by CA*}
+
+text{*Message @{text SET_CR2}: C can check CA's signature if he has received
+     CA's certificate.*}
+lemma CA_Says_2_lemma:
+     "[| Crypt (priSK (CA i)) (Hash{|Agent C, Nonce NC1|})
+           \<in> parts (knows Spy evs);
+         evs \<in> set_cr; (CA i) \<notin> bad |]
+     ==> \<exists>Y. Says (CA i) C {|sign (priSK (CA i)) {|Agent C, Nonce NC1|}, Y|}
+                 \<in> set evs"
+apply (erule rev_mp)
+apply (erule set_cr.induct, auto)
+done
+
+text{*Ever used?*}
+lemma CA_Says_2:
+     "[| Crypt (invKey SK) (Hash{|Agent C, Nonce NC1|})
+           \<in> parts (knows Spy evs);
+         cert (CA i) SK onlySig (priSK RCA) \<in> parts (knows Spy evs);
+         evs \<in> set_cr; (CA i) \<notin> bad |]
+      ==> \<exists>Y. Says (CA i) C {|sign (priSK (CA i)) {|Agent C, Nonce NC1|}, Y|}
+                  \<in> set evs"
+by (blast dest!: certificate_valid_pubSK intro!: CA_Says_2_lemma)
+
+
+text{*Message @{text SET_CR4}: C can check CA's signature if he has received
+      CA's certificate.*}
+lemma CA_Says_4_lemma:
+     "[| Crypt (priSK (CA i)) (Hash{|Agent C, Nonce NC2, Nonce NCA|})
+           \<in> parts (knows Spy evs);
+         evs \<in> set_cr; (CA i) \<notin> bad |]
+      ==> \<exists>Y. Says (CA i) C {|sign (priSK (CA i))
+                     {|Agent C, Nonce NC2, Nonce NCA|}, Y|} \<in> set evs"
+apply (erule rev_mp)
+apply (erule set_cr.induct, auto)
+done
+
+text{*NEVER USED*}
+lemma CA_Says_4:
+     "[| Crypt (invKey SK) (Hash{|Agent C, Nonce NC2, Nonce NCA|})
+           \<in> parts (knows Spy evs);
+         cert (CA i) SK onlySig (priSK RCA) \<in> parts (knows Spy evs);
+         evs \<in> set_cr; (CA i) \<notin> bad |]
+      ==> \<exists>Y. Says (CA i) C {|sign (priSK (CA i))
+                   {|Agent C, Nonce NC2, Nonce NCA|}, Y|} \<in> set evs"
+by (blast dest!: certificate_valid_pubSK intro!: CA_Says_4_lemma)
+
+
+text{*Message @{text SET_CR6}: C can check CA's signature if he has
+      received CA's certificate.*}
+lemma CA_Says_6_lemma:
+     "[| Crypt (priSK (CA i)) 
+               (Hash{|Agent C, Nonce NC3, Agent (CA i), Nonce NonceCCA|})
+           \<in> parts (knows Spy evs);
+         evs \<in> set_cr; (CA i) \<notin> bad |]
+      ==> \<exists>Y K. Says (CA i) C (Crypt K {|sign (priSK (CA i))
+      {|Agent C, Nonce NC3, Agent (CA i), Nonce NonceCCA|}, Y|}) \<in> set evs"
+apply (erule rev_mp)
+apply (erule set_cr.induct, auto)
+done
+
+text{*NEVER USED*}
+lemma CA_Says_6:
+     "[| Crypt (invKey SK) (Hash{|Agent C, Nonce NC3, Agent (CA i), Nonce NonceCCA|})
+           \<in> parts (knows Spy evs);
+         cert (CA i) SK onlySig (priSK RCA) \<in> parts (knows Spy evs);
+         evs \<in> set_cr; (CA i) \<notin> bad |]
+      ==> \<exists>Y K. Says (CA i) C (Crypt K {|sign (priSK (CA i))
+                    {|Agent C, Nonce NC3, Agent (CA i), Nonce NonceCCA|}, Y|}) \<in> set evs"
+by (blast dest!: certificate_valid_pubSK intro!: CA_Says_6_lemma)
+(*>*)
+
+
+subsection{*Useful lemmas *}
+
+text{*Rewriting rule for private encryption keys.  Analogous rewriting rules
+for other keys aren't needed.*}
+
+lemma parts_image_priEK:
+     "[|Key (priEK C) \<in> parts (Key`KK Un (knows Spy evs));
+        evs \<in> set_cr|] ==> priEK C \<in> KK | C \<in> bad"
+by auto
+
+text{*trivial proof because (priEK C) never appears even in (parts evs)*}
+lemma analz_image_priEK:
+     "evs \<in> set_cr ==>
+          (Key (priEK C) \<in> analz (Key`KK Un (knows Spy evs))) =
+          (priEK C \<in> KK | C \<in> bad)"
+by (blast dest!: parts_image_priEK intro: analz_mono [THEN [2] rev_subsetD])
+
+
+subsection{*Secrecy of Session Keys *}
+
+subsubsection{*Lemmas about the predicate KeyCryptKey *}
+
+text{*A fresh DK cannot be associated with any other
+  (with respect to a given trace). *}
+lemma DK_fresh_not_KeyCryptKey:
+     "[| Key DK \<notin> used evs; evs \<in> set_cr |] ==> ~ KeyCryptKey DK K evs"
+apply (erule rev_mp)
+apply (erule set_cr.induct)
+apply (simp_all (no_asm_simp))
+apply (blast dest: Crypt_analz_imp_used)+
+done
+
+text{*A fresh K cannot be associated with any other.  The assumption that
+  DK isn't a private encryption key may be an artifact of the particular
+  definition of KeyCryptKey.*}
+lemma K_fresh_not_KeyCryptKey:
+     "[|\<forall>C. DK \<noteq> priEK C; Key K \<notin> used evs|] ==> ~ KeyCryptKey DK K evs"
+apply (induct evs)
+apply (auto simp add: parts_insert2 split add: event.split)
+done
+
+
+text{*This holds because if (priEK (CA i)) appears in any traffic then it must
+  be known to the Spy, by @{term Spy_see_private_Key}*}
+lemma cardSK_neq_priEK:
+     "[|Key cardSK \<notin> analz (knows Spy evs);
+        Key cardSK : parts (knows Spy evs);
+        evs \<in> set_cr|] ==> cardSK \<noteq> priEK C"
+by blast
+
+lemma not_KeyCryptKey_cardSK [rule_format (no_asm)]:
+     "[|cardSK \<notin> symKeys;  \<forall>C. cardSK \<noteq> priEK C;  evs \<in> set_cr|] ==>
+      Key cardSK \<notin> analz (knows Spy evs) --> ~ KeyCryptKey cardSK K evs"
+by (erule set_cr.induct, analz_mono_contra, auto)
+
+text{*Lemma for message 5: pubSK C is never used to encrypt Keys.*}
+lemma pubSK_not_KeyCryptKey [simp]: "~ KeyCryptKey (pubSK C) K evs"
+apply (induct_tac "evs")
+apply (auto simp add: parts_insert2 split add: event.split)
+done
+
+text{*Lemma for message 6: either cardSK is compromised (when we don't care)
+  or else cardSK hasn't been used to encrypt K.  Previously we treated
+  message 5 in the same way, but the current model assumes that rule
+  @{text SET_CR5} is executed only by honest agents.*}
+lemma msg6_KeyCryptKey_disj:
+     "[|Gets B {|Crypt KC3 {|Agent C, Nonce N, Key KC2, Key cardSK, X|}, Y|}
+          \<in> set evs;
+        cardSK \<notin> symKeys;  evs \<in> set_cr|]
+      ==> Key cardSK \<in> analz (knows Spy evs) |
+          (\<forall>K. ~ KeyCryptKey cardSK K evs)"
+by (blast dest: not_KeyCryptKey_cardSK intro: cardSK_neq_priEK)
+
+text{*As usual: we express the property as a logical equivalence*}
+lemma Key_analz_image_Key_lemma:
+     "P --> (Key K \<in> analz (Key`KK Un H)) --> (K \<in> KK | Key K \<in> analz H)
+      ==>
+      P --> (Key K \<in> analz (Key`KK Un H)) = (K \<in> KK | Key K \<in> analz H)"
+by (blast intro: analz_mono [THEN [2] rev_subsetD])
+
+method_setup valid_certificate_tac = {*
+  Args.goal_spec >> (fn quant => K (SIMPLE_METHOD'' quant
+    (fn i =>
+      EVERY [ftac @{thm Gets_certificate_valid} i,
+             assume_tac i,
+             etac conjE i, REPEAT (hyp_subst_tac i)])))
+*} ""
+
+text{*The @{text "(no_asm)"} attribute is essential, since it retains
+  the quantifier and allows the simprule's condition to itself be simplified.*}
+lemma symKey_compromise [rule_format (no_asm)]:
+     "evs \<in> set_cr ==>
+      (\<forall>SK KK. SK \<in> symKeys \<longrightarrow> (\<forall>K \<in> KK. ~ KeyCryptKey K SK evs)   -->
+               (Key SK \<in> analz (Key`KK Un (knows Spy evs))) =
+               (SK \<in> KK | Key SK \<in> analz (knows Spy evs)))"
+apply (erule set_cr.induct)
+apply (rule_tac [!] allI) +
+apply (rule_tac [!] impI [THEN Key_analz_image_Key_lemma, THEN impI])+
+apply (valid_certificate_tac [8]) --{*for message 5*}
+apply (valid_certificate_tac [6]) --{*for message 5*}
+apply (erule_tac [9] msg6_KeyCryptKey_disj [THEN disjE])
+apply (simp_all
+         del: image_insert image_Un imp_disjL
+         add: analz_image_keys_simps analz_knows_absorb
+              analz_Key_image_insert_eq notin_image_iff
+              K_fresh_not_KeyCryptKey
+              DK_fresh_not_KeyCryptKey ball_conj_distrib
+              analz_image_priEK disj_simps)
+  --{*9 seconds on a 1.6GHz machine*}
+apply spy_analz
+apply blast  --{*3*}
+apply blast  --{*5*}
+done
+
+text{*The remaining quantifiers seem to be essential.
+  NO NEED to assume the cardholder's OK: bad cardholders don't do anything
+  wrong!!*}
+lemma symKey_secrecy [rule_format]:
+     "[|CA i \<notin> bad;  K \<in> symKeys;  evs \<in> set_cr|]
+      ==> \<forall>X c. Says (Cardholder c) (CA i) X \<in> set evs -->
+                Key K \<in> parts{X} -->
+                Cardholder c \<notin> bad -->
+                Key K \<notin> analz (knows Spy evs)"
+apply (erule set_cr.induct)
+apply (frule_tac [8] Gets_certificate_valid) --{*for message 5*}
+apply (frule_tac [6] Gets_certificate_valid) --{*for message 3*}
+apply (erule_tac [11] msg6_KeyCryptKey_disj [THEN disjE])
+apply (simp_all del: image_insert image_Un imp_disjL
+         add: symKey_compromise fresh_notin_analz_knows_Spy
+              analz_image_keys_simps analz_knows_absorb
+              analz_Key_image_insert_eq notin_image_iff
+              K_fresh_not_KeyCryptKey
+              DK_fresh_not_KeyCryptKey
+              analz_image_priEK)
+  --{*2.5 seconds on a 1.6GHz machine*}
+apply spy_analz  --{*Fake*}
+apply (auto intro: analz_into_parts [THEN usedI] in_parts_Says_imp_used)
+done
+
+
+subsection{*Primary Goals of Cardholder Registration *}
+
+text{*The cardholder's certificate really was created by the CA, provided the
+    CA is uncompromised *}
+
+text{*Lemma concerning the actual signed message digest*}
+lemma cert_valid_lemma:
+     "[|Crypt (priSK (CA i)) {|Hash {|Nonce N, Pan(pan C)|}, Key cardSK, N1|}
+          \<in> parts (knows Spy evs);
+        CA i \<notin> bad; evs \<in> set_cr|]
+  ==> \<exists>KC2 X Y. Says (CA i) C
+                     (Crypt KC2 
+                       {|X, certC (pan C) cardSK N onlySig (priSK (CA i)), Y|})
+                  \<in> set evs"
+apply (erule rev_mp)
+apply (erule set_cr.induct)
+apply (simp_all (no_asm_simp))
+apply auto
+done
+
+text{*Pre-packaged version for cardholder.  We don't try to confirm the values
+  of KC2, X and Y, since they are not important.*}
+lemma certificate_valid_cardSK:
+    "[|Gets C (Crypt KC2 {|X, certC (pan C) cardSK N onlySig (invKey SKi),
+                              cert (CA i) SKi onlySig (priSK RCA)|}) \<in> set evs;
+        CA i \<notin> bad; evs \<in> set_cr|]
+  ==> \<exists>KC2 X Y. Says (CA i) C
+                     (Crypt KC2 
+                       {|X, certC (pan C) cardSK N onlySig (priSK (CA i)), Y|})
+                   \<in> set evs"
+by (force dest!: Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Body]
+                    certificate_valid_pubSK cert_valid_lemma)
+
+
+lemma Hash_imp_parts [rule_format]:
+     "evs \<in> set_cr
+      ==> Hash{|X, Nonce N|} \<in> parts (knows Spy evs) -->
+          Nonce N \<in> parts (knows Spy evs)"
+apply (erule set_cr.induct, force)
+apply (simp_all (no_asm_simp))
+apply (blast intro: parts_mono [THEN [2] rev_subsetD])
+done
+
+lemma Hash_imp_parts2 [rule_format]:
+     "evs \<in> set_cr
+      ==> Hash{|X, Nonce M, Y, Nonce N|} \<in> parts (knows Spy evs) -->
+          Nonce M \<in> parts (knows Spy evs) & Nonce N \<in> parts (knows Spy evs)"
+apply (erule set_cr.induct, force)
+apply (simp_all (no_asm_simp))
+apply (blast intro: parts_mono [THEN [2] rev_subsetD])
+done
+
+
+subsection{*Secrecy of Nonces*}
+
+subsubsection{*Lemmas about the predicate KeyCryptNonce *}
+
+text{*A fresh DK cannot be associated with any other
+  (with respect to a given trace). *}
+lemma DK_fresh_not_KeyCryptNonce:
+     "[| DK \<in> symKeys; Key DK \<notin> used evs; evs \<in> set_cr |]
+      ==> ~ KeyCryptNonce DK K evs"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule set_cr.induct)
+apply (simp_all (no_asm_simp))
+apply blast
+apply blast
+apply (auto simp add: DK_fresh_not_KeyCryptKey)
+done
+
+text{*A fresh N cannot be associated with any other
+      (with respect to a given trace). *}
+lemma N_fresh_not_KeyCryptNonce:
+     "\<forall>C. DK \<noteq> priEK C ==> Nonce N \<notin> used evs --> ~ KeyCryptNonce DK N evs"
+apply (induct_tac "evs")
+apply (case_tac [2] "a")
+apply (auto simp add: parts_insert2)
+done
+
+lemma not_KeyCryptNonce_cardSK [rule_format (no_asm)]:
+     "[|cardSK \<notin> symKeys;  \<forall>C. cardSK \<noteq> priEK C;  evs \<in> set_cr|] ==>
+      Key cardSK \<notin> analz (knows Spy evs) --> ~ KeyCryptNonce cardSK N evs"
+apply (erule set_cr.induct, analz_mono_contra, simp_all)
+apply (blast dest: not_KeyCryptKey_cardSK)  --{*6*}
+done
+
+subsubsection{*Lemmas for message 5 and 6:
+  either cardSK is compromised (when we don't care)
+  or else cardSK hasn't been used to encrypt K. *}
+
+text{*Lemma for message 5: pubSK C is never used to encrypt Nonces.*}
+lemma pubSK_not_KeyCryptNonce [simp]: "~ KeyCryptNonce (pubSK C) N evs"
+apply (induct_tac "evs")
+apply (auto simp add: parts_insert2 split add: event.split)
+done
+
+text{*Lemma for message 6: either cardSK is compromised (when we don't care)
+  or else cardSK hasn't been used to encrypt K.*}
+lemma msg6_KeyCryptNonce_disj:
+     "[|Gets B {|Crypt KC3 {|Agent C, Nonce N, Key KC2, Key cardSK, X|}, Y|}
+          \<in> set evs;
+        cardSK \<notin> symKeys;  evs \<in> set_cr|]
+      ==> Key cardSK \<in> analz (knows Spy evs) |
+          ((\<forall>K. ~ KeyCryptKey cardSK K evs) &
+           (\<forall>N. ~ KeyCryptNonce cardSK N evs))"
+by (blast dest: not_KeyCryptKey_cardSK not_KeyCryptNonce_cardSK
+          intro: cardSK_neq_priEK)
+
+
+text{*As usual: we express the property as a logical equivalence*}
+lemma Nonce_analz_image_Key_lemma:
+     "P --> (Nonce N \<in> analz (Key`KK Un H)) --> (Nonce N \<in> analz H)
+      ==> P --> (Nonce N \<in> analz (Key`KK Un H)) = (Nonce N \<in> analz H)"
+by (blast intro: analz_mono [THEN [2] rev_subsetD])
+
+
+text{*The @{text "(no_asm)"} attribute is essential, since it retains
+  the quantifier and allows the simprule's condition to itself be simplified.*}
+lemma Nonce_compromise [rule_format (no_asm)]:
+     "evs \<in> set_cr ==>
+      (\<forall>N KK. (\<forall>K \<in> KK. ~ KeyCryptNonce K N evs)   -->
+               (Nonce N \<in> analz (Key`KK Un (knows Spy evs))) =
+               (Nonce N \<in> analz (knows Spy evs)))"
+apply (erule set_cr.induct)
+apply (rule_tac [!] allI)+
+apply (rule_tac [!] impI [THEN Nonce_analz_image_Key_lemma])+
+apply (frule_tac [8] Gets_certificate_valid) --{*for message 5*}
+apply (frule_tac [6] Gets_certificate_valid) --{*for message 3*}
+apply (frule_tac [11] msg6_KeyCryptNonce_disj)
+apply (erule_tac [13] disjE)
+apply (simp_all del: image_insert image_Un
+         add: symKey_compromise
+              analz_image_keys_simps analz_knows_absorb
+              analz_Key_image_insert_eq notin_image_iff
+              N_fresh_not_KeyCryptNonce
+              DK_fresh_not_KeyCryptNonce K_fresh_not_KeyCryptKey
+              ball_conj_distrib analz_image_priEK)
+  --{*14 seconds on a 1.6GHz machine*}
+apply spy_analz  --{*Fake*}
+apply blast  --{*3*}
+apply blast  --{*5*}
+txt{*Message 6*}
+apply (metis symKey_compromise)
+  --{*cardSK compromised*}
+txt{*Simplify again--necessary because the previous simplification introduces
+  some logical connectives*} 
+apply (force simp del: image_insert image_Un imp_disjL
+          simp add: analz_image_keys_simps symKey_compromise)
+done
+
+
+subsection{*Secrecy of CardSecret: the Cardholder's secret*}
+
+lemma NC2_not_CardSecret:
+     "[|Crypt EKj {|Key K, Pan p, Hash {|Agent D, Nonce N|}|}
+          \<in> parts (knows Spy evs);
+        Key K \<notin> analz (knows Spy evs);
+        Nonce N \<notin> analz (knows Spy evs);
+       evs \<in> set_cr|]
+      ==> Crypt EKi {|Key K', Pan p', Nonce N|} \<notin> parts (knows Spy evs)"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule set_cr.induct, analz_mono_contra, simp_all)
+apply (blast dest: Hash_imp_parts)+
+done
+
+lemma KC2_secure_lemma [rule_format]:
+     "[|U = Crypt KC3 {|Agent C, Nonce N, Key KC2, X|};
+        U \<in> parts (knows Spy evs);
+        evs \<in> set_cr|]
+  ==> Nonce N \<notin> analz (knows Spy evs) -->
+      (\<exists>k i W. Says (Cardholder k) (CA i) {|U,W|} \<in> set evs & 
+               Cardholder k \<notin> bad & CA i \<notin> bad)"
+apply (erule_tac P = "U \<in> ?H" in rev_mp)
+apply (erule set_cr.induct)
+apply (valid_certificate_tac [8])  --{*for message 5*}
+apply (simp_all del: image_insert image_Un imp_disjL
+         add: analz_image_keys_simps analz_knows_absorb
+              analz_knows_absorb2 notin_image_iff)
+  --{*4 seconds on a 1.6GHz machine*}
+apply (simp_all (no_asm_simp)) --{*leaves 4 subgoals*}
+apply (blast intro!: analz_insertI)+
+done
+
+lemma KC2_secrecy:
+     "[|Gets B {|Crypt K {|Agent C, Nonce N, Key KC2, X|}, Y|} \<in> set evs;
+        Nonce N \<notin> analz (knows Spy evs);  KC2 \<in> symKeys;
+        evs \<in> set_cr|]
+       ==> Key KC2 \<notin> analz (knows Spy evs)"
+by (force dest!: refl [THEN KC2_secure_lemma] symKey_secrecy)
+
+
+text{*Inductive version*}
+lemma CardSecret_secrecy_lemma [rule_format]:
+     "[|CA i \<notin> bad;  evs \<in> set_cr|]
+      ==> Key K \<notin> analz (knows Spy evs) -->
+          Crypt (pubEK (CA i)) {|Key K, Pan p, Nonce CardSecret|}
+             \<in> parts (knows Spy evs) -->
+          Nonce CardSecret \<notin> analz (knows Spy evs)"
+apply (erule set_cr.induct, analz_mono_contra)
+apply (valid_certificate_tac [8]) --{*for message 5*}
+apply (valid_certificate_tac [6]) --{*for message 5*}
+apply (frule_tac [9] msg6_KeyCryptNonce_disj [THEN disjE])
+apply (simp_all
+         del: image_insert image_Un imp_disjL
+         add: analz_image_keys_simps analz_knows_absorb
+              analz_Key_image_insert_eq notin_image_iff
+              EXHcrypt_def Crypt_notin_image_Key
+              N_fresh_not_KeyCryptNonce DK_fresh_not_KeyCryptNonce
+              ball_conj_distrib Nonce_compromise symKey_compromise
+              analz_image_priEK)
+  --{*2.5 seconds on a 1.6GHz machine*}
+apply spy_analz  --{*Fake*}
+apply (simp_all (no_asm_simp))
+apply blast  --{*1*}
+apply (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj])  --{*2*}
+apply blast  --{*3*}
+apply (blast dest: NC2_not_CardSecret Gets_imp_knows_Spy [THEN analz.Inj] analz_symKeys_Decrypt)  --{*4*}
+apply blast  --{*5*}
+apply (blast dest: KC2_secrecy)+  --{*Message 6: two cases*}
+done
+
+
+text{*Packaged version for cardholder*}
+lemma CardSecret_secrecy:
+     "[|Cardholder k \<notin> bad;  CA i \<notin> bad;
+        Says (Cardholder k) (CA i)
+           {|X, Crypt EKi {|Key KC3, Pan p, Nonce CardSecret|}|} \<in> set evs;
+        Gets A {|Z, cert (CA i) EKi onlyEnc (priSK RCA),
+                    cert (CA i) SKi onlySig (priSK RCA)|} \<in> set evs;
+        KC3 \<in> symKeys;  evs \<in> set_cr|]
+      ==> Nonce CardSecret \<notin> analz (knows Spy evs)"
+apply (frule Gets_certificate_valid, assumption)
+apply (subgoal_tac "Key KC3 \<notin> analz (knows Spy evs) ")
+apply (blast dest: CardSecret_secrecy_lemma)
+apply (rule symKey_secrecy)
+apply (auto simp add: parts_insert2)
+done
+
+
+subsection{*Secrecy of NonceCCA [the CA's secret] *}
+
+lemma NC2_not_NonceCCA:
+     "[|Hash {|Agent C', Nonce N', Agent C, Nonce N|}
+          \<in> parts (knows Spy evs);
+        Nonce N \<notin> analz (knows Spy evs);
+       evs \<in> set_cr|]
+      ==> Crypt KC1 {|{|Agent B, Nonce N|}, Hash p|} \<notin> parts (knows Spy evs)"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule set_cr.induct, analz_mono_contra, simp_all)
+apply (blast dest: Hash_imp_parts2)+
+done
+
+
+text{*Inductive version*}
+lemma NonceCCA_secrecy_lemma [rule_format]:
+     "[|CA i \<notin> bad;  evs \<in> set_cr|]
+      ==> Key K \<notin> analz (knows Spy evs) -->
+          Crypt K
+            {|sign (priSK (CA i))
+                   {|Agent C, Nonce N, Agent(CA i), Nonce NonceCCA|},
+              X, Y|}
+             \<in> parts (knows Spy evs) -->
+          Nonce NonceCCA \<notin> analz (knows Spy evs)"
+apply (erule set_cr.induct, analz_mono_contra)
+apply (valid_certificate_tac [8]) --{*for message 5*}
+apply (valid_certificate_tac [6]) --{*for message 5*}
+apply (frule_tac [9] msg6_KeyCryptNonce_disj [THEN disjE])
+apply (simp_all
+         del: image_insert image_Un imp_disjL
+         add: analz_image_keys_simps analz_knows_absorb sign_def
+              analz_Key_image_insert_eq notin_image_iff
+              EXHcrypt_def Crypt_notin_image_Key
+              N_fresh_not_KeyCryptNonce DK_fresh_not_KeyCryptNonce
+              ball_conj_distrib Nonce_compromise symKey_compromise
+              analz_image_priEK)
+  --{*3 seconds on a 1.6GHz machine*}
+apply spy_analz  --{*Fake*}
+apply blast  --{*1*}
+apply (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj])  --{*2*}
+apply blast  --{*3*}
+apply (blast dest: NC2_not_NonceCCA)  --{*4*}
+apply blast  --{*5*}
+apply (blast dest: KC2_secrecy)+  --{*Message 6: two cases*}
+done
+
+
+text{*Packaged version for cardholder*}
+lemma NonceCCA_secrecy:
+     "[|Cardholder k \<notin> bad;  CA i \<notin> bad;
+        Gets (Cardholder k)
+           (Crypt KC2
+            {|sign (priSK (CA i)) {|Agent C, Nonce N, Agent(CA i), Nonce NonceCCA|},
+              X, Y|}) \<in> set evs;
+        Says (Cardholder k) (CA i)
+           {|Crypt KC3 {|Agent C, Nonce NC3, Key KC2, X'|}, Y'|} \<in> set evs;
+        Gets A {|Z, cert (CA i) EKi onlyEnc (priSK RCA),
+                    cert (CA i) SKi onlySig (priSK RCA)|} \<in> set evs;
+        KC2 \<in> symKeys;  evs \<in> set_cr|]
+      ==> Nonce NonceCCA \<notin> analz (knows Spy evs)"
+apply (frule Gets_certificate_valid, assumption)
+apply (subgoal_tac "Key KC2 \<notin> analz (knows Spy evs) ")
+apply (blast dest: NonceCCA_secrecy_lemma)
+apply (rule symKey_secrecy)
+apply (auto simp add: parts_insert2)
+done
+
+text{*We don't bother to prove guarantees for the CA.  He doesn't care about
+  the PANSecret: it isn't his credit card!*}
+
+
+subsection{*Rewriting Rule for PANs*}
+
+text{*Lemma for message 6: either cardSK isn't a CA's private encryption key,
+  or if it is then (because it appears in traffic) that CA is bad,
+  and so the Spy knows that key already.  Either way, we can simplify
+  the expression @{term "analz (insert (Key cardSK) X)"}.*}
+lemma msg6_cardSK_disj:
+     "[|Gets A {|Crypt K {|c, n, k', Key cardSK, X|}, Y|}
+          \<in> set evs;  evs \<in> set_cr |]
+      ==> cardSK \<notin> range(invKey o pubEK o CA) | Key cardSK \<in> knows Spy evs"
+by auto
+
+lemma analz_image_pan_lemma:
+     "(Pan P \<in> analz (Key`nE Un H)) --> (Pan P \<in> analz H)  ==>
+      (Pan P \<in> analz (Key`nE Un H)) =   (Pan P \<in> analz H)"
+by (blast intro: analz_mono [THEN [2] rev_subsetD])
+
+lemma analz_image_pan [rule_format]:
+     "evs \<in> set_cr ==>
+       \<forall>KK. KK <= - invKey ` pubEK ` range CA -->
+            (Pan P \<in> analz (Key`KK Un (knows Spy evs))) =
+            (Pan P \<in> analz (knows Spy evs))"
+apply (erule set_cr.induct)
+apply (rule_tac [!] allI impI)+
+apply (rule_tac [!] analz_image_pan_lemma)
+apply (valid_certificate_tac [8]) --{*for message 5*}
+apply (valid_certificate_tac [6]) --{*for message 5*}
+apply (erule_tac [9] msg6_cardSK_disj [THEN disjE])
+apply (simp_all
+         del: image_insert image_Un
+         add: analz_image_keys_simps disjoint_image_iff
+              notin_image_iff analz_image_priEK)
+  --{*6 seconds on a 1.6GHz machine*}
+apply spy_analz
+apply (simp add: insert_absorb)  --{*6*}
+done
+
+lemma analz_insert_pan:
+     "[| evs \<in> set_cr;  K \<notin> invKey ` pubEK ` range CA |] ==>
+          (Pan P \<in> analz (insert (Key K) (knows Spy evs))) =
+          (Pan P \<in> analz (knows Spy evs))"
+by (simp del: image_insert image_Un
+         add: analz_image_keys_simps analz_image_pan)
+
+
+text{*Confidentiality of the PAN\@.  Maybe we could combine the statements of
+  this theorem with @{term analz_image_pan}, requiring a single induction but
+  a much more difficult proof.*}
+lemma pan_confidentiality:
+     "[| Pan (pan C) \<in> analz(knows Spy evs); C \<noteq>Spy; evs :set_cr|]
+    ==> \<exists>i X K HN.
+        Says C (CA i) {|X, Crypt (pubEK (CA i)) {|Key K, Pan (pan C), HN|} |}
+           \<in> set evs
+      & (CA i) \<in> bad"
+apply (erule rev_mp)
+apply (erule set_cr.induct)
+apply (valid_certificate_tac [8]) --{*for message 5*}
+apply (valid_certificate_tac [6]) --{*for message 5*}
+apply (erule_tac [9] msg6_cardSK_disj [THEN disjE])
+apply (simp_all
+         del: image_insert image_Un
+         add: analz_image_keys_simps analz_insert_pan analz_image_pan
+              notin_image_iff analz_image_priEK)
+  --{*3.5 seconds on a 1.6GHz machine*}
+apply spy_analz  --{*fake*}
+apply blast  --{*3*}
+apply blast  --{*5*}
+apply (simp (no_asm_simp) add: insert_absorb)  --{*6*}
+done
+
+
+subsection{*Unicity*}
+
+lemma CR6_Says_imp_Notes:
+     "[|Says (CA i) C (Crypt KC2
+          {|sign (priSK (CA i)) {|Agent C, Nonce NC3, Agent (CA i), Nonce Y|},
+            certC (pan C) cardSK X onlySig (priSK (CA i)),
+            cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|})  \<in> set evs;
+        evs \<in> set_cr |]
+      ==> Notes (CA i) (Key cardSK) \<in> set evs"
+apply (erule rev_mp)
+apply (erule set_cr.induct)
+apply (simp_all (no_asm_simp))
+done
+
+text{*Unicity of cardSK: it uniquely identifies the other components.  
+      This holds because a CA accepts a cardSK at most once.*}
+lemma cardholder_key_unicity:
+     "[|Says (CA i) C (Crypt KC2
+          {|sign (priSK (CA i)) {|Agent C, Nonce NC3, Agent (CA i), Nonce Y|},
+            certC (pan C) cardSK X onlySig (priSK (CA i)),
+            cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|})
+          \<in> set evs;
+        Says (CA i) C' (Crypt KC2'
+          {|sign (priSK (CA i)) {|Agent C', Nonce NC3', Agent (CA i), Nonce Y'|},
+            certC (pan C') cardSK X' onlySig (priSK (CA i)),
+            cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|})
+          \<in> set evs;
+        evs \<in> set_cr |] ==> C=C' & NC3=NC3' & X=X' & KC2=KC2' & Y=Y'"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule set_cr.induct)
+apply (simp_all (no_asm_simp))
+apply (blast dest!: CR6_Says_imp_Notes)
+done
+
+
+(*<*)
+text{*UNUSED unicity result*}
+lemma unique_KC1:
+     "[|Says C B {|Crypt KC1 X, Crypt EK {|Key KC1, Y|}|}
+          \<in> set evs;
+        Says C B' {|Crypt KC1 X', Crypt EK' {|Key KC1, Y'|}|}
+          \<in> set evs;
+        C \<notin> bad;  evs \<in> set_cr|] ==> B'=B & Y'=Y"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule set_cr.induct, auto)
+done
+
+text{*UNUSED unicity result*}
+lemma unique_KC2:
+     "[|Says C B {|Crypt K {|Agent C, nn, Key KC2, X|}, Y|} \<in> set evs;
+        Says C B' {|Crypt K' {|Agent C, nn', Key KC2, X'|}, Y'|} \<in> set evs;
+        C \<notin> bad;  evs \<in> set_cr|] ==> B'=B & X'=X"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule set_cr.induct, auto)
+done
+(*>*)
+
+
+text{*Cannot show cardSK to be secret because it isn't assumed to be fresh
+  it could be a previously compromised cardSK [e.g. involving a bad CA]*}
+
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SET_Protocol/Event_SET.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,198 @@
+(*  Title:      HOL/SET_Protocol/Event_SET.thy
+    Author:     Giampaolo Bella
+    Author:     Fabio Massacci
+    Author:     Lawrence C Paulson
+*)
+
+header{*Theory of Events for SET*}
+
+theory Event_SET
+imports Message_SET
+begin
+
+text{*The Root Certification Authority*}
+syntax        RCA :: agent
+translations "RCA" == "CA 0"
+
+
+text{*Message events*}
+datatype
+  event = Says  agent agent msg
+        | Gets  agent       msg
+        | Notes agent       msg
+
+
+text{*compromised agents: keys known, Notes visible*}
+consts bad :: "agent set"
+
+text{*Spy has access to his own key for spoof messages, but RCA is secure*}
+specification (bad)
+  Spy_in_bad     [iff]: "Spy \<in> bad"
+  RCA_not_bad [iff]: "RCA \<notin> bad"
+    by (rule exI [of _ "{Spy}"], simp)
+
+
+subsection{*Agents' Knowledge*}
+
+consts  (*Initial states of agents -- parameter of the construction*)
+  initState :: "agent => msg set"
+  knows  :: "[agent, event list] => msg set"
+
+(* Message reception does not extend spy's knowledge because of
+   reception invariant enforced by Reception rule in protocol definition*)
+primrec
+
+knows_Nil:
+  "knows A []       = initState A"
+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))"
+
+
+subsection{*Used Messages*}
+
+consts
+  (*Set of items that might be visible to somebody:
+    complement of the set of fresh items*)
+  used :: "event list => msg set"
+
+(* As above, message reception does extend used items *)
+primrec
+  used_Nil:  "used []         = (UN B. parts (initState B))"
+  used_Cons: "used (ev # evs) =
+                 (case ev of
+                    Says A B X => parts {X} Un (used evs)
+                  | Gets A X   => used evs
+                  | Notes A X  => parts {X} Un (used evs))"
+
+
+
+(* Inserted by default but later removed.  This declaration lets the file
+be re-loaded. Addsimps [knows_Cons, used_Nil, *)
+
+(** Simplifying   parts (insert X (knows Spy evs))
+      = parts {X} Un parts (knows Spy evs) -- since general case loops*)
+
+lemmas parts_insert_knows_A = parts_insert [of _ "knows A evs", standard]
+
+lemma knows_Spy_Says [simp]:
+     "knows Spy (Says A B X # evs) = insert X (knows Spy evs)"
+by auto
+
+text{*Letting the Spy see "bad" agents' notes avoids redundant case-splits
+      on whether @{term "A=Spy"} and whether @{term "A\<in>bad"}*}
+lemma knows_Spy_Notes [simp]:
+     "knows Spy (Notes A X # evs) =
+          (if A:bad then insert X (knows Spy evs) else knows Spy evs)"
+apply auto
+done
+
+lemma knows_Spy_Gets [simp]: "knows Spy (Gets A X # evs) = knows Spy evs"
+by auto
+
+lemma initState_subset_knows: "initState A <= knows A evs"
+apply (induct_tac "evs")
+apply (auto split: event.split) 
+done
+
+lemma knows_Spy_subset_knows_Spy_Says:
+     "knows Spy evs <= knows Spy (Says A B X # evs)"
+by auto
+
+lemma knows_Spy_subset_knows_Spy_Notes:
+     "knows Spy evs <= knows Spy (Notes A X # evs)"
+by auto
+
+lemma knows_Spy_subset_knows_Spy_Gets:
+     "knows Spy evs <= knows Spy (Gets A X # evs)"
+by auto
+
+(*Spy sees what is sent on the traffic*)
+lemma Says_imp_knows_Spy [rule_format]:
+     "Says A B X \<in> set evs --> X \<in> knows Spy evs"
+apply (induct_tac "evs")
+apply (auto split: event.split) 
+done
+
+(*Use with addSEs to derive contradictions from old Says events containing
+  items known to be fresh*)
+lemmas knows_Spy_partsEs =
+     Says_imp_knows_Spy [THEN parts.Inj, THEN revcut_rl, standard] 
+     parts.Body [THEN revcut_rl, standard]
+
+
+subsection{*The Function @{term used}*}
+
+lemma parts_knows_Spy_subset_used: "parts (knows Spy evs) <= used evs"
+apply (induct_tac "evs")
+apply (auto simp add: parts_insert_knows_A split: event.split) 
+done
+
+lemmas usedI = parts_knows_Spy_subset_used [THEN subsetD, intro]
+
+lemma initState_subset_used: "parts (initState B) <= used evs"
+apply (induct_tac "evs")
+apply (auto split: event.split) 
+done
+
+lemmas initState_into_used = initState_subset_used [THEN subsetD]
+
+lemma used_Says [simp]: "used (Says A B X # evs) = parts{X} Un used evs"
+by auto
+
+lemma used_Notes [simp]: "used (Notes A X # evs) = parts{X} Un used evs"
+by auto
+
+lemma used_Gets [simp]: "used (Gets A X # evs) = used evs"
+by auto
+
+
+lemma Notes_imp_parts_subset_used [rule_format]:
+     "Notes A X \<in> set evs --> parts {X} <= used evs"
+apply (induct_tac "evs")
+apply (induct_tac [2] "a", auto)
+done
+
+text{*NOTE REMOVAL--laws above are cleaner, as they don't involve "case"*}
+declare knows_Cons [simp del]
+        used_Nil [simp del] used_Cons [simp del]
+
+
+text{*For proving theorems of the form @{term "X \<notin> analz (knows Spy evs) --> P"}
+  New events added by induction to "evs" are discarded.  Provided 
+  this information isn't needed, the proof will be much shorter, since
+  it will omit complicated reasoning about @{term analz}.*}
+
+lemmas analz_mono_contra =
+       knows_Spy_subset_knows_Spy_Says [THEN analz_mono, THEN contra_subsetD]
+       knows_Spy_subset_knows_Spy_Notes [THEN analz_mono, THEN contra_subsetD]
+       knows_Spy_subset_knows_Spy_Gets [THEN analz_mono, THEN contra_subsetD]
+
+lemmas analz_impI = impI [where P = "Y \<notin> analz (knows Spy evs)", standard]
+
+ML
+{*
+val analz_mono_contra_tac = 
+  rtac @{thm analz_impI} THEN' 
+  REPEAT1 o (dresolve_tac @{thms analz_mono_contra})
+  THEN' mp_tac
+*}
+
+method_setup analz_mono_contra = {*
+    Scan.succeed (K (SIMPLE_METHOD (REPEAT_FIRST analz_mono_contra_tac))) *}
+    "for proving theorems of the form X \<notin> analz (knows Spy evs) --> P"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SET_Protocol/Merchant_Registration.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,429 @@
+(*  Title:      HOL/SET_Protocol/Merchant_Registration.thy
+    Author:     Giampaolo Bella
+    Author:     Fabio Massacci
+    Author:     Lawrence C Paulson
+*)
+
+header{*The SET Merchant Registration Protocol*}
+
+theory Merchant_Registration
+imports Public_SET
+begin
+
+text{*Copmpared with Cardholder Reigstration, @{text KeyCryptKey} is not
+  needed: no session key encrypts another.  Instead we
+  prove the "key compromise" theorems for sets KK that contain no private
+  encryption keys (@{term "priEK C"}). *}
+
+
+inductive_set
+  set_mr :: "event list set"
+where
+
+  Nil:    --{*Initial trace is empty*}
+           "[] \<in> set_mr"
+
+
+| Fake:    --{*The spy MAY say anything he CAN say.*}
+           "[| evsf \<in> set_mr; X \<in> synth (analz (knows Spy evsf)) |]
+            ==> Says Spy B X  # evsf \<in> set_mr"
+        
+
+| Reception: --{*If A sends a message X to B, then B might receive it*}
+             "[| evsr \<in> set_mr; Says A B X \<in> set evsr |]
+              ==> Gets B X  # evsr \<in> set_mr"
+
+
+| SET_MR1: --{*RegFormReq: M requires a registration form to a CA*}
+           "[| evs1 \<in> set_mr; M = Merchant k; Nonce NM1 \<notin> used evs1 |]
+            ==> Says M (CA i) {|Agent M, Nonce NM1|} # evs1 \<in> set_mr"
+
+
+| SET_MR2: --{*RegFormRes: CA replies with the registration form and the 
+               certificates for her keys*}
+  "[| evs2 \<in> set_mr; Nonce NCA \<notin> used evs2;
+      Gets (CA i) {|Agent M, Nonce NM1|} \<in> set evs2 |]
+   ==> Says (CA i) M {|sign (priSK (CA i)) {|Agent M, Nonce NM1, Nonce NCA|},
+                       cert (CA i) (pubEK (CA i)) onlyEnc (priSK RCA),
+                       cert (CA i) (pubSK (CA i)) onlySig (priSK RCA) |}
+         # evs2 \<in> set_mr"
+
+| SET_MR3:
+         --{*CertReq: M submits the key pair to be certified.  The Notes
+             event allows KM1 to be lost if M is compromised. Piero remarks
+             that the agent mentioned inside the signature is not verified to
+             correspond to M.  As in CR, each Merchant has fixed key pairs.  M
+             is only optionally required to send NCA back, so M doesn't do so
+             in the model*}
+  "[| evs3 \<in> set_mr; M = Merchant k; Nonce NM2 \<notin> used evs3;
+      Key KM1 \<notin> used evs3;  KM1 \<in> symKeys;
+      Gets M {|sign (invKey SKi) {|Agent X, Nonce NM1, Nonce NCA|},
+               cert (CA i) EKi onlyEnc (priSK RCA),
+               cert (CA i) SKi onlySig (priSK RCA) |}
+        \<in> set evs3;
+      Says M (CA i) {|Agent M, Nonce NM1|} \<in> set evs3 |]
+   ==> Says M (CA i)
+            {|Crypt KM1 (sign (priSK M) {|Agent M, Nonce NM2,
+                                          Key (pubSK M), Key (pubEK M)|}),
+              Crypt EKi (Key KM1)|}
+         # Notes M {|Key KM1, Agent (CA i)|}
+         # evs3 \<in> set_mr"
+
+| SET_MR4:
+         --{*CertRes: CA issues the certificates for merSK and merEK,
+             while checking never to have certified the m even
+             separately. NOTE: In Cardholder Registration the
+             corresponding rule (6) doesn't use the "sign" primitive. "The
+             CertRes shall be signed but not encrypted if the EE is a Merchant
+             or Payment Gateway."-- Programmer's Guide, page 191.*}
+    "[| evs4 \<in> set_mr; M = Merchant k;
+        merSK \<notin> symKeys;  merEK \<notin> symKeys;
+        Notes (CA i) (Key merSK) \<notin> set evs4;
+        Notes (CA i) (Key merEK) \<notin> set evs4;
+        Gets (CA i) {|Crypt KM1 (sign (invKey merSK)
+                                 {|Agent M, Nonce NM2, Key merSK, Key merEK|}),
+                      Crypt (pubEK (CA i)) (Key KM1) |}
+          \<in> set evs4 |]
+    ==> Says (CA i) M {|sign (priSK(CA i)) {|Agent M, Nonce NM2, Agent(CA i)|},
+                        cert  M      merSK    onlySig (priSK (CA i)),
+                        cert  M      merEK    onlyEnc (priSK (CA i)),
+                        cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|}
+          # Notes (CA i) (Key merSK)
+          # Notes (CA i) (Key merEK)
+          # evs4 \<in> set_mr"
+
+
+text{*Note possibility proofs are missing.*}
+
+declare Says_imp_knows_Spy [THEN parts.Inj, dest]
+declare parts.Body [dest]
+declare analz_into_parts [dest]
+declare Fake_parts_insert_in_Un [dest]
+
+text{*General facts about message reception*}
+lemma Gets_imp_Says:
+     "[| Gets B X \<in> set evs; evs \<in> set_mr |] ==> \<exists>A. Says A B X \<in> set evs"
+apply (erule rev_mp)
+apply (erule set_mr.induct, auto)
+done
+
+lemma Gets_imp_knows_Spy:
+     "[| Gets B X \<in> set evs; evs \<in> set_mr |]  ==> X \<in> knows Spy evs"
+by (blast dest!: Gets_imp_Says Says_imp_knows_Spy)
+
+
+declare Gets_imp_knows_Spy [THEN parts.Inj, dest]
+
+subsubsection{*Proofs on keys *}
+
+text{*Spy never sees an agent's private keys! (unless it's bad at start)*}
+lemma Spy_see_private_Key [simp]:
+     "evs \<in> set_mr
+      ==> (Key(invKey (publicKey b A)) \<in> parts(knows Spy evs)) = (A \<in> bad)"
+apply (erule set_mr.induct)
+apply (auto dest!: Gets_imp_knows_Spy [THEN parts.Inj])
+done
+
+lemma Spy_analz_private_Key [simp]:
+     "evs \<in> set_mr ==>
+     (Key(invKey (publicKey b A)) \<in> analz(knows Spy evs)) = (A \<in> bad)"
+by auto
+
+declare Spy_see_private_Key [THEN [2] rev_iffD1, dest!]
+declare Spy_analz_private_Key [THEN [2] rev_iffD1, dest!]
+
+(*This is to state that the signed keys received in step 4
+  are into parts - rather than installing sign_def each time.
+  Needed in Spy_see_priSK_RCA, Spy_see_priEK and in Spy_see_priSK
+Goal "[|Gets C \<lbrace>Crypt KM1
+                (sign K \<lbrace>Agent M, Nonce NM2, Key merSK, Key merEK\<rbrace>), X\<rbrace>
+          \<in> set evs;  evs \<in> set_mr |]
+    ==> Key merSK \<in> parts (knows Spy evs) \<and>
+        Key merEK \<in> parts (knows Spy evs)"
+by (fast_tac (claset() addss (simpset())) 1);
+qed "signed_keys_in_parts";
+???*)
+
+text{*Proofs on certificates -
+  they hold, as in CR, because RCA's keys are secure*}
+
+lemma Crypt_valid_pubEK:
+     "[| Crypt (priSK RCA) {|Agent (CA i), Key EKi, onlyEnc|}
+           \<in> parts (knows Spy evs);
+         evs \<in> set_mr |] ==> EKi = pubEK (CA i)"
+apply (erule rev_mp)
+apply (erule set_mr.induct, auto)
+done
+
+lemma certificate_valid_pubEK:
+    "[| cert (CA i) EKi onlyEnc (priSK RCA) \<in> parts (knows Spy evs);
+        evs \<in> set_mr |]
+     ==> EKi = pubEK (CA i)"
+apply (unfold cert_def signCert_def)
+apply (blast dest!: Crypt_valid_pubEK)
+done
+
+lemma Crypt_valid_pubSK:
+     "[| Crypt (priSK RCA) {|Agent (CA i), Key SKi, onlySig|}
+           \<in> parts (knows Spy evs);
+         evs \<in> set_mr |] ==> SKi = pubSK (CA i)"
+apply (erule rev_mp)
+apply (erule set_mr.induct, auto)
+done
+
+lemma certificate_valid_pubSK:
+    "[| cert (CA i) SKi onlySig (priSK RCA) \<in> parts (knows Spy evs);
+        evs \<in> set_mr |] ==> SKi = pubSK (CA i)"
+apply (unfold cert_def signCert_def)
+apply (blast dest!: Crypt_valid_pubSK)
+done
+
+lemma Gets_certificate_valid:
+     "[| Gets A {| X, cert (CA i) EKi onlyEnc (priSK RCA),
+                      cert (CA i) SKi onlySig (priSK RCA)|} \<in> set evs;
+         evs \<in> set_mr |]
+      ==> EKi = pubEK (CA i) & SKi = pubSK (CA i)"
+by (blast dest: certificate_valid_pubEK certificate_valid_pubSK)
+
+
+text{*Nobody can have used non-existent keys!*}
+lemma new_keys_not_used [rule_format,simp]:
+     "evs \<in> set_mr
+      ==> Key K \<notin> used evs --> K \<in> symKeys -->
+          K \<notin> keysFor (parts (knows Spy evs))"
+apply (erule set_mr.induct, simp_all)
+apply (force dest!: usedI keysFor_parts_insert)  --{*Fake*}
+apply force  --{*Message 2*}
+apply (blast dest: Gets_certificate_valid)  --{*Message 3*}
+apply force  --{*Message 4*}
+done
+
+
+subsubsection{*New Versions: As Above, but Generalized with the Kk Argument*}
+
+lemma gen_new_keys_not_used [rule_format]:
+     "evs \<in> set_mr
+      ==> Key K \<notin> used evs --> K \<in> symKeys -->
+          K \<notin> keysFor (parts (Key`KK Un knows Spy evs))"
+by auto
+
+lemma gen_new_keys_not_analzd:
+     "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_mr |]
+      ==> K \<notin> keysFor (analz (Key`KK Un knows Spy evs))"
+by (blast intro: keysFor_mono [THEN [2] rev_subsetD]
+          dest: gen_new_keys_not_used)
+
+lemma analz_Key_image_insert_eq:
+     "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_mr |]
+      ==> analz (Key ` (insert K KK) \<union> knows Spy evs) =
+          insert (Key K) (analz (Key ` KK \<union> knows Spy evs))"
+by (simp add: gen_new_keys_not_analzd)
+
+
+lemma Crypt_parts_imp_used:
+     "[|Crypt K X \<in> parts (knows Spy evs);
+        K \<in> symKeys; evs \<in> set_mr |] ==> Key K \<in> used evs"
+apply (rule ccontr)
+apply (force dest: new_keys_not_used Crypt_imp_invKey_keysFor)
+done
+
+lemma Crypt_analz_imp_used:
+     "[|Crypt K X \<in> analz (knows Spy evs);
+        K \<in> symKeys; evs \<in> set_mr |] ==> Key K \<in> used evs"
+by (blast intro: Crypt_parts_imp_used)
+
+text{*Rewriting rule for private encryption keys.  Analogous rewriting rules
+for other keys aren't needed.*}
+
+lemma parts_image_priEK:
+     "[|Key (priEK (CA i)) \<in> parts (Key`KK Un (knows Spy evs));
+        evs \<in> set_mr|] ==> priEK (CA i) \<in> KK | CA i \<in> bad"
+by auto
+
+text{*trivial proof because (priEK (CA i)) never appears even in (parts evs)*}
+lemma analz_image_priEK:
+     "evs \<in> set_mr ==>
+          (Key (priEK (CA i)) \<in> analz (Key`KK Un (knows Spy evs))) =
+          (priEK (CA i) \<in> KK | CA i \<in> bad)"
+by (blast dest!: parts_image_priEK intro: analz_mono [THEN [2] rev_subsetD])
+
+
+subsection{*Secrecy of Session Keys*}
+
+text{*This holds because if (priEK (CA i)) appears in any traffic then it must
+  be known to the Spy, by @{text Spy_see_private_Key}*}
+lemma merK_neq_priEK:
+     "[|Key merK \<notin> analz (knows Spy evs);
+        Key merK \<in> parts (knows Spy evs);
+        evs \<in> set_mr|] ==> merK \<noteq> priEK C"
+by blast
+
+text{*Lemma for message 4: either merK is compromised (when we don't care)
+  or else merK hasn't been used to encrypt K.*}
+lemma msg4_priEK_disj:
+     "[|Gets B {|Crypt KM1
+                       (sign K {|Agent M, Nonce NM2, Key merSK, Key merEK|}),
+                 Y|} \<in> set evs;
+        evs \<in> set_mr|]
+  ==> (Key merSK \<in> analz (knows Spy evs) | merSK \<notin> range(\<lambda>C. priEK C))
+   &  (Key merEK \<in> analz (knows Spy evs) | merEK \<notin> range(\<lambda>C. priEK C))"
+apply (unfold sign_def)
+apply (blast dest: merK_neq_priEK)
+done
+
+
+lemma Key_analz_image_Key_lemma:
+     "P --> (Key K \<in> analz (Key`KK Un H)) --> (K\<in>KK | Key K \<in> analz H)
+      ==>
+      P --> (Key K \<in> analz (Key`KK Un H)) = (K\<in>KK | Key K \<in> analz H)"
+by (blast intro: analz_mono [THEN [2] rev_subsetD])
+
+lemma symKey_compromise:
+     "evs \<in> set_mr ==>
+      (\<forall>SK KK. SK \<in> symKeys \<longrightarrow> (\<forall>K \<in> KK. K \<notin> range(\<lambda>C. priEK C)) -->
+               (Key SK \<in> analz (Key`KK Un (knows Spy evs))) =
+               (SK \<in> KK | Key SK \<in> analz (knows Spy evs)))"
+apply (erule set_mr.induct)
+apply (safe del: impI intro!: Key_analz_image_Key_lemma [THEN impI])
+apply (drule_tac [7] msg4_priEK_disj)
+apply (frule_tac [6] Gets_certificate_valid)
+apply (safe del: impI)
+apply (simp_all del: image_insert image_Un imp_disjL
+         add: analz_image_keys_simps abbrev_simps analz_knows_absorb
+              analz_knows_absorb2 analz_Key_image_insert_eq notin_image_iff
+              Spy_analz_private_Key analz_image_priEK)
+  --{*5 seconds on a 1.6GHz machine*}
+apply spy_analz  --{*Fake*}
+apply auto  --{*Message 3*}
+done
+
+lemma symKey_secrecy [rule_format]:
+     "[|CA i \<notin> bad; K \<in> symKeys;  evs \<in> set_mr|]
+      ==> \<forall>X m. Says (Merchant m) (CA i) X \<in> set evs -->
+                Key K \<in> parts{X} -->
+                Merchant m \<notin> bad -->
+                Key K \<notin> analz (knows Spy evs)"
+apply (erule set_mr.induct)
+apply (drule_tac [7] msg4_priEK_disj)
+apply (frule_tac [6] Gets_certificate_valid)
+apply (safe del: impI)
+apply (simp_all del: image_insert image_Un imp_disjL
+         add: analz_image_keys_simps abbrev_simps analz_knows_absorb
+              analz_knows_absorb2 analz_Key_image_insert_eq
+              symKey_compromise notin_image_iff Spy_analz_private_Key
+              analz_image_priEK)
+apply spy_analz  --{*Fake*}
+apply force  --{*Message 1*}
+apply (auto intro: analz_into_parts [THEN usedI] in_parts_Says_imp_used)  --{*Message 3*}
+done
+
+subsection{*Unicity *}
+
+lemma msg4_Says_imp_Notes:
+ "[|Says (CA i) M {|sign (priSK (CA i)) {|Agent M, Nonce NM2, Agent (CA i)|},
+                    cert  M      merSK    onlySig (priSK (CA i)),
+                    cert  M      merEK    onlyEnc (priSK (CA i)),
+                    cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|} \<in> set evs;
+    evs \<in> set_mr |]
+  ==> Notes (CA i) (Key merSK) \<in> set evs
+   &  Notes (CA i) (Key merEK) \<in> set evs"
+apply (erule rev_mp)
+apply (erule set_mr.induct)
+apply (simp_all (no_asm_simp))
+done
+
+text{*Unicity of merSK wrt a given CA:
+  merSK uniquely identifies the other components, including merEK*}
+lemma merSK_unicity:
+ "[|Says (CA i) M {|sign (priSK(CA i)) {|Agent M, Nonce NM2, Agent (CA i)|},
+                    cert  M      merSK    onlySig (priSK (CA i)),
+                    cert  M      merEK    onlyEnc (priSK (CA i)),
+                    cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|} \<in> set evs;
+    Says (CA i) M' {|sign (priSK(CA i)) {|Agent M', Nonce NM2', Agent (CA i)|},
+                    cert  M'      merSK    onlySig (priSK (CA i)),
+                    cert  M'      merEK'    onlyEnc (priSK (CA i)),
+                    cert (CA i) (pubSK(CA i)) onlySig (priSK RCA)|} \<in> set evs;
+    evs \<in> set_mr |] ==> M=M' & NM2=NM2' & merEK=merEK'"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule set_mr.induct)
+apply (simp_all (no_asm_simp))
+apply (blast dest!: msg4_Says_imp_Notes)
+done
+
+text{*Unicity of merEK wrt a given CA:
+  merEK uniquely identifies the other components, including merSK*}
+lemma merEK_unicity:
+ "[|Says (CA i) M {|sign (priSK(CA i)) {|Agent M, Nonce NM2, Agent (CA i)|},
+                    cert  M      merSK    onlySig (priSK (CA i)),
+                    cert  M      merEK    onlyEnc (priSK (CA i)),
+                    cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)|} \<in> set evs;
+    Says (CA i) M' {|sign (priSK(CA i)) {|Agent M', Nonce NM2', Agent (CA i)|},
+                     cert  M'      merSK'    onlySig (priSK (CA i)),
+                     cert  M'      merEK    onlyEnc (priSK (CA i)),
+                     cert (CA i) (pubSK(CA i)) onlySig (priSK RCA)|} \<in> set evs;
+    evs \<in> set_mr |] 
+  ==> M=M' & NM2=NM2' & merSK=merSK'"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule set_mr.induct)
+apply (simp_all (no_asm_simp))
+apply (blast dest!: msg4_Says_imp_Notes)
+done
+
+
+text{* -No interest on secrecy of nonces: they appear to be used
+    only for freshness.
+   -No interest on secrecy of merSK or merEK, as in CR.
+   -There's no equivalent of the PAN*}
+
+
+subsection{*Primary Goals of Merchant Registration *}
+
+subsubsection{*The merchant's certificates really were created by the CA,
+provided the CA is uncompromised *}
+
+text{*The assumption @{term "CA i \<noteq> RCA"} is required: step 2 uses 
+  certificates of the same form.*}
+lemma certificate_merSK_valid_lemma [intro]:
+     "[|Crypt (priSK (CA i)) {|Agent M, Key merSK, onlySig|}
+          \<in> parts (knows Spy evs);
+        CA i \<notin> bad;  CA i \<noteq> RCA;  evs \<in> set_mr|]
+ ==> \<exists>X Y Z. Says (CA i) M
+                  {|X, cert M merSK onlySig (priSK (CA i)), Y, Z|} \<in> set evs"
+apply (erule rev_mp)
+apply (erule set_mr.induct)
+apply (simp_all (no_asm_simp))
+apply auto
+done
+
+lemma certificate_merSK_valid:
+     "[| cert M merSK onlySig (priSK (CA i)) \<in> parts (knows Spy evs);
+         CA i \<notin> bad;  CA i \<noteq> RCA;  evs \<in> set_mr|]
+ ==> \<exists>X Y Z. Says (CA i) M
+                  {|X, cert M merSK onlySig (priSK (CA i)), Y, Z|} \<in> set evs"
+by auto
+
+lemma certificate_merEK_valid_lemma [intro]:
+     "[|Crypt (priSK (CA i)) {|Agent M, Key merEK, onlyEnc|}
+          \<in> parts (knows Spy evs);
+        CA i \<notin> bad;  CA i \<noteq> RCA;  evs \<in> set_mr|]
+ ==> \<exists>X Y Z. Says (CA i) M
+                  {|X, Y, cert M merEK onlyEnc (priSK (CA i)), Z|} \<in> set evs"
+apply (erule rev_mp)
+apply (erule set_mr.induct)
+apply (simp_all (no_asm_simp))
+apply auto
+done
+
+lemma certificate_merEK_valid:
+     "[| cert M merEK onlyEnc (priSK (CA i)) \<in> parts (knows Spy evs);
+         CA i \<notin> bad;  CA i \<noteq> RCA;  evs \<in> set_mr|]
+ ==> \<exists>X Y Z. Says (CA i) M
+                  {|X, Y, cert M merEK onlyEnc (priSK (CA i)), Z|} \<in> set evs"
+by auto
+
+text{*The two certificates - for merSK and for merEK - cannot be proved to
+  have originated together*}
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SET_Protocol/Message_SET.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,957 @@
+(*  Title:      HOL/SET_Protocol/Message_SET.thy
+    Author:     Giampaolo Bella
+    Author:     Fabio Massacci
+    Author:     Lawrence C Paulson
+*)
+
+header{*The Message Theory, Modified for SET*}
+
+theory Message_SET
+imports Main Nat_Int_Bij
+begin
+
+subsection{*General Lemmas*}
+
+text{*Needed occasionally with @{text spy_analz_tac}, e.g. in
+     @{text analz_insert_Key_newK}*}
+
+lemma Un_absorb3 [simp] : "A \<union> (B \<union> A) = B \<union> A"
+by blast
+
+text{*Collapses redundant cases in the huge protocol proofs*}
+lemmas disj_simps = disj_comms disj_left_absorb disj_assoc 
+
+text{*Effective with assumptions like @{term "K \<notin> range pubK"} and 
+   @{term "K \<notin> invKey`range pubK"}*}
+lemma notin_image_iff: "(y \<notin> f`I) = (\<forall>i\<in>I. f i \<noteq> y)"
+by blast
+
+text{*Effective with the assumption @{term "KK \<subseteq> - (range(invKey o pubK))"} *}
+lemma disjoint_image_iff: "(A <= - (f`I)) = (\<forall>i\<in>I. f i \<notin> A)"
+by blast
+
+
+
+types
+  key = nat
+
+consts
+  all_symmetric :: bool        --{*true if all keys are symmetric*}
+  invKey        :: "key=>key"  --{*inverse of a symmetric key*}
+
+specification (invKey)
+  invKey [simp]: "invKey (invKey K) = K"
+  invKey_symmetric: "all_symmetric --> invKey = id"
+    by (rule exI [of _ id], auto)
+
+
+text{*The inverse of a symmetric key is itself; that of a public key
+      is the private key and vice versa*}
+
+constdefs
+  symKeys :: "key set"
+  "symKeys == {K. invKey K = K}"
+
+text{*Agents. We allow any number of certification authorities, cardholders
+            merchants, and payment gateways.*}
+datatype
+  agent = CA nat | Cardholder nat | Merchant nat | PG nat | Spy
+
+text{*Messages*}
+datatype
+     msg = Agent  agent     --{*Agent names*}
+         | Number nat       --{*Ordinary integers, timestamps, ...*}
+         | Nonce  nat       --{*Unguessable nonces*}
+         | Pan    nat       --{*Unguessable Primary Account Numbers (??)*}
+         | Key    key       --{*Crypto keys*}
+         | Hash   msg       --{*Hashing*}
+         | MPair  msg msg   --{*Compound messages*}
+         | Crypt  key msg   --{*Encryption, public- or shared-key*}
+
+
+(*Concrete syntax: messages appear as {|A,B,NA|}, etc...*)
+syntax
+  "@MTuple"      :: "['a, args] => 'a * 'b"       ("(2{|_,/ _|})")
+
+syntax (xsymbols)
+  "@MTuple"      :: "['a, args] => 'a * 'b"       ("(2\<lbrace>_,/ _\<rbrace>)")
+
+translations
+  "{|x, y, z|}"   == "{|x, {|y, z|}|}"
+  "{|x, y|}"      == "MPair x y"
+
+
+constdefs
+  nat_of_agent :: "agent => nat"
+   "nat_of_agent == agent_case (curry nat2_to_nat 0)
+                               (curry nat2_to_nat 1)
+                               (curry nat2_to_nat 2)
+                               (curry nat2_to_nat 3)
+                               (nat2_to_nat (4,0))"
+    --{*maps each agent to a unique natural number, for specifications*}
+
+text{*The function is indeed injective*}
+lemma inj_nat_of_agent: "inj nat_of_agent"
+by (simp add: nat_of_agent_def inj_on_def curry_def
+              nat2_to_nat_inj [THEN inj_eq]  split: agent.split) 
+
+
+constdefs
+  (*Keys useful to decrypt elements of a message set*)
+  keysFor :: "msg set => key set"
+  "keysFor H == invKey ` {K. \<exists>X. Crypt K X \<in> H}"
+
+subsubsection{*Inductive definition of all "parts" of a message.*}
+
+inductive_set
+  parts :: "msg set => msg set"
+  for H :: "msg set"
+  where
+    Inj [intro]:               "X \<in> H ==> X \<in> parts H"
+  | Fst:         "{|X,Y|}   \<in> parts H ==> X \<in> parts H"
+  | Snd:         "{|X,Y|}   \<in> parts H ==> Y \<in> parts H"
+  | Body:        "Crypt K X \<in> parts H ==> X \<in> parts H"
+
+
+(*Monotonicity*)
+lemma parts_mono: "G<=H ==> parts(G) <= parts(H)"
+apply auto
+apply (erule parts.induct)
+apply (auto dest: Fst Snd Body)
+done
+
+
+subsubsection{*Inverse of keys*}
+
+(*Equations hold because constructors are injective; cannot prove for all f*)
+lemma Key_image_eq [simp]: "(Key x \<in> Key`A) = (x\<in>A)"
+by auto
+
+lemma Nonce_Key_image_eq [simp]: "(Nonce x \<notin> Key`A)"
+by auto
+
+lemma Cardholder_image_eq [simp]: "(Cardholder x \<in> Cardholder`A) = (x \<in> A)"
+by auto
+
+lemma CA_image_eq [simp]: "(CA x \<in> CA`A) = (x \<in> A)"
+by auto
+
+lemma Pan_image_eq [simp]: "(Pan x \<in> Pan`A) = (x \<in> A)"
+by auto
+
+lemma Pan_Key_image_eq [simp]: "(Pan x \<notin> Key`A)"
+by auto
+
+lemma Nonce_Pan_image_eq [simp]: "(Nonce x \<notin> Pan`A)"
+by auto
+
+lemma invKey_eq [simp]: "(invKey K = invKey K') = (K=K')"
+apply safe
+apply (drule_tac f = invKey in arg_cong, simp)
+done
+
+
+subsection{*keysFor operator*}
+
+lemma keysFor_empty [simp]: "keysFor {} = {}"
+by (unfold keysFor_def, blast)
+
+lemma keysFor_Un [simp]: "keysFor (H \<union> H') = keysFor H \<union> keysFor H'"
+by (unfold keysFor_def, blast)
+
+lemma keysFor_UN [simp]: "keysFor (\<Union>i\<in>A. H i) = (\<Union>i\<in>A. keysFor (H i))"
+by (unfold keysFor_def, blast)
+
+(*Monotonicity*)
+lemma keysFor_mono: "G\<subseteq>H ==> keysFor(G) \<subseteq> keysFor(H)"
+by (unfold keysFor_def, blast)
+
+lemma keysFor_insert_Agent [simp]: "keysFor (insert (Agent A) H) = keysFor H"
+by (unfold keysFor_def, auto)
+
+lemma keysFor_insert_Nonce [simp]: "keysFor (insert (Nonce N) H) = keysFor H"
+by (unfold keysFor_def, auto)
+
+lemma keysFor_insert_Number [simp]: "keysFor (insert (Number N) H) = keysFor H"
+by (unfold keysFor_def, auto)
+
+lemma keysFor_insert_Key [simp]: "keysFor (insert (Key K) H) = keysFor H"
+by (unfold keysFor_def, auto)
+
+lemma keysFor_insert_Pan [simp]: "keysFor (insert (Pan A) H) = keysFor H"
+by (unfold keysFor_def, auto)
+
+lemma keysFor_insert_Hash [simp]: "keysFor (insert (Hash X) H) = keysFor H"
+by (unfold keysFor_def, auto)
+
+lemma keysFor_insert_MPair [simp]: "keysFor (insert {|X,Y|} H) = keysFor H"
+by (unfold keysFor_def, auto)
+
+lemma keysFor_insert_Crypt [simp]:
+    "keysFor (insert (Crypt K X) H) = insert (invKey K) (keysFor H)"
+by (unfold keysFor_def, auto)
+
+lemma keysFor_image_Key [simp]: "keysFor (Key`E) = {}"
+by (unfold keysFor_def, auto)
+
+lemma Crypt_imp_invKey_keysFor: "Crypt K X \<in> H ==> invKey K \<in> keysFor H"
+by (unfold keysFor_def, blast)
+
+
+subsection{*Inductive relation "parts"*}
+
+lemma MPair_parts:
+     "[| {|X,Y|} \<in> parts H;
+         [| X \<in> parts H; Y \<in> parts H |] ==> P |] ==> P"
+by (blast dest: parts.Fst parts.Snd)
+
+declare MPair_parts [elim!]  parts.Body [dest!]
+text{*NB These two rules are UNSAFE in the formal sense, as they discard the
+     compound message.  They work well on THIS FILE.
+  @{text MPair_parts} is left as SAFE because it speeds up proofs.
+  The Crypt rule is normally kept UNSAFE to avoid breaking up certificates.*}
+
+lemma parts_increasing: "H \<subseteq> parts(H)"
+by blast
+
+lemmas parts_insertI = subset_insertI [THEN parts_mono, THEN subsetD, standard]
+
+lemma parts_empty [simp]: "parts{} = {}"
+apply safe
+apply (erule parts.induct, blast+)
+done
+
+lemma parts_emptyE [elim!]: "X\<in> parts{} ==> P"
+by simp
+
+(*WARNING: loops if H = {Y}, therefore must not be repeated!*)
+lemma parts_singleton: "X\<in> parts H ==> \<exists>Y\<in>H. X\<in> parts {Y}"
+by (erule parts.induct, fast+)
+
+
+subsubsection{*Unions*}
+
+lemma parts_Un_subset1: "parts(G) \<union> parts(H) \<subseteq> parts(G \<union> H)"
+by (intro Un_least parts_mono Un_upper1 Un_upper2)
+
+lemma parts_Un_subset2: "parts(G \<union> H) \<subseteq> parts(G) \<union> parts(H)"
+apply (rule subsetI)
+apply (erule parts.induct, blast+)
+done
+
+lemma parts_Un [simp]: "parts(G \<union> H) = parts(G) \<union> parts(H)"
+by (intro equalityI parts_Un_subset1 parts_Un_subset2)
+
+lemma parts_insert: "parts (insert X H) = parts {X} \<union> parts H"
+apply (subst insert_is_Un [of _ H])
+apply (simp only: parts_Un)
+done
+
+(*TWO inserts to avoid looping.  This rewrite is better than nothing.
+  Not suitable for Addsimps: its behaviour can be strange.*)
+lemma parts_insert2:
+     "parts (insert X (insert Y H)) = parts {X} \<union> parts {Y} \<union> parts H"
+apply (simp add: Un_assoc)
+apply (simp add: parts_insert [symmetric])
+done
+
+lemma parts_UN_subset1: "(\<Union>x\<in>A. parts(H x)) \<subseteq> parts(\<Union>x\<in>A. H x)"
+by (intro UN_least parts_mono UN_upper)
+
+lemma parts_UN_subset2: "parts(\<Union>x\<in>A. H x) \<subseteq> (\<Union>x\<in>A. parts(H x))"
+apply (rule subsetI)
+apply (erule parts.induct, blast+)
+done
+
+lemma parts_UN [simp]: "parts(\<Union>x\<in>A. H x) = (\<Union>x\<in>A. parts(H x))"
+by (intro equalityI parts_UN_subset1 parts_UN_subset2)
+
+(*Added to simplify arguments to parts, analz and synth.
+  NOTE: the UN versions are no longer used!*)
+
+
+text{*This allows @{text blast} to simplify occurrences of
+  @{term "parts(G\<union>H)"} in the assumption.*}
+declare parts_Un [THEN equalityD1, THEN subsetD, THEN UnE, elim!]
+
+
+lemma parts_insert_subset: "insert X (parts H) \<subseteq> parts(insert X H)"
+by (blast intro: parts_mono [THEN [2] rev_subsetD])
+
+subsubsection{*Idempotence and transitivity*}
+
+lemma parts_partsD [dest!]: "X\<in> parts (parts H) ==> X\<in> parts H"
+by (erule parts.induct, blast+)
+
+lemma parts_idem [simp]: "parts (parts H) = parts H"
+by blast
+
+lemma parts_trans: "[| X\<in> parts G;  G \<subseteq> parts H |] ==> X\<in> parts H"
+by (drule parts_mono, blast)
+
+(*Cut*)
+lemma parts_cut:
+     "[| Y\<in> parts (insert X G);  X\<in> parts H |] ==> Y\<in> parts (G \<union> H)"
+by (erule parts_trans, auto)
+
+lemma parts_cut_eq [simp]: "X\<in> parts H ==> parts (insert X H) = parts H"
+by (force dest!: parts_cut intro: parts_insertI)
+
+
+subsubsection{*Rewrite rules for pulling out atomic messages*}
+
+lemmas parts_insert_eq_I = equalityI [OF subsetI parts_insert_subset]
+
+
+lemma parts_insert_Agent [simp]:
+     "parts (insert (Agent agt) H) = insert (Agent agt) (parts H)"
+apply (rule parts_insert_eq_I)
+apply (erule parts.induct, auto)
+done
+
+lemma parts_insert_Nonce [simp]:
+     "parts (insert (Nonce N) H) = insert (Nonce N) (parts H)"
+apply (rule parts_insert_eq_I)
+apply (erule parts.induct, auto)
+done
+
+lemma parts_insert_Number [simp]:
+     "parts (insert (Number N) H) = insert (Number N) (parts H)"
+apply (rule parts_insert_eq_I)
+apply (erule parts.induct, auto)
+done
+
+lemma parts_insert_Key [simp]:
+     "parts (insert (Key K) H) = insert (Key K) (parts H)"
+apply (rule parts_insert_eq_I)
+apply (erule parts.induct, auto)
+done
+
+lemma parts_insert_Pan [simp]:
+     "parts (insert (Pan A) H) = insert (Pan A) (parts H)"
+apply (rule parts_insert_eq_I)
+apply (erule parts.induct, auto)
+done
+
+lemma parts_insert_Hash [simp]:
+     "parts (insert (Hash X) H) = insert (Hash X) (parts H)"
+apply (rule parts_insert_eq_I)
+apply (erule parts.induct, auto)
+done
+
+lemma parts_insert_Crypt [simp]:
+     "parts (insert (Crypt K X) H) =
+          insert (Crypt K X) (parts (insert X H))"
+apply (rule equalityI)
+apply (rule subsetI)
+apply (erule parts.induct, auto)
+apply (erule parts.induct)
+apply (blast intro: parts.Body)+
+done
+
+lemma parts_insert_MPair [simp]:
+     "parts (insert {|X,Y|} H) =
+          insert {|X,Y|} (parts (insert X (insert Y H)))"
+apply (rule equalityI)
+apply (rule subsetI)
+apply (erule parts.induct, auto)
+apply (erule parts.induct)
+apply (blast intro: parts.Fst parts.Snd)+
+done
+
+lemma parts_image_Key [simp]: "parts (Key`N) = Key`N"
+apply auto
+apply (erule parts.induct, auto)
+done
+
+lemma parts_image_Pan [simp]: "parts (Pan`A) = Pan`A"
+apply auto
+apply (erule parts.induct, auto)
+done
+
+
+(*In any message, there is an upper bound N on its greatest nonce.*)
+lemma msg_Nonce_supply: "\<exists>N. \<forall>n. N\<le>n --> Nonce n \<notin> parts {msg}"
+apply (induct_tac "msg")
+apply (simp_all (no_asm_simp) add: exI parts_insert2)
+(*MPair case: blast_tac works out the necessary sum itself!*)
+prefer 2 apply (blast elim!: add_leE)
+(*Nonce case*)
+apply (rule_tac x = "N + Suc nat" in exI)
+apply (auto elim!: add_leE)
+done
+
+(* Ditto, for numbers.*)
+lemma msg_Number_supply: "\<exists>N. \<forall>n. N<=n --> Number n \<notin> parts {msg}"
+apply (induct_tac "msg")
+apply (simp_all (no_asm_simp) add: exI parts_insert2)
+prefer 2 apply (blast elim!: add_leE)
+apply (rule_tac x = "N + Suc nat" in exI, auto)
+done
+
+subsection{*Inductive relation "analz"*}
+
+text{*Inductive definition of "analz" -- what can be broken down from a set of
+    messages, including keys.  A form of downward closure.  Pairs can
+    be taken apart; messages decrypted with known keys.*}
+
+inductive_set
+  analz :: "msg set => msg set"
+  for H :: "msg set"
+  where
+    Inj [intro,simp] :    "X \<in> H ==> X \<in> analz H"
+  | Fst:     "{|X,Y|} \<in> analz H ==> X \<in> analz H"
+  | Snd:     "{|X,Y|} \<in> analz H ==> Y \<in> analz H"
+  | Decrypt [dest]:
+             "[|Crypt K X \<in> analz H; Key(invKey K): analz H|] ==> X \<in> analz H"
+
+
+(*Monotonicity; Lemma 1 of Lowe's paper*)
+lemma analz_mono: "G<=H ==> analz(G) <= analz(H)"
+apply auto
+apply (erule analz.induct)
+apply (auto dest: Fst Snd)
+done
+
+text{*Making it safe speeds up proofs*}
+lemma MPair_analz [elim!]:
+     "[| {|X,Y|} \<in> analz H;
+             [| X \<in> analz H; Y \<in> analz H |] ==> P
+          |] ==> P"
+by (blast dest: analz.Fst analz.Snd)
+
+lemma analz_increasing: "H \<subseteq> analz(H)"
+by blast
+
+lemma analz_subset_parts: "analz H \<subseteq> parts H"
+apply (rule subsetI)
+apply (erule analz.induct, blast+)
+done
+
+lemmas analz_into_parts = analz_subset_parts [THEN subsetD, standard]
+
+lemmas not_parts_not_analz = analz_subset_parts [THEN contra_subsetD, standard]
+
+
+lemma parts_analz [simp]: "parts (analz H) = parts H"
+apply (rule equalityI)
+apply (rule analz_subset_parts [THEN parts_mono, THEN subset_trans], simp)
+apply (blast intro: analz_increasing [THEN parts_mono, THEN subsetD])
+done
+
+lemma analz_parts [simp]: "analz (parts H) = parts H"
+apply auto
+apply (erule analz.induct, auto)
+done
+
+lemmas analz_insertI = subset_insertI [THEN analz_mono, THEN [2] rev_subsetD, standard]
+
+subsubsection{*General equational properties*}
+
+lemma analz_empty [simp]: "analz{} = {}"
+apply safe
+apply (erule analz.induct, blast+)
+done
+
+(*Converse fails: we can analz more from the union than from the
+  separate parts, as a key in one might decrypt a message in the other*)
+lemma analz_Un: "analz(G) \<union> analz(H) \<subseteq> analz(G \<union> H)"
+by (intro Un_least analz_mono Un_upper1 Un_upper2)
+
+lemma analz_insert: "insert X (analz H) \<subseteq> analz(insert X H)"
+by (blast intro: analz_mono [THEN [2] rev_subsetD])
+
+subsubsection{*Rewrite rules for pulling out atomic messages*}
+
+lemmas analz_insert_eq_I = equalityI [OF subsetI analz_insert]
+
+lemma analz_insert_Agent [simp]:
+     "analz (insert (Agent agt) H) = insert (Agent agt) (analz H)"
+apply (rule analz_insert_eq_I)
+apply (erule analz.induct, auto)
+done
+
+lemma analz_insert_Nonce [simp]:
+     "analz (insert (Nonce N) H) = insert (Nonce N) (analz H)"
+apply (rule analz_insert_eq_I)
+apply (erule analz.induct, auto)
+done
+
+lemma analz_insert_Number [simp]:
+     "analz (insert (Number N) H) = insert (Number N) (analz H)"
+apply (rule analz_insert_eq_I)
+apply (erule analz.induct, auto)
+done
+
+lemma analz_insert_Hash [simp]:
+     "analz (insert (Hash X) H) = insert (Hash X) (analz H)"
+apply (rule analz_insert_eq_I)
+apply (erule analz.induct, auto)
+done
+
+(*Can only pull out Keys if they are not needed to decrypt the rest*)
+lemma analz_insert_Key [simp]:
+    "K \<notin> keysFor (analz H) ==>
+          analz (insert (Key K) H) = insert (Key K) (analz H)"
+apply (unfold keysFor_def)
+apply (rule analz_insert_eq_I)
+apply (erule analz.induct, auto)
+done
+
+lemma analz_insert_MPair [simp]:
+     "analz (insert {|X,Y|} H) =
+          insert {|X,Y|} (analz (insert X (insert Y H)))"
+apply (rule equalityI)
+apply (rule subsetI)
+apply (erule analz.induct, auto)
+apply (erule analz.induct)
+apply (blast intro: analz.Fst analz.Snd)+
+done
+
+(*Can pull out enCrypted message if the Key is not known*)
+lemma analz_insert_Crypt:
+     "Key (invKey K) \<notin> analz H
+      ==> analz (insert (Crypt K X) H) = insert (Crypt K X) (analz H)"
+apply (rule analz_insert_eq_I)
+apply (erule analz.induct, auto)
+done
+
+lemma analz_insert_Pan [simp]:
+     "analz (insert (Pan A) H) = insert (Pan A) (analz H)"
+apply (rule analz_insert_eq_I)
+apply (erule analz.induct, auto)
+done
+
+lemma lemma1: "Key (invKey K) \<in> analz H ==>
+               analz (insert (Crypt K X) H) \<subseteq>
+               insert (Crypt K X) (analz (insert X H))"
+apply (rule subsetI)
+apply (erule_tac x = x in analz.induct, auto)
+done
+
+lemma lemma2: "Key (invKey K) \<in> analz H ==>
+               insert (Crypt K X) (analz (insert X H)) \<subseteq>
+               analz (insert (Crypt K X) H)"
+apply auto
+apply (erule_tac x = x in analz.induct, auto)
+apply (blast intro: analz_insertI analz.Decrypt)
+done
+
+lemma analz_insert_Decrypt:
+     "Key (invKey K) \<in> analz H ==>
+               analz (insert (Crypt K X) H) =
+               insert (Crypt K X) (analz (insert X H))"
+by (intro equalityI lemma1 lemma2)
+
+(*Case analysis: either the message is secure, or it is not!
+  Effective, but can cause subgoals to blow up!
+  Use with split_if;  apparently split_tac does not cope with patterns
+  such as "analz (insert (Crypt K X) H)" *)
+lemma analz_Crypt_if [simp]:
+     "analz (insert (Crypt K X) H) =
+          (if (Key (invKey K) \<in> analz H)
+           then insert (Crypt K X) (analz (insert X H))
+           else insert (Crypt K X) (analz H))"
+by (simp add: analz_insert_Crypt analz_insert_Decrypt)
+
+
+(*This rule supposes "for the sake of argument" that we have the key.*)
+lemma analz_insert_Crypt_subset:
+     "analz (insert (Crypt K X) H) \<subseteq>
+           insert (Crypt K X) (analz (insert X H))"
+apply (rule subsetI)
+apply (erule analz.induct, auto)
+done
+
+lemma analz_image_Key [simp]: "analz (Key`N) = Key`N"
+apply auto
+apply (erule analz.induct, auto)
+done
+
+lemma analz_image_Pan [simp]: "analz (Pan`A) = Pan`A"
+apply auto
+apply (erule analz.induct, auto)
+done
+
+
+subsubsection{*Idempotence and transitivity*}
+
+lemma analz_analzD [dest!]: "X\<in> analz (analz H) ==> X\<in> analz H"
+by (erule analz.induct, blast+)
+
+lemma analz_idem [simp]: "analz (analz H) = analz H"
+by blast
+
+lemma analz_trans: "[| X\<in> analz G;  G \<subseteq> analz H |] ==> X\<in> analz H"
+by (drule analz_mono, blast)
+
+(*Cut; Lemma 2 of Lowe*)
+lemma analz_cut: "[| Y\<in> analz (insert X H);  X\<in> analz H |] ==> Y\<in> analz H"
+by (erule analz_trans, blast)
+
+(*Cut can be proved easily by induction on
+   "Y: analz (insert X H) ==> X: analz H --> Y: analz H"
+*)
+
+(*This rewrite rule helps in the simplification of messages that involve
+  the forwarding of unknown components (X).  Without it, removing occurrences
+  of X can be very complicated. *)
+lemma analz_insert_eq: "X\<in> analz H ==> analz (insert X H) = analz H"
+by (blast intro: analz_cut analz_insertI)
+
+
+text{*A congruence rule for "analz"*}
+
+lemma analz_subset_cong:
+     "[| analz G \<subseteq> analz G'; analz H \<subseteq> analz H'
+               |] ==> analz (G \<union> H) \<subseteq> analz (G' \<union> H')"
+apply clarify
+apply (erule analz.induct)
+apply (best intro: analz_mono [THEN subsetD])+
+done
+
+lemma analz_cong:
+     "[| analz G = analz G'; analz H = analz H'
+               |] ==> analz (G \<union> H) = analz (G' \<union> H')"
+by (intro equalityI analz_subset_cong, simp_all)
+
+lemma analz_insert_cong:
+     "analz H = analz H' ==> analz(insert X H) = analz(insert X H')"
+by (force simp only: insert_def intro!: analz_cong)
+
+(*If there are no pairs or encryptions then analz does nothing*)
+lemma analz_trivial:
+     "[| \<forall>X Y. {|X,Y|} \<notin> H;  \<forall>X K. Crypt K X \<notin> H |] ==> analz H = H"
+apply safe
+apply (erule analz.induct, blast+)
+done
+
+(*These two are obsolete (with a single Spy) but cost little to prove...*)
+lemma analz_UN_analz_lemma:
+     "X\<in> analz (\<Union>i\<in>A. analz (H i)) ==> X\<in> analz (\<Union>i\<in>A. H i)"
+apply (erule analz.induct)
+apply (blast intro: analz_mono [THEN [2] rev_subsetD])+
+done
+
+lemma analz_UN_analz [simp]: "analz (\<Union>i\<in>A. analz (H i)) = analz (\<Union>i\<in>A. H i)"
+by (blast intro: analz_UN_analz_lemma analz_mono [THEN [2] rev_subsetD])
+
+
+subsection{*Inductive relation "synth"*}
+
+text{*Inductive definition of "synth" -- what can be built up from a set of
+    messages.  A form of upward closure.  Pairs can be built, messages
+    encrypted with known keys.  Agent names are public domain.
+    Numbers can be guessed, but Nonces cannot be.*}
+
+inductive_set
+  synth :: "msg set => msg set"
+  for H :: "msg set"
+  where
+    Inj    [intro]:   "X \<in> H ==> X \<in> synth H"
+  | Agent  [intro]:   "Agent agt \<in> synth H"
+  | Number [intro]:   "Number n  \<in> synth H"
+  | Hash   [intro]:   "X \<in> synth H ==> Hash X \<in> synth H"
+  | MPair  [intro]:   "[|X \<in> synth H;  Y \<in> synth H|] ==> {|X,Y|} \<in> synth H"
+  | Crypt  [intro]:   "[|X \<in> synth H;  Key(K) \<in> H|] ==> Crypt K X \<in> synth H"
+
+(*Monotonicity*)
+lemma synth_mono: "G<=H ==> synth(G) <= synth(H)"
+apply auto
+apply (erule synth.induct)
+apply (auto dest: Fst Snd Body)
+done
+
+(*NO Agent_synth, as any Agent name can be synthesized.  Ditto for Number*)
+inductive_cases Nonce_synth [elim!]: "Nonce n \<in> synth H"
+inductive_cases Key_synth   [elim!]: "Key K \<in> synth H"
+inductive_cases Hash_synth  [elim!]: "Hash X \<in> synth H"
+inductive_cases MPair_synth [elim!]: "{|X,Y|} \<in> synth H"
+inductive_cases Crypt_synth [elim!]: "Crypt K X \<in> synth H"
+inductive_cases Pan_synth   [elim!]: "Pan A \<in> synth H"
+
+
+lemma synth_increasing: "H \<subseteq> synth(H)"
+by blast
+
+subsubsection{*Unions*}
+
+(*Converse fails: we can synth more from the union than from the
+  separate parts, building a compound message using elements of each.*)
+lemma synth_Un: "synth(G) \<union> synth(H) \<subseteq> synth(G \<union> H)"
+by (intro Un_least synth_mono Un_upper1 Un_upper2)
+
+lemma synth_insert: "insert X (synth H) \<subseteq> synth(insert X H)"
+by (blast intro: synth_mono [THEN [2] rev_subsetD])
+
+subsubsection{*Idempotence and transitivity*}
+
+lemma synth_synthD [dest!]: "X\<in> synth (synth H) ==> X\<in> synth H"
+by (erule synth.induct, blast+)
+
+lemma synth_idem: "synth (synth H) = synth H"
+by blast
+
+lemma synth_trans: "[| X\<in> synth G;  G \<subseteq> synth H |] ==> X\<in> synth H"
+by (drule synth_mono, blast)
+
+(*Cut; Lemma 2 of Lowe*)
+lemma synth_cut: "[| Y\<in> synth (insert X H);  X\<in> synth H |] ==> Y\<in> synth H"
+by (erule synth_trans, blast)
+
+lemma Agent_synth [simp]: "Agent A \<in> synth H"
+by blast
+
+lemma Number_synth [simp]: "Number n \<in> synth H"
+by blast
+
+lemma Nonce_synth_eq [simp]: "(Nonce N \<in> synth H) = (Nonce N \<in> H)"
+by blast
+
+lemma Key_synth_eq [simp]: "(Key K \<in> synth H) = (Key K \<in> H)"
+by blast
+
+lemma Crypt_synth_eq [simp]: "Key K \<notin> H ==> (Crypt K X \<in> synth H) = (Crypt K X \<in> H)"
+by blast
+
+lemma Pan_synth_eq [simp]: "(Pan A \<in> synth H) = (Pan A \<in> H)"
+by blast
+
+lemma keysFor_synth [simp]:
+    "keysFor (synth H) = keysFor H \<union> invKey`{K. Key K \<in> H}"
+by (unfold keysFor_def, blast)
+
+
+subsubsection{*Combinations of parts, analz and synth*}
+
+lemma parts_synth [simp]: "parts (synth H) = parts H \<union> synth H"
+apply (rule equalityI)
+apply (rule subsetI)
+apply (erule parts.induct)
+apply (blast intro: synth_increasing [THEN parts_mono, THEN subsetD]
+                    parts.Fst parts.Snd parts.Body)+
+done
+
+lemma analz_analz_Un [simp]: "analz (analz G \<union> H) = analz (G \<union> H)"
+apply (intro equalityI analz_subset_cong)+
+apply simp_all
+done
+
+lemma analz_synth_Un [simp]: "analz (synth G \<union> H) = analz (G \<union> H) \<union> synth G"
+apply (rule equalityI)
+apply (rule subsetI)
+apply (erule analz.induct)
+prefer 5 apply (blast intro: analz_mono [THEN [2] rev_subsetD])
+apply (blast intro: analz.Fst analz.Snd analz.Decrypt)+
+done
+
+lemma analz_synth [simp]: "analz (synth H) = analz H \<union> synth H"
+apply (cut_tac H = "{}" in analz_synth_Un)
+apply (simp (no_asm_use))
+done
+
+
+subsubsection{*For reasoning about the Fake rule in traces*}
+
+lemma parts_insert_subset_Un: "X\<in> G ==> parts(insert X H) \<subseteq> parts G \<union> parts H"
+by (rule subset_trans [OF parts_mono parts_Un_subset2], blast)
+
+(*More specifically for Fake.  Very occasionally we could do with a version
+  of the form  parts{X} \<subseteq> synth (analz H) \<union> parts H *)
+lemma Fake_parts_insert: "X \<in> synth (analz H) ==>
+      parts (insert X H) \<subseteq> synth (analz H) \<union> parts H"
+apply (drule parts_insert_subset_Un)
+apply (simp (no_asm_use))
+apply blast
+done
+
+lemma Fake_parts_insert_in_Un:
+     "[|Z \<in> parts (insert X H);  X: synth (analz H)|] 
+      ==> Z \<in>  synth (analz H) \<union> parts H";
+by (blast dest: Fake_parts_insert [THEN subsetD, dest])
+
+(*H is sometimes (Key ` KK \<union> spies evs), so can't put G=H*)
+lemma Fake_analz_insert:
+     "X\<in> synth (analz G) ==>
+      analz (insert X H) \<subseteq> synth (analz G) \<union> analz (G \<union> H)"
+apply (rule subsetI)
+apply (subgoal_tac "x \<in> analz (synth (analz G) \<union> H) ")
+prefer 2 apply (blast intro: analz_mono [THEN [2] rev_subsetD] analz_mono [THEN synth_mono, THEN [2] rev_subsetD])
+apply (simp (no_asm_use))
+apply blast
+done
+
+lemma analz_conj_parts [simp]:
+     "(X \<in> analz H & X \<in> parts H) = (X \<in> analz H)"
+by (blast intro: analz_subset_parts [THEN subsetD])
+
+lemma analz_disj_parts [simp]:
+     "(X \<in> analz H | X \<in> parts H) = (X \<in> parts H)"
+by (blast intro: analz_subset_parts [THEN subsetD])
+
+(*Without this equation, other rules for synth and analz would yield
+  redundant cases*)
+lemma MPair_synth_analz [iff]:
+     "({|X,Y|} \<in> synth (analz H)) =
+      (X \<in> synth (analz H) & Y \<in> synth (analz H))"
+by blast
+
+lemma Crypt_synth_analz:
+     "[| Key K \<in> analz H;  Key (invKey K) \<in> analz H |]
+       ==> (Crypt K X \<in> synth (analz H)) = (X \<in> synth (analz H))"
+by blast
+
+
+lemma Hash_synth_analz [simp]:
+     "X \<notin> synth (analz H)
+      ==> (Hash{|X,Y|} \<in> synth (analz H)) = (Hash{|X,Y|} \<in> analz H)"
+by blast
+
+
+(*We do NOT want Crypt... messages broken up in protocols!!*)
+declare parts.Body [rule del]
+
+
+text{*Rewrites to push in Key and Crypt messages, so that other messages can
+    be pulled out using the @{text analz_insert} rules*}
+
+lemmas pushKeys [standard] =
+  insert_commute [of "Key K" "Agent C"]
+  insert_commute [of "Key K" "Nonce N"]
+  insert_commute [of "Key K" "Number N"]
+  insert_commute [of "Key K" "Pan PAN"]
+  insert_commute [of "Key K" "Hash X"]
+  insert_commute [of "Key K" "MPair X Y"]
+  insert_commute [of "Key K" "Crypt X K'"]
+
+lemmas pushCrypts [standard] =
+  insert_commute [of "Crypt X K" "Agent C"]
+  insert_commute [of "Crypt X K" "Nonce N"]
+  insert_commute [of "Crypt X K" "Number N"]
+  insert_commute [of "Crypt X K" "Pan PAN"]
+  insert_commute [of "Crypt X K" "Hash X'"]
+  insert_commute [of "Crypt X K" "MPair X' Y"]
+
+text{*Cannot be added with @{text "[simp]"} -- messages should not always be
+  re-ordered.*}
+lemmas pushes = pushKeys pushCrypts
+
+
+subsection{*Tactics useful for many protocol proofs*}
+(*<*)
+ML
+{*
+structure MessageSET =
+struct
+
+(*Prove base case (subgoal i) and simplify others.  A typical base case
+  concerns  Crypt K X \<notin> Key`shrK`bad  and cannot be proved by rewriting
+  alone.*)
+fun prove_simple_subgoals_tac (cs, ss) i =
+    force_tac (cs, ss addsimps [@{thm image_eq_UN}]) i THEN
+    ALLGOALS (asm_simp_tac ss)
+
+(*Analysis of Fake cases.  Also works for messages that forward unknown parts,
+  but this application is no longer necessary if analz_insert_eq is used.
+  Abstraction over i is ESSENTIAL: it delays the dereferencing of claset
+  DEPENDS UPON "X" REFERRING TO THE FRADULENT MESSAGE *)
+
+fun impOfSubs th = th RSN (2, @{thm rev_subsetD})
+
+(*Apply rules to break down assumptions of the form
+  Y \<in> parts(insert X H)  and  Y \<in> analz(insert X H)
+*)
+val Fake_insert_tac =
+    dresolve_tac [impOfSubs @{thm Fake_analz_insert},
+                  impOfSubs @{thm Fake_parts_insert}] THEN'
+    eresolve_tac [asm_rl, @{thm synth.Inj}];
+
+fun Fake_insert_simp_tac ss i =
+    REPEAT (Fake_insert_tac i) THEN asm_full_simp_tac ss i;
+
+fun atomic_spy_analz_tac (cs,ss) = SELECT_GOAL
+    (Fake_insert_simp_tac ss 1
+     THEN
+     IF_UNSOLVED (Blast.depth_tac
+                  (cs addIs [@{thm analz_insertI},
+                                   impOfSubs @{thm analz_subset_parts}]) 4 1))
+
+fun spy_analz_tac (cs,ss) i =
+  DETERM
+   (SELECT_GOAL
+     (EVERY
+      [  (*push in occurrences of X...*)
+       (REPEAT o CHANGED)
+           (res_inst_tac (Simplifier.the_context ss)
+             [(("x", 1), "X")] (insert_commute RS ssubst) 1),
+       (*...allowing further simplifications*)
+       simp_tac ss 1,
+       REPEAT (FIRSTGOAL (resolve_tac [allI,impI,notI,conjI,iffI])),
+       DEPTH_SOLVE (atomic_spy_analz_tac (cs,ss) 1)]) i)
+
+end
+*}
+(*>*)
+
+
+(*By default only o_apply is built-in.  But in the presence of eta-expansion
+  this means that some terms displayed as (f o g) will be rewritten, and others
+  will not!*)
+declare o_def [simp]
+
+
+lemma Crypt_notin_image_Key [simp]: "Crypt K X \<notin> Key ` A"
+by auto
+
+lemma Hash_notin_image_Key [simp] :"Hash X \<notin> Key ` A"
+by auto
+
+lemma synth_analz_mono: "G<=H ==> synth (analz(G)) <= synth (analz(H))"
+by (simp add: synth_mono analz_mono)
+
+lemma Fake_analz_eq [simp]:
+     "X \<in> synth(analz H) ==> synth (analz (insert X H)) = synth (analz H)"
+apply (drule Fake_analz_insert[of _ _ "H"])
+apply (simp add: synth_increasing[THEN Un_absorb2])
+apply (drule synth_mono)
+apply (simp add: synth_idem)
+apply (blast intro: synth_analz_mono [THEN [2] rev_subsetD])
+done
+
+text{*Two generalizations of @{text analz_insert_eq}*}
+lemma gen_analz_insert_eq [rule_format]:
+     "X \<in> analz H ==> ALL G. H \<subseteq> G --> analz (insert X G) = analz G";
+by (blast intro: analz_cut analz_insertI analz_mono [THEN [2] rev_subsetD])
+
+lemma synth_analz_insert_eq [rule_format]:
+     "X \<in> synth (analz H)
+      ==> ALL G. H \<subseteq> G --> (Key K \<in> analz (insert X G)) = (Key K \<in> analz G)";
+apply (erule synth.induct)
+apply (simp_all add: gen_analz_insert_eq subset_trans [OF _ subset_insertI])
+done
+
+lemma Fake_parts_sing:
+     "X \<in> synth (analz H) ==> parts{X} \<subseteq> synth (analz H) \<union> parts H";
+apply (rule subset_trans)
+ apply (erule_tac [2] Fake_parts_insert)
+apply (simp add: parts_mono)
+done
+
+lemmas Fake_parts_sing_imp_Un = Fake_parts_sing [THEN [2] rev_subsetD]
+
+method_setup spy_analz = {*
+    Scan.succeed (fn ctxt =>
+        SIMPLE_METHOD' (MessageSET.spy_analz_tac (clasimpset_of ctxt))) *}
+    "for proving the Fake case when analz is involved"
+
+method_setup atomic_spy_analz = {*
+    Scan.succeed (fn ctxt =>
+        SIMPLE_METHOD' (MessageSET.atomic_spy_analz_tac (clasimpset_of ctxt))) *}
+    "for debugging spy_analz"
+
+method_setup Fake_insert_simp = {*
+    Scan.succeed (fn ctxt =>
+        SIMPLE_METHOD' (MessageSET.Fake_insert_simp_tac (simpset_of ctxt))) *}
+    "for debugging spy_analz"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SET_Protocol/Public_SET.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,529 @@
+(*  Title:      HOL/SET_Protocol/Public_SET.thy
+    Author:     Giampaolo Bella
+    Author:     Fabio Massacci
+    Author:     Lawrence C Paulson
+*)
+
+header{*The Public-Key Theory, Modified for SET*}
+
+theory Public_SET
+imports Event_SET
+begin
+
+subsection{*Symmetric and Asymmetric Keys*}
+
+text{*definitions influenced by the wish to assign asymmetric keys 
+  - since the beginning - only to RCA and CAs, namely we need a partial 
+  function on type Agent*}
+
+
+text{*The SET specs mention two signature keys for CAs - we only have one*}
+
+consts
+  publicKey :: "[bool, agent] => key"
+    --{*the boolean is TRUE if a signing key*}
+
+syntax
+  pubEK :: "agent => key"
+  pubSK :: "agent => key"
+  priEK :: "agent => key"
+  priSK :: "agent => key"
+
+translations
+  "pubEK"  == "publicKey False"
+  "pubSK"  == "publicKey True"
+
+  (*BEWARE!! priEK, priSK DON'T WORK with inj, range, image, etc.*)
+  "priEK A"  == "invKey (pubEK A)"
+  "priSK A"  == "invKey (pubSK A)"
+
+text{*By freeness of agents, no two agents have the same key. Since
+ @{term "True\<noteq>False"}, no agent has the same signing and encryption keys.*}
+
+specification (publicKey)
+  injective_publicKey:
+    "publicKey b A = publicKey c A' ==> b=c & A=A'"
+(*<*)
+   apply (rule exI [of _ "%b A. 2 * nat_of_agent A + (if b then 1 else 0)"]) 
+   apply (auto simp add: inj_on_def inj_nat_of_agent [THEN inj_eq] split: agent.split) 
+   apply (drule_tac f="%x. x mod 2" in arg_cong, simp add: mod_Suc)+
+(*or this, but presburger won't abstract out the function applications
+   apply presburger+
+*)
+   done                       
+(*>*)
+
+axioms
+  (*No private key equals any public key (essential to ensure that private
+    keys are private!) *)
+  privateKey_neq_publicKey [iff]:
+      "invKey (publicKey b A) \<noteq> publicKey b' A'"
+
+declare privateKey_neq_publicKey [THEN not_sym, iff]
+
+  
+subsection{*Initial Knowledge*}
+
+text{*This information is not necessary.  Each protocol distributes any needed
+certificates, and anyway our proofs require a formalization of the Spy's 
+knowledge only.  However, the initial knowledge is as follows:
+   All agents know RCA's public keys;
+   RCA and CAs know their own respective keys;
+   RCA (has already certified and therefore) knows all CAs public keys; 
+   Spy knows all keys of all bad agents.*}
+primrec    
+(*<*)
+  initState_CA:
+    "initState (CA i)  =
+       (if i=0 then Key ` ({priEK RCA, priSK RCA} Un
+                            pubEK ` (range CA) Un pubSK ` (range CA))
+        else {Key (priEK (CA i)), Key (priSK (CA i)),
+              Key (pubEK (CA i)), Key (pubSK (CA i)),
+              Key (pubEK RCA), Key (pubSK RCA)})"
+
+  initState_Cardholder:
+    "initState (Cardholder i)  =    
+       {Key (priEK (Cardholder i)), Key (priSK (Cardholder i)),
+        Key (pubEK (Cardholder i)), Key (pubSK (Cardholder i)),
+        Key (pubEK RCA), Key (pubSK RCA)}"
+
+  initState_Merchant:
+    "initState (Merchant i)  =    
+       {Key (priEK (Merchant i)), Key (priSK (Merchant i)),
+        Key (pubEK (Merchant i)), Key (pubSK (Merchant i)),
+        Key (pubEK RCA), Key (pubSK RCA)}"
+
+  initState_PG:
+    "initState (PG i)  = 
+       {Key (priEK (PG i)), Key (priSK (PG i)),
+        Key (pubEK (PG i)), Key (pubSK (PG i)),
+        Key (pubEK RCA), Key (pubSK RCA)}"
+(*>*)
+  initState_Spy:
+    "initState Spy = Key ` (invKey ` pubEK ` bad Un
+                            invKey ` pubSK ` bad Un
+                            range pubEK Un range pubSK)"
+
+
+text{*Injective mapping from agents to PANs: an agent can have only one card*}
+
+consts  pan :: "agent => nat"
+
+specification (pan)
+  inj_pan: "inj pan"
+  --{*No two agents have the same PAN*}
+(*<*)
+   apply (rule exI [of _ "nat_of_agent"]) 
+   apply (simp add: inj_on_def inj_nat_of_agent [THEN inj_eq]) 
+   done
+(*>*)
+
+declare inj_pan [THEN inj_eq, iff]
+
+consts
+  XOR :: "nat*nat => nat"  --{*no properties are assumed of exclusive-or*}
+
+
+subsection{*Signature Primitives*}
+
+constdefs 
+
+ (* Signature = Message + signed Digest *)
+  sign :: "[key, msg]=>msg"
+    "sign K X == {|X, Crypt K (Hash X) |}"
+
+ (* Signature Only = signed Digest Only *)
+  signOnly :: "[key, msg]=>msg"
+    "signOnly K X == Crypt K (Hash X)"
+
+ (* Signature for Certificates = Message + signed Message *)
+  signCert :: "[key, msg]=>msg"
+    "signCert K X == {|X, Crypt K X |}"
+
+ (* Certification Authority's Certificate.
+    Contains agent name, a key, a number specifying the key's target use,
+              a key to sign the entire certificate.
+
+    Should prove if signK=priSK RCA and C=CA i,
+                  then Ka=pubEK i or pubSK i depending on T  ??
+ *)
+  cert :: "[agent, key, msg, key] => msg"
+    "cert A Ka T signK == signCert signK {|Agent A, Key Ka, T|}"
+
+
+ (* Cardholder's Certificate.
+    Contains a PAN, the certified key Ka, the PANSecret PS,
+    a number specifying the target use for Ka, the signing key signK.
+ *)
+  certC :: "[nat, key, nat, msg, key] => msg"
+    "certC PAN Ka PS T signK ==
+     signCert signK {|Hash {|Nonce PS, Pan PAN|}, Key Ka, T|}"
+
+  (*cert and certA have no repeated elements, so they could be translations,
+    but that's tricky and makes proofs slower*)
+
+syntax
+  "onlyEnc" :: msg      
+  "onlySig" :: msg
+  "authCode" :: msg
+
+translations
+  "onlyEnc"   == "Number 0"
+  "onlySig"  == "Number (Suc 0)"
+  "authCode" == "Number (Suc (Suc 0))"
+
+subsection{*Encryption Primitives*}
+
+constdefs
+
+  EXcrypt :: "[key,key,msg,msg] => msg"
+  --{*Extra Encryption*}
+    (*K: the symmetric key   EK: the public encryption key*)
+    "EXcrypt K EK M m ==
+       {|Crypt K {|M, Hash m|}, Crypt EK {|Key K, m|}|}"
+
+  EXHcrypt :: "[key,key,msg,msg] => msg"
+  --{*Extra Encryption with Hashing*}
+    (*K: the symmetric key   EK: the public encryption key*)
+    "EXHcrypt K EK M m ==
+       {|Crypt K {|M, Hash m|}, Crypt EK {|Key K, m, Hash M|}|}"
+
+  Enc :: "[key,key,key,msg] => msg"
+  --{*Simple Encapsulation with SIGNATURE*}
+    (*SK: the sender's signing key
+      K: the symmetric key
+      EK: the public encryption key*)
+    "Enc SK K EK M ==
+       {|Crypt K (sign SK M), Crypt EK (Key K)|}"
+
+  EncB :: "[key,key,key,msg,msg] => msg"
+  --{*Encapsulation with Baggage.  Keys as above, and baggage b.*}
+    "EncB SK K EK M b == 
+       {|Enc SK K EK {|M, Hash b|}, b|}"
+
+
+subsection{*Basic Properties of pubEK, pubSK, priEK and priSK *}
+
+lemma publicKey_eq_iff [iff]:
+     "(publicKey b A = publicKey b' A') = (b=b' & A=A')"
+by (blast dest: injective_publicKey)
+
+lemma privateKey_eq_iff [iff]:
+     "(invKey (publicKey b A) = invKey (publicKey b' A')) = (b=b' & A=A')"
+by auto
+
+lemma not_symKeys_publicKey [iff]: "publicKey b A \<notin> symKeys"
+by (simp add: symKeys_def)
+
+lemma not_symKeys_privateKey [iff]: "invKey (publicKey b A) \<notin> symKeys"
+by (simp add: symKeys_def)
+
+lemma symKeys_invKey_eq [simp]: "K \<in> symKeys ==> invKey K = K"
+by (simp add: symKeys_def)
+
+lemma symKeys_invKey_iff [simp]: "(invKey K \<in> symKeys) = (K \<in> symKeys)"
+by (unfold symKeys_def, auto)
+
+text{*Can be slow (or even loop) as a simprule*}
+lemma symKeys_neq_imp_neq: "(K \<in> symKeys) \<noteq> (K' \<in> symKeys) ==> K \<noteq> K'"
+by blast
+
+text{*These alternatives to @{text symKeys_neq_imp_neq} don't seem any better
+in practice.*}
+lemma publicKey_neq_symKey: "K \<in> symKeys ==> publicKey b A \<noteq> K"
+by blast
+
+lemma symKey_neq_publicKey: "K \<in> symKeys ==> K \<noteq> publicKey b A"
+by blast
+
+lemma privateKey_neq_symKey: "K \<in> symKeys ==> invKey (publicKey b A) \<noteq> K"
+by blast
+
+lemma symKey_neq_privateKey: "K \<in> symKeys ==> K \<noteq> invKey (publicKey b A)"
+by blast
+
+lemma analz_symKeys_Decrypt:
+     "[| Crypt K X \<in> analz H;  K \<in> symKeys;  Key K \<in> analz H |]  
+      ==> X \<in> analz H"
+by auto
+
+
+subsection{*"Image" Equations That Hold for Injective Functions *}
+
+lemma invKey_image_eq [iff]: "(invKey x \<in> invKey`A) = (x\<in>A)"
+by auto
+
+text{*holds because invKey is injective*}
+lemma publicKey_image_eq [iff]:
+     "(publicKey b A \<in> publicKey c ` AS) = (b=c & A\<in>AS)"
+by auto
+
+lemma privateKey_image_eq [iff]:
+     "(invKey (publicKey b A) \<in> invKey ` publicKey c ` AS) = (b=c & A\<in>AS)"
+by auto
+
+lemma privateKey_notin_image_publicKey [iff]:
+     "invKey (publicKey b A) \<notin> publicKey c ` AS"
+by auto
+
+lemma publicKey_notin_image_privateKey [iff]:
+     "publicKey b A \<notin> invKey ` publicKey c ` AS"
+by auto
+
+lemma keysFor_parts_initState [simp]: "keysFor (parts (initState C)) = {}"
+apply (simp add: keysFor_def)
+apply (induct_tac "C")
+apply (auto intro: range_eqI)
+done
+
+text{*for proving @{text new_keys_not_used}*}
+lemma keysFor_parts_insert:
+     "[| K \<in> keysFor (parts (insert X H));  X \<in> synth (analz H) |]  
+      ==> K \<in> keysFor (parts H) | Key (invKey K) \<in> parts H"
+by (force dest!: 
+         parts_insert_subset_Un [THEN keysFor_mono, THEN [2] rev_subsetD]
+         analz_subset_parts [THEN keysFor_mono, THEN [2] rev_subsetD] 
+            intro: analz_into_parts)
+
+lemma Crypt_imp_keysFor [intro]:
+     "[|K \<in> symKeys; Crypt K X \<in> H|] ==> K \<in> keysFor H"
+by (drule Crypt_imp_invKey_keysFor, simp)
+
+text{*Agents see their own private keys!*}
+lemma privateKey_in_initStateCA [iff]:
+     "Key (invKey (publicKey b A)) \<in> initState A"
+by (case_tac "A", auto)
+
+text{*Agents see their own public keys!*}
+lemma publicKey_in_initStateCA [iff]: "Key (publicKey b A) \<in> initState A"
+by (case_tac "A", auto)
+
+text{*RCA sees CAs' public keys! *}
+lemma pubK_CA_in_initState_RCA [iff]:
+     "Key (publicKey b (CA i)) \<in> initState RCA"
+by auto
+
+
+text{*Spy knows all public keys*}
+lemma knows_Spy_pubEK_i [iff]: "Key (publicKey b A) \<in> knows Spy evs"
+apply (induct_tac "evs")
+apply (simp_all add: imageI knows_Cons split add: event.split)
+done
+
+declare knows_Spy_pubEK_i [THEN analz.Inj, iff]
+                            (*needed????*)
+
+text{*Spy sees private keys of bad agents! [and obviously public keys too]*}
+lemma knows_Spy_bad_privateKey [intro!]:
+     "A \<in> bad ==> Key (invKey (publicKey b A)) \<in> knows Spy evs"
+by (rule initState_subset_knows [THEN subsetD], simp)
+
+
+subsection{*Fresh Nonces for Possibility Theorems*}
+
+lemma Nonce_notin_initState [iff]: "Nonce N \<notin> parts (initState B)"
+by (induct_tac "B", auto)
+
+lemma Nonce_notin_used_empty [simp]: "Nonce N \<notin> used []"
+by (simp add: used_Nil)
+
+text{*In any trace, there is an upper bound N on the greatest nonce in use.*}
+lemma Nonce_supply_lemma: "\<exists>N. \<forall>n. N<=n --> Nonce n \<notin> used evs"
+apply (induct_tac "evs")
+apply (rule_tac x = 0 in exI)
+apply (simp_all add: used_Cons split add: event.split, safe)
+apply (rule msg_Nonce_supply [THEN exE], blast elim!: add_leE)+
+done
+
+lemma Nonce_supply1: "\<exists>N. Nonce N \<notin> used evs"
+by (rule Nonce_supply_lemma [THEN exE], blast)
+
+lemma Nonce_supply: "Nonce (@ N. Nonce N \<notin> used evs) \<notin> used evs"
+apply (rule Nonce_supply_lemma [THEN exE])
+apply (rule someI, fast)
+done
+
+
+subsection{*Specialized Methods for Possibility Theorems*}
+
+ML
+{*
+(*Tactic for possibility theorems*)
+fun possibility_tac ctxt =
+    REPEAT (*omit used_Says so that Nonces start from different traces!*)
+    (ALLGOALS (simp_tac (simpset_of ctxt delsimps [@{thm used_Says}, @{thm used_Notes}]))
+     THEN
+     REPEAT_FIRST (eq_assume_tac ORELSE' 
+                   resolve_tac [refl, conjI, @{thm Nonce_supply}]))
+
+(*For harder protocols (such as SET_CR!), where we have to set up some
+  nonces and keys initially*)
+fun basic_possibility_tac ctxt =
+    REPEAT 
+    (ALLGOALS (asm_simp_tac (simpset_of ctxt setSolver safe_solver))
+     THEN
+     REPEAT_FIRST (resolve_tac [refl, conjI]))
+*}
+
+method_setup possibility = {*
+    Scan.succeed (SIMPLE_METHOD o possibility_tac) *}
+    "for proving possibility theorems"
+
+method_setup basic_possibility = {*
+    Scan.succeed (SIMPLE_METHOD o basic_possibility_tac) *}
+    "for proving possibility theorems"
+
+
+subsection{*Specialized Rewriting for Theorems About @{term analz} and Image*}
+
+lemma insert_Key_singleton: "insert (Key K) H = Key ` {K} Un H"
+by blast
+
+lemma insert_Key_image:
+     "insert (Key K) (Key`KK Un C) = Key ` (insert K KK) Un C"
+by blast
+
+text{*Needed for @{text DK_fresh_not_KeyCryptKey}*}
+lemma publicKey_in_used [iff]: "Key (publicKey b A) \<in> used evs"
+by auto
+
+lemma privateKey_in_used [iff]: "Key (invKey (publicKey b A)) \<in> used evs"
+by (blast intro!: initState_into_used)
+
+text{*Reverse the normal simplification of "image" to build up (not break down)
+  the set of keys.  Based on @{text analz_image_freshK_ss}, but simpler.*}
+lemmas analz_image_keys_simps =
+       simp_thms mem_simps --{*these two allow its use with @{text "only:"}*}
+       image_insert [THEN sym] image_Un [THEN sym] 
+       rangeI symKeys_neq_imp_neq
+       insert_Key_singleton insert_Key_image Un_assoc [THEN sym]
+
+
+(*General lemmas proved by Larry*)
+
+subsection{*Controlled Unfolding of Abbreviations*}
+
+text{*A set is expanded only if a relation is applied to it*}
+lemma def_abbrev_simp_relation:
+     "A == B ==> (A \<in> X) = (B \<in> X) &  
+                 (u = A) = (u = B) &  
+                 (A = u) = (B = u)"
+by auto
+
+text{*A set is expanded only if one of the given functions is applied to it*}
+lemma def_abbrev_simp_function:
+     "A == B  
+      ==> parts (insert A X) = parts (insert B X) &  
+          analz (insert A X) = analz (insert B X) &  
+          keysFor (insert A X) = keysFor (insert B X)"
+by auto
+
+subsubsection{*Special Simplification Rules for @{term signCert}*}
+
+text{*Avoids duplicating X and its components!*}
+lemma parts_insert_signCert:
+     "parts (insert (signCert K X) H) =  
+      insert {|X, Crypt K X|} (parts (insert (Crypt K X) H))"
+by (simp add: signCert_def insert_commute [of X])
+
+text{*Avoids a case split! [X is always available]*}
+lemma analz_insert_signCert:
+     "analz (insert (signCert K X) H) =  
+      insert {|X, Crypt K X|} (insert (Crypt K X) (analz (insert X H)))"
+by (simp add: signCert_def insert_commute [of X])
+
+lemma keysFor_insert_signCert: "keysFor (insert (signCert K X) H) = keysFor H"
+by (simp add: signCert_def)
+
+text{*Controlled rewrite rules for @{term signCert}, just the definitions
+  of the others. Encryption primitives are just expanded, despite their huge
+  redundancy!*}
+lemmas abbrev_simps [simp] =
+    parts_insert_signCert analz_insert_signCert keysFor_insert_signCert
+    sign_def     [THEN def_abbrev_simp_relation]
+    sign_def     [THEN def_abbrev_simp_function]
+    signCert_def [THEN def_abbrev_simp_relation]
+    signCert_def [THEN def_abbrev_simp_function]
+    certC_def    [THEN def_abbrev_simp_relation]
+    certC_def    [THEN def_abbrev_simp_function]
+    cert_def     [THEN def_abbrev_simp_relation]
+    cert_def     [THEN def_abbrev_simp_function]
+    EXcrypt_def  [THEN def_abbrev_simp_relation]
+    EXcrypt_def  [THEN def_abbrev_simp_function]
+    EXHcrypt_def [THEN def_abbrev_simp_relation]
+    EXHcrypt_def [THEN def_abbrev_simp_function]
+    Enc_def      [THEN def_abbrev_simp_relation]
+    Enc_def      [THEN def_abbrev_simp_function]
+    EncB_def     [THEN def_abbrev_simp_relation]
+    EncB_def     [THEN def_abbrev_simp_function]
+
+
+subsubsection{*Elimination Rules for Controlled Rewriting *}
+
+lemma Enc_partsE: 
+     "!!R. [|Enc SK K EK M \<in> parts H;  
+             [|Crypt K (sign SK M) \<in> parts H;  
+               Crypt EK (Key K) \<in> parts H|] ==> R|]  
+           ==> R"
+
+by (unfold Enc_def, blast)
+
+lemma EncB_partsE: 
+     "!!R. [|EncB SK K EK M b \<in> parts H;  
+             [|Crypt K (sign SK {|M, Hash b|}) \<in> parts H;  
+               Crypt EK (Key K) \<in> parts H;  
+               b \<in> parts H|] ==> R|]  
+           ==> R"
+by (unfold EncB_def Enc_def, blast)
+
+lemma EXcrypt_partsE: 
+     "!!R. [|EXcrypt K EK M m \<in> parts H;  
+             [|Crypt K {|M, Hash m|} \<in> parts H;  
+               Crypt EK {|Key K, m|} \<in> parts H|] ==> R|]  
+           ==> R"
+by (unfold EXcrypt_def, blast)
+
+
+subsection{*Lemmas to Simplify Expressions Involving @{term analz} *}
+
+lemma analz_knows_absorb:
+     "Key K \<in> analz (knows Spy evs)  
+      ==> analz (Key ` (insert K H) \<union> knows Spy evs) =  
+          analz (Key ` H \<union> knows Spy evs)"
+by (simp add: analz_insert_eq Un_upper2 [THEN analz_mono, THEN subsetD])
+
+lemma analz_knows_absorb2:
+     "Key K \<in> analz (knows Spy evs)  
+      ==> analz (Key ` (insert X (insert K H)) \<union> knows Spy evs) =  
+          analz (Key ` (insert X H) \<union> knows Spy evs)"
+apply (subst insert_commute)
+apply (erule analz_knows_absorb)
+done
+
+lemma analz_insert_subset_eq:
+     "[|X \<in> analz (knows Spy evs);  knows Spy evs \<subseteq> H|]  
+      ==> analz (insert X H) = analz H"
+apply (rule analz_insert_eq)
+apply (blast intro: analz_mono [THEN [2] rev_subsetD])
+done
+
+lemmas analz_insert_simps = 
+         analz_insert_subset_eq Un_upper2
+         subset_insertI [THEN [2] subset_trans] 
+
+
+subsection{*Freshness Lemmas*}
+
+lemma in_parts_Says_imp_used:
+     "[|Key K \<in> parts {X}; Says A B X \<in> set evs|] ==> Key K \<in> used evs"
+by (blast intro: parts_trans dest!: Says_imp_knows_Spy [THEN parts.Inj])
+
+text{*A useful rewrite rule with @{term analz_image_keys_simps}*}
+lemma Crypt_notin_image_Key: "Crypt K X \<notin> Key ` KK"
+by auto
+
+lemma fresh_notin_analz_knows_Spy:
+     "Key K \<notin> used evs ==> Key K \<notin> analz (knows Spy evs)"
+by (auto dest: analz_into_parts)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SET_Protocol/Purchase.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,1172 @@
+(*  Title:      HOL/SET_Protocol/Purchase.thy
+    Author:     Giampaolo Bella
+    Author:     Fabio Massacci
+    Author:     Lawrence C Paulson
+*)
+
+header{*Purchase Phase of SET*}
+
+theory Purchase
+imports Public_SET
+begin
+
+text{*
+Note: nonces seem to consist of 20 bytes.  That includes both freshness
+challenges (Chall-EE, etc.) and important secrets (CardSecret, PANsecret)
+
+This version omits @{text LID_C} but retains @{text LID_M}. At first glance
+(Programmer's Guide page 267) it seems that both numbers are just introduced
+for the respective convenience of the Cardholder's and Merchant's
+system. However, omitting both of them would create a problem of
+identification: how can the Merchant's system know what transaction is it
+supposed to process?
+
+Further reading (Programmer's guide page 309) suggest that there is an outside
+bootstrapping message (SET initiation message) which is used by the Merchant
+and the Cardholder to agree on the actual transaction. This bootstrapping
+message is described in the SET External Interface Guide and ought to generate
+@{text LID_M}. According SET Extern Interface Guide, this number might be a
+cookie, an invoice number etc. The Programmer's Guide on page 310, states that
+in absence of @{text LID_M} the protocol must somehow ("outside SET") identify
+the transaction from OrderDesc, which is assumed to be a searchable text only
+field. Thus, it is assumed that the Merchant or the Cardholder somehow agreed
+out-of-bad on the value of @{text LID_M} (for instance a cookie in a web
+transaction etc.). This out-of-band agreement is expressed with a preliminary
+start action in which the merchant and the Cardholder agree on the appropriate
+values. Agreed values are stored with a suitable notes action.
+
+"XID is a transaction ID that is usually generated by the Merchant system,
+unless there is no PInitRes, in which case it is generated by the Cardholder
+system. It is a randomly generated 20 byte variable that is globally unique
+(statistically). Merchant and Cardholder systems shall use appropriate random
+number generators to ensure the global uniqueness of XID."
+--Programmer's Guide, page 267.
+
+PI (Payment Instruction) is the most central and sensitive data structure in
+SET. It is used to pass the data required to authorize a payment card payment
+from the Cardholder to the Payment Gateway, which will use the data to
+initiate a payment card transaction through the traditional payment card
+financial network. The data is encrypted by the Cardholder and sent via the
+Merchant, such that the data is hidden from the Merchant unless the Acquirer
+passes the data back to the Merchant.
+--Programmer's Guide, page 271.*}
+
+consts
+
+    CardSecret :: "nat => nat"
+     --{*Maps Cardholders to CardSecrets.
+         A CardSecret of 0 means no cerificate, must use unsigned format.*}
+
+    PANSecret :: "nat => nat"
+     --{*Maps Cardholders to PANSecrets.*}
+
+inductive_set
+  set_pur :: "event list set"
+where
+
+  Nil:   --{*Initial trace is empty*}
+         "[] \<in> set_pur"
+
+| Fake:  --{*The spy MAY say anything he CAN say.*}
+         "[| evsf \<in> set_pur;  X \<in> synth(analz(knows Spy evsf)) |]
+          ==> Says Spy B X  # evsf \<in> set_pur"
+
+
+| Reception: --{*If A sends a message X to B, then B might receive it*}
+             "[| evsr \<in> set_pur;  Says A B X \<in> set evsr |]
+              ==> Gets B X  # evsr \<in> set_pur"
+
+| Start: 
+      --{*Added start event which is out-of-band for SET: the Cardholder and
+          the merchant agree on the amounts and uses @{text LID_M} as an
+          identifier.
+          This is suggested by the External Interface Guide. The Programmer's
+          Guide, in absence of @{text LID_M}, states that the merchant uniquely
+          identifies the order out of some data contained in OrderDesc.*}
+   "[|evsStart \<in> set_pur;
+      Number LID_M \<notin> used evsStart;
+      C = Cardholder k; M = Merchant i; P = PG j;
+      Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
+      LID_M \<notin> range CardSecret;
+      LID_M \<notin> range PANSecret |]
+     ==> Notes C {|Number LID_M, Transaction|}
+       # Notes M {|Number LID_M, Agent P, Transaction|}
+       # evsStart \<in> set_pur"
+
+| PInitReq:
+     --{*Purchase initialization, page 72 of Formal Protocol Desc.*}
+   "[|evsPIReq \<in> set_pur;
+      Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
+      Nonce Chall_C \<notin> used evsPIReq;
+      Chall_C \<notin> range CardSecret; Chall_C \<notin> range PANSecret;
+      Notes C {|Number LID_M, Transaction |} \<in> set evsPIReq |]
+    ==> Says C M {|Number LID_M, Nonce Chall_C|} # evsPIReq \<in> set_pur"
+
+| PInitRes:
+     --{*Merchant replies with his own label XID and the encryption
+         key certificate of his chosen Payment Gateway. Page 74 of Formal
+         Protocol Desc. We use @{text LID_M} to identify Cardholder*}
+   "[|evsPIRes \<in> set_pur;
+      Gets M {|Number LID_M, Nonce Chall_C|} \<in> set evsPIRes;
+      Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
+      Notes M {|Number LID_M, Agent P, Transaction|} \<in> set evsPIRes;
+      Nonce Chall_M \<notin> used evsPIRes;
+      Chall_M \<notin> range CardSecret; Chall_M \<notin> range PANSecret;
+      Number XID \<notin> used evsPIRes;
+      XID \<notin> range CardSecret; XID \<notin> range PANSecret|]
+    ==> Says M C (sign (priSK M)
+                       {|Number LID_M, Number XID,
+                         Nonce Chall_C, Nonce Chall_M,
+                         cert P (pubEK P) onlyEnc (priSK RCA)|})
+          # evsPIRes \<in> set_pur"
+
+| PReqUns:
+      --{*UNSIGNED Purchase request (CardSecret = 0).
+        Page 79 of Formal Protocol Desc.
+        Merchant never sees the amount in clear. This holds of the real
+        protocol, where XID identifies the transaction. We omit
+        Hash{|Number XID, Nonce (CardSecret k)|} from PIHead because
+        the CardSecret is 0 and because AuthReq treated the unsigned case
+        very differently from the signed one anyway.*}
+   "!!Chall_C Chall_M OrderDesc P PurchAmt XID evsPReqU.
+    [|evsPReqU \<in> set_pur;
+      C = Cardholder k; CardSecret k = 0;
+      Key KC1 \<notin> used evsPReqU;  KC1 \<in> symKeys;
+      Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
+      HOD = Hash{|Number OrderDesc, Number PurchAmt|};
+      OIData = {|Number LID_M, Number XID, Nonce Chall_C, HOD,Nonce Chall_M|};
+      PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M|};
+      Gets C (sign (priSK M)
+                   {|Number LID_M, Number XID,
+                     Nonce Chall_C, Nonce Chall_M,
+                     cert P EKj onlyEnc (priSK RCA)|})
+        \<in> set evsPReqU;
+      Says C M {|Number LID_M, Nonce Chall_C|} \<in> set evsPReqU;
+      Notes C {|Number LID_M, Transaction|} \<in> set evsPReqU |]
+    ==> Says C M
+             {|EXHcrypt KC1 EKj {|PIHead, Hash OIData|} (Pan (pan C)),
+               OIData, Hash{|PIHead, Pan (pan C)|} |}
+          # Notes C {|Key KC1, Agent M|}
+          # evsPReqU \<in> set_pur"
+
+| PReqS:
+      --{*SIGNED Purchase request.  Page 77 of Formal Protocol Desc.
+          We could specify the equation
+          @{term "PIReqSigned = {| PIDualSigned, OIDualSigned |}"}, since the
+          Formal Desc. gives PIHead the same format in the unsigned case.
+          However, there's little point, as P treats the signed and 
+          unsigned cases differently.*}
+   "!!C Chall_C Chall_M EKj HOD KC2 LID_M M OIData
+      OIDualSigned OrderDesc P PANData PIData PIDualSigned
+      PIHead PurchAmt Transaction XID evsPReqS k.
+    [|evsPReqS \<in> set_pur;
+      C = Cardholder k;
+      CardSecret k \<noteq> 0;  Key KC2 \<notin> used evsPReqS;  KC2 \<in> symKeys;
+      Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
+      HOD = Hash{|Number OrderDesc, Number PurchAmt|};
+      OIData = {|Number LID_M, Number XID, Nonce Chall_C, HOD, Nonce Chall_M|};
+      PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M,
+                  Hash{|Number XID, Nonce (CardSecret k)|}|};
+      PANData = {|Pan (pan C), Nonce (PANSecret k)|};
+      PIData = {|PIHead, PANData|};
+      PIDualSigned = {|sign (priSK C) {|Hash PIData, Hash OIData|},
+                       EXcrypt KC2 EKj {|PIHead, Hash OIData|} PANData|};
+      OIDualSigned = {|OIData, Hash PIData|};
+      Gets C (sign (priSK M)
+                   {|Number LID_M, Number XID,
+                     Nonce Chall_C, Nonce Chall_M,
+                     cert P EKj onlyEnc (priSK RCA)|})
+        \<in> set evsPReqS;
+      Says C M {|Number LID_M, Nonce Chall_C|} \<in> set evsPReqS;
+      Notes C {|Number LID_M, Transaction|} \<in> set evsPReqS |]
+    ==> Says C M {|PIDualSigned, OIDualSigned|}
+          # Notes C {|Key KC2, Agent M|}
+          # evsPReqS \<in> set_pur"
+
+  --{*Authorization Request.  Page 92 of Formal Protocol Desc.
+    Sent in response to Purchase Request.*}
+| AuthReq:
+   "[| evsAReq \<in> set_pur;
+       Key KM \<notin> used evsAReq;  KM \<in> symKeys;
+       Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
+       HOD = Hash{|Number OrderDesc, Number PurchAmt|};
+       OIData = {|Number LID_M, Number XID, Nonce Chall_C, HOD,
+                  Nonce Chall_M|};
+       CardSecret k \<noteq> 0 -->
+         P_I = {|sign (priSK C) {|HPIData, Hash OIData|}, encPANData|};
+       Gets M {|P_I, OIData, HPIData|} \<in> set evsAReq;
+       Says M C (sign (priSK M) {|Number LID_M, Number XID,
+                                  Nonce Chall_C, Nonce Chall_M,
+                                  cert P EKj onlyEnc (priSK RCA)|})
+         \<in> set evsAReq;
+        Notes M {|Number LID_M, Agent P, Transaction|}
+           \<in> set evsAReq |]
+    ==> Says M P
+             (EncB (priSK M) KM (pubEK P)
+               {|Number LID_M, Number XID, Hash OIData, HOD|}   P_I)
+          # evsAReq \<in> set_pur"
+
+  --{*Authorization Response has two forms: for UNSIGNED and SIGNED PIs.
+    Page 99 of Formal Protocol Desc.
+    PI is a keyword (product!), so we call it @{text P_I}. The hashes HOD and
+    HOIData occur independently in @{text P_I} and in M's message.
+    The authCode in AuthRes represents the baggage of EncB, which in the
+    full protocol is [CapToken], [AcqCardMsg], [AuthToken]:
+    optional items for split shipments, recurring payments, etc.*}
+
+| AuthResUns:
+    --{*Authorization Response, UNSIGNED*}
+   "[| evsAResU \<in> set_pur;
+       C = Cardholder k; M = Merchant i;
+       Key KP \<notin> used evsAResU;  KP \<in> symKeys;
+       CardSecret k = 0;  KC1 \<in> symKeys;  KM \<in> symKeys;
+       PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M|};
+       P_I = EXHcrypt KC1 EKj {|PIHead, HOIData|} (Pan (pan C));
+       Gets P (EncB (priSK M) KM (pubEK P)
+               {|Number LID_M, Number XID, HOIData, HOD|} P_I)
+           \<in> set evsAResU |]
+   ==> Says P M
+            (EncB (priSK P) KP (pubEK M)
+              {|Number LID_M, Number XID, Number PurchAmt|}
+              authCode)
+       # evsAResU \<in> set_pur"
+
+| AuthResS:
+    --{*Authorization Response, SIGNED*}
+   "[| evsAResS \<in> set_pur;
+       C = Cardholder k;
+       Key KP \<notin> used evsAResS;  KP \<in> symKeys;
+       CardSecret k \<noteq> 0;  KC2 \<in> symKeys;  KM \<in> symKeys;
+       P_I = {|sign (priSK C) {|Hash PIData, HOIData|},
+               EXcrypt KC2 (pubEK P) {|PIHead, HOIData|} PANData|};
+       PANData = {|Pan (pan C), Nonce (PANSecret k)|};
+       PIData = {|PIHead, PANData|};
+       PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M,
+                  Hash{|Number XID, Nonce (CardSecret k)|}|};
+       Gets P (EncB (priSK M) KM (pubEK P)
+                {|Number LID_M, Number XID, HOIData, HOD|}
+               P_I)
+           \<in> set evsAResS |]
+   ==> Says P M
+            (EncB (priSK P) KP (pubEK M)
+              {|Number LID_M, Number XID, Number PurchAmt|}
+              authCode)
+       # evsAResS \<in> set_pur"
+
+| PRes:
+    --{*Purchase response.*}
+   "[| evsPRes \<in> set_pur;  KP \<in> symKeys;  M = Merchant i;
+       Transaction = {|Agent M, Agent C, Number OrderDesc, Number PurchAmt|};
+       Gets M (EncB (priSK P) KP (pubEK M)
+              {|Number LID_M, Number XID, Number PurchAmt|}
+              authCode)
+          \<in> set evsPRes;
+       Gets M {|Number LID_M, Nonce Chall_C|} \<in> set evsPRes;
+       Says M P
+            (EncB (priSK M) KM (pubEK P)
+              {|Number LID_M, Number XID, Hash OIData, HOD|} P_I)
+         \<in> set evsPRes;
+       Notes M {|Number LID_M, Agent P, Transaction|}
+          \<in> set evsPRes
+      |]
+   ==> Says M C
+         (sign (priSK M) {|Number LID_M, Number XID, Nonce Chall_C,
+                           Hash (Number PurchAmt)|})
+         # evsPRes \<in> set_pur"
+
+
+specification (CardSecret PANSecret)
+  inj_CardSecret:  "inj CardSecret"
+  inj_PANSecret:   "inj PANSecret"
+  CardSecret_neq_PANSecret: "CardSecret k \<noteq> PANSecret k'"
+    --{*No CardSecret equals any PANSecret*}
+  apply (rule_tac x="curry nat2_to_nat 0" in exI)
+  apply (rule_tac x="curry nat2_to_nat 1" in exI)
+  apply (simp add: nat2_to_nat_inj [THEN inj_eq] inj_on_def)
+  done
+
+declare Says_imp_knows_Spy [THEN parts.Inj, dest]
+declare parts.Body [dest]
+declare analz_into_parts [dest]
+declare Fake_parts_insert_in_Un [dest]
+
+declare CardSecret_neq_PANSecret [iff] 
+        CardSecret_neq_PANSecret [THEN not_sym, iff]
+declare inj_CardSecret [THEN inj_eq, iff] 
+        inj_PANSecret [THEN inj_eq, iff]
+
+
+subsection{*Possibility Properties*}
+
+lemma Says_to_Gets:
+     "Says A B X # evs \<in> set_pur ==> Gets B X # Says A B X # evs \<in> set_pur"
+by (rule set_pur.Reception, auto)
+
+text{*Possibility for UNSIGNED purchases. Note that we need to ensure
+that XID differs from OrderDesc and PurchAmt, since it is supposed to be
+a unique number!*}
+lemma possibility_Uns:
+    "[| CardSecret k = 0;
+        C = Cardholder k;  M = Merchant i;
+        Key KC \<notin> used []; Key KM \<notin> used []; Key KP \<notin> used []; 
+        KC \<in> symKeys; KM \<in> symKeys; KP \<in> symKeys; 
+        KC < KM; KM < KP;
+        Nonce Chall_C \<notin> used []; Chall_C \<notin> range CardSecret \<union> range PANSecret;
+        Nonce Chall_M \<notin> used []; Chall_M \<notin> range CardSecret \<union> range PANSecret;
+        Chall_C < Chall_M; 
+        Number LID_M \<notin> used []; LID_M \<notin> range CardSecret \<union> range PANSecret;
+        Number XID \<notin> used []; XID \<notin> range CardSecret \<union> range PANSecret;
+        LID_M < XID; XID < OrderDesc; OrderDesc < PurchAmt |] 
+   ==> \<exists>evs \<in> set_pur.
+          Says M C
+               (sign (priSK M)
+                    {|Number LID_M, Number XID, Nonce Chall_C, 
+                      Hash (Number PurchAmt)|})
+                  \<in> set evs" 
+apply (intro exI bexI)
+apply (rule_tac [2]
+        set_pur.Nil
+         [THEN set_pur.Start [of _ LID_M C k M i _ _ _ OrderDesc PurchAmt], 
+          THEN set_pur.PInitReq [of concl: C M LID_M Chall_C],
+          THEN Says_to_Gets, 
+          THEN set_pur.PInitRes [of concl: M C LID_M XID Chall_C Chall_M], 
+          THEN Says_to_Gets,
+          THEN set_pur.PReqUns [of concl: C M KC],
+          THEN Says_to_Gets, 
+          THEN set_pur.AuthReq [of concl: M "PG j" KM LID_M XID], 
+          THEN Says_to_Gets, 
+          THEN set_pur.AuthResUns [of concl: "PG j" M KP LID_M XID],
+          THEN Says_to_Gets, 
+          THEN set_pur.PRes]) 
+apply basic_possibility
+apply (simp_all add: used_Cons symKeys_neq_imp_neq) 
+done
+
+lemma possibility_S:
+    "[| CardSecret k \<noteq> 0;
+        C = Cardholder k;  M = Merchant i;
+        Key KC \<notin> used []; Key KM \<notin> used []; Key KP \<notin> used []; 
+        KC \<in> symKeys; KM \<in> symKeys; KP \<in> symKeys; 
+        KC < KM; KM < KP;
+        Nonce Chall_C \<notin> used []; Chall_C \<notin> range CardSecret \<union> range PANSecret;
+        Nonce Chall_M \<notin> used []; Chall_M \<notin> range CardSecret \<union> range PANSecret;
+        Chall_C < Chall_M; 
+        Number LID_M \<notin> used []; LID_M \<notin> range CardSecret \<union> range PANSecret;
+        Number XID \<notin> used []; XID \<notin> range CardSecret \<union> range PANSecret;
+        LID_M < XID; XID < OrderDesc; OrderDesc < PurchAmt |] 
+   ==>  \<exists>evs \<in> set_pur.
+            Says M C
+                 (sign (priSK M) {|Number LID_M, Number XID, Nonce Chall_C, 
+                                   Hash (Number PurchAmt)|})
+               \<in> set evs"
+apply (intro exI bexI)
+apply (rule_tac [2]
+        set_pur.Nil
+         [THEN set_pur.Start [of _ LID_M C k M i _ _ _ OrderDesc PurchAmt], 
+          THEN set_pur.PInitReq [of concl: C M LID_M Chall_C],
+          THEN Says_to_Gets, 
+          THEN set_pur.PInitRes [of concl: M C LID_M XID Chall_C Chall_M], 
+          THEN Says_to_Gets,
+          THEN set_pur.PReqS [of concl: C M _ _ KC],
+          THEN Says_to_Gets, 
+          THEN set_pur.AuthReq [of concl: M "PG j" KM LID_M XID], 
+          THEN Says_to_Gets, 
+          THEN set_pur.AuthResS [of concl: "PG j" M KP LID_M XID],
+          THEN Says_to_Gets, 
+          THEN set_pur.PRes]) 
+apply basic_possibility
+apply (auto simp add: used_Cons symKeys_neq_imp_neq) 
+done
+
+text{*General facts about message reception*}
+lemma Gets_imp_Says:
+     "[| Gets B X \<in> set evs; evs \<in> set_pur |]
+   ==> \<exists>A. Says A B X \<in> set evs"
+apply (erule rev_mp)
+apply (erule set_pur.induct, auto)
+done
+
+lemma Gets_imp_knows_Spy:
+     "[| Gets B X \<in> set evs; evs \<in> set_pur |]  ==> X \<in> knows Spy evs"
+by (blast dest!: Gets_imp_Says Says_imp_knows_Spy)
+
+declare Gets_imp_knows_Spy [THEN parts.Inj, dest]
+
+text{*Forwarding lemmas, to aid simplification*}
+
+lemma AuthReq_msg_in_parts_spies:
+     "[|Gets M {|P_I, OIData, HPIData|} \<in> set evs;
+        evs \<in> set_pur|] ==> P_I \<in> parts (knows Spy evs)"
+by auto
+
+lemma AuthReq_msg_in_analz_spies:
+     "[|Gets M {|P_I, OIData, HPIData|} \<in> set evs;
+        evs \<in> set_pur|] ==> P_I \<in> analz (knows Spy evs)"
+by (blast dest: Gets_imp_knows_Spy [THEN analz.Inj])
+
+
+subsection{*Proofs on Asymmetric Keys*}
+
+text{*Private Keys are Secret*}
+
+text{*Spy never sees an agent's private keys! (unless it's bad at start)*}
+lemma Spy_see_private_Key [simp]:
+     "evs \<in> set_pur
+      ==> (Key(invKey (publicKey b A)) \<in> parts(knows Spy evs)) = (A \<in> bad)"
+apply (erule set_pur.induct)
+apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
+apply auto
+done
+declare Spy_see_private_Key [THEN [2] rev_iffD1, dest!]
+
+lemma Spy_analz_private_Key [simp]:
+     "evs \<in> set_pur ==>
+     (Key(invKey (publicKey b A)) \<in> analz(knows Spy evs)) = (A \<in> bad)"
+by auto
+declare Spy_analz_private_Key [THEN [2] rev_iffD1, dest!]
+
+text{*rewriting rule for priEK's*}
+lemma parts_image_priEK:
+     "[|Key (priEK C) \<in> parts (Key`KK Un (knows Spy evs));
+        evs \<in> set_pur|] ==> priEK C \<in> KK | C \<in> bad"
+by auto
+
+text{*trivial proof because @{term"priEK C"} never appears even in
+  @{term "parts evs"}. *}
+lemma analz_image_priEK:
+     "evs \<in> set_pur ==>
+          (Key (priEK C) \<in> analz (Key`KK Un (knows Spy evs))) =
+          (priEK C \<in> KK | C \<in> bad)"
+by (blast dest!: parts_image_priEK intro: analz_mono [THEN [2] rev_subsetD])
+
+
+subsection{*Public Keys in Certificates are Correct*}
+
+lemma Crypt_valid_pubEK [dest!]:
+     "[| Crypt (priSK RCA) {|Agent C, Key EKi, onlyEnc|}
+           \<in> parts (knows Spy evs);
+         evs \<in> set_pur |] ==> EKi = pubEK C"
+by (erule rev_mp, erule set_pur.induct, auto)
+
+lemma Crypt_valid_pubSK [dest!]:
+     "[| Crypt (priSK RCA) {|Agent C, Key SKi, onlySig|}
+           \<in> parts (knows Spy evs);
+         evs \<in> set_pur |] ==> SKi = pubSK C"
+by (erule rev_mp, erule set_pur.induct, auto)
+
+lemma certificate_valid_pubEK:
+    "[| cert C EKi onlyEnc (priSK RCA) \<in> parts (knows Spy evs);
+        evs \<in> set_pur |]
+     ==> EKi = pubEK C"
+by (unfold cert_def signCert_def, auto)
+
+lemma certificate_valid_pubSK:
+    "[| cert C SKi onlySig (priSK RCA) \<in> parts (knows Spy evs);
+        evs \<in> set_pur |] ==> SKi = pubSK C"
+by (unfold cert_def signCert_def, auto)
+
+lemma Says_certificate_valid [simp]:
+     "[| Says A B (sign SK {|lid, xid, cc, cm,
+                           cert C EK onlyEnc (priSK RCA)|}) \<in> set evs;
+         evs \<in> set_pur |]
+      ==> EK = pubEK C"
+by (unfold sign_def, auto)
+
+lemma Gets_certificate_valid [simp]:
+     "[| Gets A (sign SK {|lid, xid, cc, cm,
+                           cert C EK onlyEnc (priSK RCA)|}) \<in> set evs;
+         evs \<in> set_pur |]
+      ==> EK = pubEK C"
+by (frule Gets_imp_Says, auto)
+
+method_setup valid_certificate_tac = {*
+  Args.goal_spec >> (fn quant =>
+    K (SIMPLE_METHOD'' quant (fn i =>
+      EVERY [ftac @{thm Gets_certificate_valid} i,
+             assume_tac i, REPEAT (hyp_subst_tac i)])))
+*} ""
+
+
+subsection{*Proofs on Symmetric Keys*}
+
+text{*Nobody can have used non-existent keys!*}
+lemma new_keys_not_used [rule_format,simp]:
+     "evs \<in> set_pur
+      ==> Key K \<notin> used evs --> K \<in> symKeys -->
+          K \<notin> keysFor (parts (knows Spy evs))"
+apply (erule set_pur.induct)
+apply (valid_certificate_tac [8]) --{*PReqS*}
+apply (valid_certificate_tac [7]) --{*PReqUns*}
+apply auto
+apply (force dest!: usedI keysFor_parts_insert) --{*Fake*}
+done
+
+lemma new_keys_not_analzd:
+     "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_pur |]
+      ==> K \<notin> keysFor (analz (knows Spy evs))"
+by (blast intro: keysFor_mono [THEN [2] rev_subsetD] dest: new_keys_not_used)
+
+lemma Crypt_parts_imp_used:
+     "[|Crypt K X \<in> parts (knows Spy evs);
+        K \<in> symKeys; evs \<in> set_pur |] ==> Key K \<in> used evs"
+apply (rule ccontr)
+apply (force dest: new_keys_not_used Crypt_imp_invKey_keysFor)
+done
+
+lemma Crypt_analz_imp_used:
+     "[|Crypt K X \<in> analz (knows Spy evs);
+        K \<in> symKeys; evs \<in> set_pur |] ==> Key K \<in> used evs"
+by (blast intro: Crypt_parts_imp_used)
+
+text{*New versions: as above, but generalized to have the KK argument*}
+
+lemma gen_new_keys_not_used:
+     "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_pur |]
+      ==> Key K \<notin> used evs --> K \<in> symKeys -->
+          K \<notin> keysFor (parts (Key`KK Un knows Spy evs))"
+by auto
+
+lemma gen_new_keys_not_analzd:
+     "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_pur |]
+      ==> K \<notin> keysFor (analz (Key`KK Un knows Spy evs))"
+by (blast intro: keysFor_mono [THEN subsetD] dest: gen_new_keys_not_used)
+
+lemma analz_Key_image_insert_eq:
+     "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_pur |]
+      ==> analz (Key ` (insert K KK) \<union> knows Spy evs) =
+          insert (Key K) (analz (Key ` KK \<union> knows Spy evs))"
+by (simp add: gen_new_keys_not_analzd)
+
+
+subsection{*Secrecy of Symmetric Keys*}
+
+lemma Key_analz_image_Key_lemma:
+     "P --> (Key K \<in> analz (Key`KK Un H)) --> (K\<in>KK | Key K \<in> analz H)
+      ==>
+      P --> (Key K \<in> analz (Key`KK Un H)) = (K\<in>KK | Key K \<in> analz H)"
+by (blast intro: analz_mono [THEN [2] rev_subsetD])
+
+
+lemma symKey_compromise:
+     "evs \<in> set_pur \<Longrightarrow>
+      (\<forall>SK KK. SK \<in> symKeys \<longrightarrow>
+        (\<forall>K \<in> KK. K \<notin> range(\<lambda>C. priEK C)) \<longrightarrow>
+               (Key SK \<in> analz (Key`KK \<union> (knows Spy evs))) =
+               (SK \<in> KK \<or> Key SK \<in> analz (knows Spy evs)))"
+apply (erule set_pur.induct)
+apply (rule_tac [!] allI)+
+apply (rule_tac [!] impI [THEN Key_analz_image_Key_lemma, THEN impI])+
+apply (frule_tac [9] AuthReq_msg_in_analz_spies) --{*AReq*}
+apply (valid_certificate_tac [8]) --{*PReqS*}
+apply (valid_certificate_tac [7]) --{*PReqUns*}
+apply (simp_all
+         del: image_insert image_Un imp_disjL
+         add: analz_image_keys_simps disj_simps
+              analz_Key_image_insert_eq notin_image_iff
+              analz_insert_simps analz_image_priEK)
+  --{*8 seconds on a 1.6GHz machine*}
+apply spy_analz --{*Fake*}
+apply (blast elim!: ballE)+ --{*PReq: unsigned and signed*}
+done
+
+
+
+subsection{*Secrecy of Nonces*}
+
+text{*As usual: we express the property as a logical equivalence*}
+lemma Nonce_analz_image_Key_lemma:
+     "P --> (Nonce N \<in> analz (Key`KK Un H)) --> (Nonce N \<in> analz H)
+      ==> P --> (Nonce N \<in> analz (Key`KK Un H)) = (Nonce N \<in> analz H)"
+by (blast intro: analz_mono [THEN [2] rev_subsetD])
+
+text{*The @{text "(no_asm)"} attribute is essential, since it retains
+  the quantifier and allows the simprule's condition to itself be simplified.*}
+lemma Nonce_compromise [rule_format (no_asm)]:
+     "evs \<in> set_pur ==>
+      (\<forall>N KK. (\<forall>K \<in> KK. K \<notin> range(\<lambda>C. priEK C))   -->
+              (Nonce N \<in> analz (Key`KK \<union> (knows Spy evs))) =
+              (Nonce N \<in> analz (knows Spy evs)))"
+apply (erule set_pur.induct)
+apply (rule_tac [!] allI)+
+apply (rule_tac [!] impI [THEN Nonce_analz_image_Key_lemma])+
+apply (frule_tac [9] AuthReq_msg_in_analz_spies) --{*AReq*}
+apply (valid_certificate_tac [8]) --{*PReqS*}
+apply (valid_certificate_tac [7]) --{*PReqUns*}
+apply (simp_all
+         del: image_insert image_Un imp_disjL
+         add: analz_image_keys_simps disj_simps symKey_compromise
+              analz_Key_image_insert_eq notin_image_iff
+              analz_insert_simps analz_image_priEK)
+  --{*8 seconds on a 1.6GHz machine*}
+apply spy_analz --{*Fake*}
+apply (blast elim!: ballE) --{*PReqS*}
+done
+
+lemma PANSecret_notin_spies:
+     "[|Nonce (PANSecret k) \<in> analz (knows Spy evs);  evs \<in> set_pur|]
+      ==> 
+       (\<exists>V W X Y KC2 M. \<exists>P \<in> bad.
+          Says (Cardholder k) M
+               {|{|W, EXcrypt KC2 (pubEK P) X {|Y, Nonce (PANSecret k)|}|},
+                 V|}  \<in>  set evs)"
+apply (erule rev_mp)
+apply (erule set_pur.induct)
+apply (frule_tac [9] AuthReq_msg_in_analz_spies)
+apply (valid_certificate_tac [8]) --{*PReqS*}
+apply (simp_all
+         del: image_insert image_Un imp_disjL
+         add: analz_image_keys_simps disj_simps
+              symKey_compromise pushes sign_def Nonce_compromise
+              analz_Key_image_insert_eq notin_image_iff
+              analz_insert_simps analz_image_priEK)
+  --{*2.5 seconds on a 1.6GHz machine*}
+apply spy_analz
+apply (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj])
+apply (blast dest: Says_imp_knows_Spy [THEN analz.Inj] 
+                   Gets_imp_knows_Spy [THEN analz.Inj])
+apply (blast dest: Gets_imp_knows_Spy [THEN analz.Inj]) --{*PReqS*}
+apply (blast dest: Says_imp_knows_Spy [THEN analz.Inj] 
+                   Gets_imp_knows_Spy [THEN analz.Inj]) --{*PRes*}
+done
+
+text{*This theorem is a bit silly, in that many CardSecrets are 0!
+  But then we don't care.  NOT USED*}
+lemma CardSecret_notin_spies:
+     "evs \<in> set_pur ==> Nonce (CardSecret i) \<notin> parts (knows Spy evs)"
+by (erule set_pur.induct, auto)
+
+
+subsection{*Confidentiality of PAN*}
+
+lemma analz_image_pan_lemma:
+     "(Pan P \<in> analz (Key`nE Un H)) --> (Pan P \<in> analz H)  ==>
+      (Pan P \<in> analz (Key`nE Un H)) =   (Pan P \<in> analz H)"
+by (blast intro: analz_mono [THEN [2] rev_subsetD])
+
+text{*The @{text "(no_asm)"} attribute is essential, since it retains
+  the quantifier and allows the simprule's condition to itself be simplified.*}
+lemma analz_image_pan [rule_format (no_asm)]:
+     "evs \<in> set_pur ==>
+       \<forall>KK. (\<forall>K \<in> KK. K \<notin> range(\<lambda>C. priEK C)) -->
+            (Pan P \<in> analz (Key`KK Un (knows Spy evs))) =
+            (Pan P \<in> analz (knows Spy evs))"
+apply (erule set_pur.induct)
+apply (rule_tac [!] allI impI)+
+apply (rule_tac [!] analz_image_pan_lemma)+
+apply (frule_tac [9] AuthReq_msg_in_analz_spies) --{*AReq*}
+apply (valid_certificate_tac [8]) --{*PReqS*}
+apply (valid_certificate_tac [7]) --{*PReqUns*}
+apply (simp_all
+         del: image_insert image_Un imp_disjL
+         add: analz_image_keys_simps
+              symKey_compromise pushes sign_def
+              analz_Key_image_insert_eq notin_image_iff
+              analz_insert_simps analz_image_priEK)
+  --{*7 seconds on a 1.6GHz machine*}
+apply spy_analz --{*Fake*}
+apply auto
+done
+
+lemma analz_insert_pan:
+     "[| evs \<in> set_pur;  K \<notin> range(\<lambda>C. priEK C) |] ==>
+          (Pan P \<in> analz (insert (Key K) (knows Spy evs))) =
+          (Pan P \<in> analz (knows Spy evs))"
+by (simp del: image_insert image_Un
+         add: analz_image_keys_simps analz_image_pan)
+
+text{*Confidentiality of the PAN, unsigned case.*}
+theorem pan_confidentiality_unsigned:
+     "[| Pan (pan C) \<in> analz(knows Spy evs);  C = Cardholder k;
+         CardSecret k = 0;  evs \<in> set_pur|]
+    ==> \<exists>P M KC1 K X Y.
+     Says C M {|EXHcrypt KC1 (pubEK P) X (Pan (pan C)), Y|}
+          \<in> set evs  &
+     P \<in> bad"
+apply (erule rev_mp)
+apply (erule set_pur.induct)
+apply (frule_tac [9] AuthReq_msg_in_analz_spies) --{*AReq*}
+apply (valid_certificate_tac [8]) --{*PReqS*}
+apply (valid_certificate_tac [7]) --{*PReqUns*}
+apply (simp_all
+         del: image_insert image_Un imp_disjL
+         add: analz_image_keys_simps analz_insert_pan analz_image_pan
+              notin_image_iff
+              analz_insert_simps analz_image_priEK)
+  --{*3 seconds on a 1.6GHz machine*}
+apply spy_analz --{*Fake*}
+apply blast --{*PReqUns: unsigned*}
+apply force --{*PReqS: signed*}
+done
+
+text{*Confidentiality of the PAN, signed case.*}
+theorem pan_confidentiality_signed:
+ "[|Pan (pan C) \<in> analz(knows Spy evs);  C = Cardholder k;
+    CardSecret k \<noteq> 0;  evs \<in> set_pur|]
+  ==> \<exists>P M KC2 PIDualSign_1 PIDualSign_2 other OIDualSign.
+      Says C M {|{|PIDualSign_1, 
+                   EXcrypt KC2 (pubEK P) PIDualSign_2 {|Pan (pan C), other|}|}, 
+       OIDualSign|} \<in> set evs  &  P \<in> bad"
+apply (erule rev_mp)
+apply (erule set_pur.induct)
+apply (frule_tac [9] AuthReq_msg_in_analz_spies) --{*AReq*}
+apply (valid_certificate_tac [8]) --{*PReqS*}
+apply (valid_certificate_tac [7]) --{*PReqUns*}
+apply (simp_all
+         del: image_insert image_Un imp_disjL
+         add: analz_image_keys_simps analz_insert_pan analz_image_pan
+              notin_image_iff
+              analz_insert_simps analz_image_priEK)
+  --{*3 seconds on a 1.6GHz machine*}
+apply spy_analz --{*Fake*}
+apply force --{*PReqUns: unsigned*}
+apply blast --{*PReqS: signed*}
+done
+
+text{*General goal: that C, M and PG agree on those details of the transaction
+     that they are allowed to know about.  PG knows about price and account
+     details.  M knows about the order description and price.  C knows both.*}
+
+
+subsection{*Proofs Common to Signed and Unsigned Versions*}
+
+lemma M_Notes_PG:
+     "[|Notes M {|Number LID_M, Agent P, Agent M, Agent C, etc|} \<in> set evs;
+        evs \<in> set_pur|] ==> \<exists>j. P = PG j"
+by (erule rev_mp, erule set_pur.induct, simp_all)
+
+text{*If we trust M, then @{term LID_M} determines his choice of P
+      (Payment Gateway)*}
+lemma goodM_gives_correct_PG:
+     "[| MsgPInitRes = 
+            {|Number LID_M, xid, cc, cm, cert P EKj onlyEnc (priSK RCA)|};
+         Crypt (priSK M) (Hash MsgPInitRes) \<in> parts (knows Spy evs);
+         evs \<in> set_pur; M \<notin> bad |]
+      ==> \<exists>j trans.
+            P = PG j &
+            Notes M {|Number LID_M, Agent P, trans|} \<in> set evs"
+apply clarify
+apply (erule rev_mp)
+apply (erule set_pur.induct)
+apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
+apply simp_all
+apply (blast intro: M_Notes_PG)+
+done
+
+lemma C_gets_correct_PG:
+     "[| Gets A (sign (priSK M) {|Number LID_M, xid, cc, cm,
+                              cert P EKj onlyEnc (priSK RCA)|}) \<in> set evs;
+         evs \<in> set_pur;  M \<notin> bad|]
+      ==> \<exists>j trans.
+            P = PG j &
+            Notes M {|Number LID_M, Agent P, trans|} \<in> set evs &
+            EKj = pubEK P"
+by (rule refl [THEN goodM_gives_correct_PG, THEN exE], auto)
+
+text{*When C receives PInitRes, he learns M's choice of P*}
+lemma C_verifies_PInitRes:
+ "[| MsgPInitRes = {|Number LID_M, Number XID, Nonce Chall_C, Nonce Chall_M,
+           cert P EKj onlyEnc (priSK RCA)|};
+     Crypt (priSK M) (Hash MsgPInitRes) \<in> parts (knows Spy evs);
+     evs \<in> set_pur;  M \<notin> bad|]
+  ==> \<exists>j trans.
+         Notes M {|Number LID_M, Agent P, trans|} \<in> set evs &
+         P = PG j &
+         EKj = pubEK P"
+apply clarify
+apply (erule rev_mp)
+apply (erule set_pur.induct)
+apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
+apply simp_all
+apply (blast intro: M_Notes_PG)+
+done
+
+text{*Corollary of previous one*}
+lemma Says_C_PInitRes:
+     "[|Says A C (sign (priSK M)
+                      {|Number LID_M, Number XID,
+                        Nonce Chall_C, Nonce Chall_M,
+                        cert P EKj onlyEnc (priSK RCA)|})
+           \<in> set evs;  M \<notin> bad;  evs \<in> set_pur|]
+      ==> \<exists>j trans.
+           Notes M {|Number LID_M, Agent P, trans|} \<in> set evs &
+           P = PG j &
+           EKj = pubEK (PG j)"
+apply (frule Says_certificate_valid)
+apply (auto simp add: sign_def)
+apply (blast dest: refl [THEN goodM_gives_correct_PG])
+apply (blast dest: refl [THEN C_verifies_PInitRes])
+done
+
+text{*When P receives an AuthReq, he knows that the signed part originated 
+      with M. PIRes also has a signed message from M....*}
+lemma P_verifies_AuthReq:
+     "[| AuthReqData = {|Number LID_M, Number XID, HOIData, HOD|};
+         Crypt (priSK M) (Hash {|AuthReqData, Hash P_I|})
+           \<in> parts (knows Spy evs);
+         evs \<in> set_pur;  M \<notin> bad|]
+      ==> \<exists>j trans KM OIData HPIData.
+            Notes M {|Number LID_M, Agent (PG j), trans|} \<in> set evs &
+            Gets M {|P_I, OIData, HPIData|} \<in> set evs &
+            Says M (PG j) (EncB (priSK M) KM (pubEK (PG j)) AuthReqData P_I)
+              \<in> set evs"
+apply clarify
+apply (erule rev_mp)
+apply (erule set_pur.induct, simp_all)
+apply (frule_tac [4] M_Notes_PG, auto)
+done
+
+text{*When M receives AuthRes, he knows that P signed it, including
+  the identifying tags and the purchase amount, which he can verify.
+  (Although the spec has SIGNED and UNSIGNED forms of AuthRes, they
+   send the same message to M.)  The conclusion is weak: M is existentially
+  quantified! That is because Authorization Response does not refer to M, while
+  the digital envelope weakens the link between @{term MsgAuthRes} and
+  @{term"priSK M"}.  Changing the precondition to refer to 
+  @{term "Crypt K (sign SK M)"} requires assuming @{term K} to be secure, since
+  otherwise the Spy could create that message.*}
+theorem M_verifies_AuthRes:
+  "[| MsgAuthRes = {|{|Number LID_M, Number XID, Number PurchAmt|}, 
+                     Hash authCode|};
+      Crypt (priSK (PG j)) (Hash MsgAuthRes) \<in> parts (knows Spy evs);
+      PG j \<notin> bad;  evs \<in> set_pur|]
+   ==> \<exists>M KM KP HOIData HOD P_I.
+        Gets (PG j)
+           (EncB (priSK M) KM (pubEK (PG j))
+                    {|Number LID_M, Number XID, HOIData, HOD|}
+                    P_I) \<in> set evs &
+        Says (PG j) M
+             (EncB (priSK (PG j)) KP (pubEK M)
+              {|Number LID_M, Number XID, Number PurchAmt|}
+              authCode) \<in> set evs"
+apply clarify
+apply (erule rev_mp)
+apply (erule set_pur.induct)
+apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
+apply simp_all
+apply blast+
+done
+
+
+subsection{*Proofs for Unsigned Purchases*}
+
+text{*What we can derive from the ASSUMPTION that C issued a purchase request.
+   In the unsigned case, we must trust "C": there's no authentication.*}
+lemma C_determines_EKj:
+     "[| Says C M {|EXHcrypt KC1 EKj {|PIHead, Hash OIData|} (Pan (pan C)),
+                    OIData, Hash{|PIHead, Pan (pan C)|} |} \<in> set evs;
+         PIHead = {|Number LID_M, Trans_details|};
+         evs \<in> set_pur;  C = Cardholder k;  M \<notin> bad|]
+  ==> \<exists>trans j.
+               Notes M {|Number LID_M, Agent (PG j), trans |} \<in> set evs &
+               EKj = pubEK (PG j)"
+apply clarify
+apply (erule rev_mp)
+apply (erule set_pur.induct, simp_all)
+apply (valid_certificate_tac [2]) --{*PReqUns*}
+apply auto
+apply (blast dest: Gets_imp_Says Says_C_PInitRes)
+done
+
+
+text{*Unicity of @{term LID_M} between Merchant and Cardholder notes*}
+lemma unique_LID_M:
+     "[|Notes (Merchant i) {|Number LID_M, Agent P, Trans|} \<in> set evs;
+        Notes C {|Number LID_M, Agent M, Agent C, Number OD,
+             Number PA|} \<in> set evs;
+        evs \<in> set_pur|]
+      ==> M = Merchant i & Trans = {|Agent M, Agent C, Number OD, Number PA|}"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule set_pur.induct, simp_all)
+apply (force dest!: Notes_imp_parts_subset_used)
+done
+
+text{*Unicity of @{term LID_M}, for two Merchant Notes events*}
+lemma unique_LID_M2:
+     "[|Notes M {|Number LID_M, Trans|} \<in> set evs;
+        Notes M {|Number LID_M, Trans'|} \<in> set evs;
+        evs \<in> set_pur|] ==> Trans' = Trans"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule set_pur.induct, simp_all)
+apply (force dest!: Notes_imp_parts_subset_used)
+done
+
+text{*Lemma needed below: for the case that
+  if PRes is present, then @{term LID_M} has been used.*}
+lemma signed_imp_used:
+     "[| Crypt (priSK M) (Hash X) \<in> parts (knows Spy evs);
+         M \<notin> bad;  evs \<in> set_pur|] ==> parts {X} \<subseteq> used evs"
+apply (erule rev_mp)
+apply (erule set_pur.induct)
+apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
+apply simp_all
+apply safe
+apply blast+
+done
+
+text{*Similar, with nested Hash*}
+lemma signed_Hash_imp_used:
+     "[| Crypt (priSK C) (Hash {|H, Hash X|}) \<in> parts (knows Spy evs);
+         C \<notin> bad;  evs \<in> set_pur|] ==> parts {X} \<subseteq> used evs"
+apply (erule rev_mp)
+apply (erule set_pur.induct)
+apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
+apply simp_all
+apply safe
+apply blast+
+done
+
+text{*Lemma needed below: for the case that
+  if PRes is present, then @{text LID_M} has been used.*}
+lemma PRes_imp_LID_used:
+     "[| Crypt (priSK M) (Hash {|N, X|}) \<in> parts (knows Spy evs);
+         M \<notin> bad;  evs \<in> set_pur|] ==> N \<in> used evs"
+by (drule signed_imp_used, auto)
+
+text{*When C receives PRes, he knows that M and P agreed to the purchase details.
+  He also knows that P is the same PG as before*}
+lemma C_verifies_PRes_lemma:
+     "[| Crypt (priSK M) (Hash MsgPRes) \<in> parts (knows Spy evs);
+         Notes C {|Number LID_M, Trans |} \<in> set evs;
+         Trans = {| Agent M, Agent C, Number OrderDesc, Number PurchAmt |};
+         MsgPRes = {|Number LID_M, Number XID, Nonce Chall_C,
+                Hash (Number PurchAmt)|};
+         evs \<in> set_pur;  M \<notin> bad|]
+  ==> \<exists>j KP.
+        Notes M {|Number LID_M, Agent (PG j), Trans |}
+          \<in> set evs &
+        Gets M (EncB (priSK (PG j)) KP (pubEK M)
+                {|Number LID_M, Number XID, Number PurchAmt|}
+                authCode)
+          \<in> set evs &
+        Says M C (sign (priSK M) MsgPRes) \<in> set evs"
+apply clarify
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule set_pur.induct)
+apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
+apply simp_all
+apply blast
+apply blast
+apply (blast dest: PRes_imp_LID_used)
+apply (frule M_Notes_PG, auto)
+apply (blast dest: unique_LID_M)
+done
+
+text{*When the Cardholder receives Purchase Response from an uncompromised
+Merchant, he knows that M sent it. He also knows that M received a message signed
+by a Payment Gateway chosen by M to authorize the purchase.*}
+theorem C_verifies_PRes:
+     "[| MsgPRes = {|Number LID_M, Number XID, Nonce Chall_C,
+                     Hash (Number PurchAmt)|};
+         Gets C (sign (priSK M) MsgPRes) \<in> set evs;
+         Notes C {|Number LID_M, Agent M, Agent C, Number OrderDesc,
+                   Number PurchAmt|} \<in> set evs;
+         evs \<in> set_pur;  M \<notin> bad|]
+  ==> \<exists>P KP trans.
+        Notes M {|Number LID_M,Agent P, trans|} \<in> set evs &
+        Gets M (EncB (priSK P) KP (pubEK M)
+                {|Number LID_M, Number XID, Number PurchAmt|}
+                authCode)  \<in>  set evs &
+        Says M C (sign (priSK M) MsgPRes) \<in> set evs"
+apply (rule C_verifies_PRes_lemma [THEN exE])
+apply (auto simp add: sign_def)
+done
+
+subsection{*Proofs for Signed Purchases*}
+
+text{*Some Useful Lemmas: the cardholder knows what he is doing*}
+
+lemma Crypt_imp_Says_Cardholder:
+     "[| Crypt K {|{|{|Number LID_M, others|}, Hash OIData|}, Hash PANData|}
+           \<in> parts (knows Spy evs);
+         PANData = {|Pan (pan (Cardholder k)), Nonce (PANSecret k)|};
+         Key K \<notin> analz (knows Spy evs);
+         evs \<in> set_pur|]
+  ==> \<exists>M shash EK HPIData.
+       Says (Cardholder k) M {|{|shash,
+          Crypt K
+            {|{|{|Number LID_M, others|}, Hash OIData|}, Hash PANData|},
+           Crypt EK {|Key K, PANData|}|},
+          OIData, HPIData|} \<in> set evs"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule set_pur.induct, analz_mono_contra)
+apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
+apply simp_all
+apply auto
+done
+
+lemma Says_PReqS_imp_trans_details_C:
+     "[| MsgPReqS = {|{|shash,
+                 Crypt K
+                  {|{|{|Number LID_M, PIrest|}, Hash OIData|}, hashpd|},
+            cryptek|}, data|};
+         Says (Cardholder k) M MsgPReqS \<in> set evs;
+         evs \<in> set_pur |]
+   ==> \<exists>trans.
+           Notes (Cardholder k) 
+                 {|Number LID_M, Agent M, Agent (Cardholder k), trans|}
+            \<in> set evs"
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule set_pur.induct)
+apply (simp_all (no_asm_simp))
+apply auto
+done
+
+text{*Can't happen: only Merchants create this type of Note*}
+lemma Notes_Cardholder_self_False:
+     "[|Notes (Cardholder k)
+          {|Number n, Agent P, Agent (Cardholder k), Agent C, etc|} \<in> set evs;
+        evs \<in> set_pur|] ==> False"
+by (erule rev_mp, erule set_pur.induct, auto)
+
+text{*When M sees a dual signature, he knows that it originated with C.
+  Using XID he knows it was intended for him.
+  This guarantee isn't useful to P, who never gets OIData.*}
+theorem M_verifies_Signed_PReq:
+ "[| MsgDualSign = {|HPIData, Hash OIData|};
+     OIData = {|Number LID_M, etc|};
+     Crypt (priSK C) (Hash MsgDualSign) \<in> parts (knows Spy evs);
+     Notes M {|Number LID_M, Agent P, extras|} \<in> set evs;
+     M = Merchant i;  C = Cardholder k;  C \<notin> bad;  evs \<in> set_pur|]
+  ==> \<exists>PIData PICrypt.
+        HPIData = Hash PIData &
+        Says C M {|{|sign (priSK C) MsgDualSign, PICrypt|}, OIData, Hash PIData|}
+          \<in> set evs"
+apply clarify
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule set_pur.induct)
+apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
+apply simp_all
+apply blast
+apply (metis subsetD insert_subset parts.Fst parts_increasing signed_Hash_imp_used)
+apply (metis unique_LID_M)
+apply (blast dest!: Notes_Cardholder_self_False)
+done
+
+text{*When P sees a dual signature, he knows that it originated with C.
+  and was intended for M. This guarantee isn't useful to M, who never gets
+  PIData. I don't see how to link @{term "PG j"} and @{text LID_M} without
+  assuming @{term "M \<notin> bad"}.*}
+theorem P_verifies_Signed_PReq:
+     "[| MsgDualSign = {|Hash PIData, HOIData|};
+         PIData = {|PIHead, PANData|};
+         PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M,
+                    TransStain|};
+         Crypt (priSK C) (Hash MsgDualSign) \<in> parts (knows Spy evs);
+         evs \<in> set_pur;  C \<notin> bad;  M \<notin> bad|]
+    ==> \<exists>OIData OrderDesc K j trans.
+          HOD = Hash{|Number OrderDesc, Number PurchAmt|} &
+          HOIData = Hash OIData &
+          Notes M {|Number LID_M, Agent (PG j), trans|} \<in> set evs &
+          Says C M {|{|sign (priSK C) MsgDualSign,
+                     EXcrypt K (pubEK (PG j))
+                                {|PIHead, Hash OIData|} PANData|},
+                     OIData, Hash PIData|}
+            \<in> set evs"
+apply clarify
+apply (erule rev_mp)
+apply (erule set_pur.induct, simp_all)
+apply (auto dest!: C_gets_correct_PG)
+done
+
+lemma C_determines_EKj_signed:
+     "[| Says C M {|{|sign (priSK C) text,
+                      EXcrypt K EKj {|PIHead, X|} Y|}, Z|} \<in> set evs;
+         PIHead = {|Number LID_M, Number XID, W|};
+         C = Cardholder k;  evs \<in> set_pur;  M \<notin> bad|]
+  ==> \<exists> trans j.
+         Notes M {|Number LID_M, Agent (PG j), trans|} \<in> set evs &
+         EKj = pubEK (PG j)"
+apply clarify
+apply (erule rev_mp)
+apply (erule set_pur.induct, simp_all, auto)
+apply (blast dest: C_gets_correct_PG)
+done
+
+lemma M_Says_AuthReq:
+     "[| AuthReqData = {|Number LID_M, Number XID, HOIData, HOD|};
+         sign (priSK M) {|AuthReqData, Hash P_I|} \<in> parts (knows Spy evs);
+         evs \<in> set_pur;  M \<notin> bad|]
+   ==> \<exists>j trans KM.
+           Notes M {|Number LID_M, Agent (PG j), trans |} \<in> set evs &
+             Says M (PG j)
+               (EncB (priSK M) KM (pubEK (PG j)) AuthReqData P_I)
+              \<in> set evs"
+apply (rule refl [THEN P_verifies_AuthReq, THEN exE])
+apply (auto simp add: sign_def)
+done
+
+text{*A variant of @{text M_verifies_Signed_PReq} with explicit PI information.
+  Even here we cannot be certain about what C sent to M, since a bad
+  PG could have replaced the two key fields.  (NOT USED)*}
+lemma Signed_PReq_imp_Says_Cardholder:
+     "[| MsgDualSign = {|Hash PIData, Hash OIData|};
+         OIData = {|Number LID_M, Number XID, Nonce Chall_C, HOD, etc|};
+         PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M,
+                    TransStain|};
+         PIData = {|PIHead, PANData|};
+         Crypt (priSK C) (Hash MsgDualSign) \<in> parts (knows Spy evs);
+         M = Merchant i;  C = Cardholder k;  C \<notin> bad;  evs \<in> set_pur|]
+      ==> \<exists>KC EKj.
+            Says C M {|{|sign (priSK C) MsgDualSign,
+                       EXcrypt KC EKj {|PIHead, Hash OIData|} PANData|},
+                       OIData, Hash PIData|}
+              \<in> set evs"
+apply clarify
+apply (erule rev_mp)
+apply (erule rev_mp)
+apply (erule set_pur.induct, simp_all, auto)
+done
+
+text{*When P receives an AuthReq and a dual signature, he knows that C and M
+  agree on the essential details.  PurchAmt however is never sent by M to
+  P; instead C and M both send 
+     @{term "HOD = Hash{|Number OrderDesc, Number PurchAmt|}"}
+  and P compares the two copies of HOD.
+
+  Agreement can't be proved for some things, including the symmetric keys
+  used in the digital envelopes.  On the other hand, M knows the true identity
+  of PG (namely j'), and sends AReq there; he can't, however, check that
+  the EXcrypt involves the correct PG's key.
+*}
+theorem P_sees_CM_agreement:
+     "[| AuthReqData = {|Number LID_M, Number XID, HOIData, HOD|};
+         KC \<in> symKeys;
+         Gets (PG j) (EncB (priSK M) KM (pubEK (PG j)) AuthReqData P_I)
+           \<in> set evs;
+         C = Cardholder k;
+         PI_sign = sign (priSK C) {|Hash PIData, HOIData|};
+         P_I = {|PI_sign,
+                 EXcrypt KC (pubEK (PG j)) {|PIHead, HOIData|} PANData|};
+         PANData = {|Pan (pan C), Nonce (PANSecret k)|};
+         PIData = {|PIHead, PANData|};
+         PIHead = {|Number LID_M, Number XID, HOD, Number PurchAmt, Agent M,
+                    TransStain|};
+         evs \<in> set_pur;  C \<notin> bad;  M \<notin> bad|]
+  ==> \<exists>OIData OrderDesc KM' trans j' KC' KC'' P_I' P_I''.
+           HOD = Hash{|Number OrderDesc, Number PurchAmt|} &
+           HOIData = Hash OIData &
+           Notes M {|Number LID_M, Agent (PG j'), trans|} \<in> set evs &
+           Says C M {|P_I', OIData, Hash PIData|} \<in> set evs &
+           Says M (PG j') (EncB (priSK M) KM' (pubEK (PG j'))
+                           AuthReqData P_I'')  \<in>  set evs &
+           P_I' = {|PI_sign,
+             EXcrypt KC' (pubEK (PG j')) {|PIHead, Hash OIData|} PANData|} &
+           P_I'' = {|PI_sign,
+             EXcrypt KC'' (pubEK (PG j)) {|PIHead, Hash OIData|} PANData|}"
+apply clarify
+apply (rule exE)
+apply (rule P_verifies_Signed_PReq [OF refl refl refl])
+apply (simp (no_asm_use) add: sign_def EncB_def, blast)
+apply (assumption+, clarify, simp)
+apply (drule Gets_imp_knows_Spy [THEN parts.Inj], assumption)
+apply (blast elim: EncB_partsE dest: refl [THEN M_Says_AuthReq] unique_LID_M2)
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SET_Protocol/ROOT.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,9 @@
+(*  Title:      HOL/SET_Protocol/ROOT.ML
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   2003  University of Cambridge
+
+Root file for the SET protocol proofs.
+*)
+
+no_document use_thy "Nat_Int_Bij";
+use_thys ["Cardholder_Registration", "Merchant_Registration", "Purchase"];
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SET_Protocol/document/root.tex	Wed Oct 21 08:16:25 2009 +0200
@@ -0,0 +1,27 @@
+\documentclass[10pt,a4paper,twoside]{article}
+\usepackage{graphicx}
+\usepackage{latexsym,theorem}
+\usepackage{isabelle,isabellesym}
+\usepackage{pdfsetup}\urlstyle{rm}
+
+\begin{document}
+
+\pagestyle{headings}
+\pagenumbering{arabic}
+
+\title{Verification of The SET Protocol}
+\author{Giampaolo Bella, Fabio Massacci, Lawrence C. Paulson et al.}
+\maketitle
+
+\tableofcontents
+
+\begin{center}
+  \includegraphics[scale=0.5]{session_graph}  
+\end{center}
+
+\newpage
+
+\parindent 0pt\parskip 0.5ex
+
+\input{session}
+\end{document}
--- a/src/HOL/SMT/Tools/cvc3_solver.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/SMT/Tools/cvc3_solver.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -31,7 +31,7 @@
     val p = Pretty.big_list (cex_kind real ^ " found:") (map Pretty.str ls')
   in error (Pretty.string_of p) end
 
-fun core_oracle (SMT_Solver.ProofData {context, output, recon, ...}) =
+fun core_oracle ({context, output, recon, ...} : SMT_Solver.proof_data) =
   let
     val empty_line = (fn "" => true | _ => false)
     val split_first = (fn [] => ("", []) | l :: ls => (l, ls))
@@ -43,12 +43,11 @@
     else error (solver_name ^ " failed")
   end
 
-fun smtlib_solver oracle _ =
-  SMT_Solver.SolverConfig {
-    command = {env_var=env_var, remote_name=solver_name},
-    arguments = options,
-    interface = SMTLIB_Interface.interface,
-    reconstruct = oracle }
+fun smtlib_solver oracle _ = {
+  command = {env_var=env_var, remote_name=solver_name},
+  arguments = options,
+  interface = SMTLIB_Interface.interface,
+  reconstruct = oracle }
 
 val setup =
   Thm.add_oracle (Binding.name solver_name, core_oracle) #-> (fn (_, oracle) =>
--- a/src/HOL/SMT/Tools/smt_solver.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/SMT/Tools/smt_solver.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -8,17 +8,15 @@
 sig
   exception SMT_COUNTEREXAMPLE of bool * term list
 
-  datatype interface = Interface of {
+  type interface = {
     normalize: SMT_Normalize.config list,
     translate: SMT_Translate.config }
-
-  datatype proof_data = ProofData of {
+  type proof_data = {
     context: Proof.context,
     output: string list,
     recon: SMT_Translate.recon,
     assms: thm list option }
-
-  datatype solver_config = SolverConfig of {
+  type solver_config = {
     command: {env_var: string, remote_name: string},
     arguments: string list,
     interface: interface,
@@ -58,17 +56,17 @@
 exception SMT_COUNTEREXAMPLE of bool * term list
 
 
-datatype interface = Interface of {
+type interface = {
   normalize: SMT_Normalize.config list,
   translate: SMT_Translate.config }
 
-datatype proof_data = ProofData of {
+type proof_data = {
   context: Proof.context,
   output: string list,
   recon: SMT_Translate.recon,
   assms: thm list option }
 
-datatype solver_config = SolverConfig of {
+type solver_config = {
   command: {env_var: string, remote_name: string},
   arguments: string list,
   interface: interface,
@@ -144,12 +142,12 @@
 end
 
 fun make_proof_data ctxt ((recon, thms), ls) =
-  ProofData {context=ctxt, output=ls, recon=recon, assms=SOME thms}
+  {context=ctxt, output=ls, recon=recon, assms=SOME thms}
 
 fun gen_solver solver ctxt prems =
   let
-    val SolverConfig {command, arguments, interface, reconstruct} = solver ctxt
-    val Interface {normalize=nc, translate=tc} = interface
+    val {command, arguments, interface, reconstruct} = solver ctxt
+    val {normalize=nc, translate=tc} = interface
     val thy = ProofContext.theory_of ctxt
   in
     SMT_Normalize.normalize nc ctxt prems
--- a/src/HOL/SMT/Tools/smt_translate.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/SMT/Tools/smt_translate.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -34,29 +34,29 @@
   val bv_extract: (int -> int -> string) -> builtin_fun
 
   (* configuration options *)
-  datatype prefixes = Prefixes of {
+  type prefixes = {
     var_prefix: string,
     typ_prefix: string,
     fun_prefix: string,
     pred_prefix: string }
-  datatype markers = Markers of {
+  type markers = {
     term_marker: string,
     formula_marker: string }
-  datatype builtins = Builtins of {
+  type builtins = {
     builtin_typ: typ -> string option,
     builtin_num: int * typ -> string option,
     builtin_fun: bool -> builtin_table }
-  datatype sign = Sign of {
+  type sign = {
     typs: string list,
     funs: (string * (string list * string)) list,
     preds: (string * string list) list }
-  datatype config = Config of {
+  type config = {
     strict: bool,
     prefixes: prefixes,
     markers: markers,
     builtins: builtins,
     serialize: sign -> (string, string) sterm list -> TextIO.outstream -> unit}
-  datatype recon = Recon of {typs: typ Symtab.table, terms: term Symtab.table}
+  type recon = {typs: typ Symtab.table, terms: term Symtab.table}
 
   val translate: config -> theory -> thm list -> TextIO.outstream ->
     recon * thm list
@@ -159,29 +159,29 @@
 
 (* Configuration options *)
 
-datatype prefixes = Prefixes of {
+type prefixes = {
   var_prefix: string,
   typ_prefix: string,
   fun_prefix: string,
   pred_prefix: string }
-datatype markers = Markers of {
+type markers = {
   term_marker: string,
   formula_marker: string }
-datatype builtins = Builtins of {
+type builtins = {
   builtin_typ: typ -> string option,
   builtin_num: int * typ -> string option,
   builtin_fun: bool -> builtin_table }
-datatype sign = Sign of {
+type sign = {
   typs: string list,
   funs: (string * (string list * string)) list,
   preds: (string * string list) list }
-datatype config = Config of {
+type config = {
   strict: bool,
   prefixes: prefixes,
   markers: markers,
   builtins: builtins,
   serialize: sign -> (string, string) sterm list -> TextIO.outstream -> unit}
-datatype recon = Recon of {typs: typ Symtab.table, terms: term Symtab.table}
+type recon = {typs: typ Symtab.table, terms: term Symtab.table}
 
 
 (* Translate Isabelle/HOL terms into SMT intermediate terms.
@@ -409,7 +409,7 @@
   fun add_typ x (typs, funs, preds) = (Typtab.update x typs, funs, preds)
   fun add_fun true x (typs, funs, preds) = (typs, Termtab.update x funs, preds)
     | add_fun false x (typs, funs, preds) = (typs, funs, Termtab.update x preds)
-  fun make_sign (typs, funs, preds) = Sign {
+  fun make_sign (typs, funs, preds) = {
     typs = map snd (Typtab.dest typs),
     funs = map snd (Termtab.dest funs),
     preds = map (apsnd fst o snd) (Termtab.dest preds) }
@@ -418,11 +418,11 @@
       val rTs = Typtab.dest typs |> map swap |> Symtab.make
       val rts = Termtab.dest funs @ Termtab.dest preds
         |> map (apfst fst o swap) |> Symtab.make
-    in Recon {typs=rTs, terms=rts} end
+    in {typs=rTs, terms=rts} end
 
   fun either f g x = (case f x of NONE => g x | y => y)
 
-  fun rep_typ (Builtins {builtin_typ, ...}) T (st as (vars, ns, sgn)) =
+  fun rep_typ ({builtin_typ, ...} : builtins) T (st as (vars, ns, sgn)) =
     (case either builtin_typ (lookup_typ sgn) T of
       SOME n => (n, st)
     | NONE =>
@@ -444,16 +444,16 @@
           val (n, ns') = fresh_fun loc ns
         in (n, (vars, ns', add_fun loc (t, (n, uns)) sgn)) end)
 
-  fun rep_num (bs as Builtins {builtin_num, ...}) (i, T) st =
+  fun rep_num (bs as {builtin_num, ...} : builtins) (i, T) st =
     (case builtin_num (i, T) of
       SOME n => (n, st)
     | NONE => rep_fun bs true (HOLogic.mk_number T i) T 0 st)
 in
 fun signature_of prefixes markers builtins thy ts =
   let
-    val Prefixes {var_prefix, typ_prefix, fun_prefix, pred_prefix} = prefixes
-    val Markers {formula_marker, term_marker} = markers
-    val Builtins {builtin_fun, ...} = builtins
+    val {var_prefix, typ_prefix, fun_prefix, pred_prefix} = prefixes
+    val {formula_marker, term_marker} = markers
+    val {builtin_fun, ...} = builtins
 
     fun sign loc t =
       (case t of
@@ -493,7 +493,7 @@
 (* Combination of all translation functions and invocation of serialization. *)
 
 fun translate config thy thms stream =
-  let val Config {strict, prefixes, markers, builtins, serialize} = config
+  let val {strict, prefixes, markers, builtins, serialize} = config
   in
     map Thm.prop_of thms
     |> SMT_Monomorph.monomorph thy
--- a/src/HOL/SMT/Tools/smtlib_interface.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/SMT/Tools/smtlib_interface.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -109,7 +109,7 @@
           | wr_pat (T.SNoPat ts) = wrp "nopat" ts
       in par (wr_quant q #> fold wr_var vs #> wre b #> fold wr_pat ps) end)
 
-fun serialize attributes (T.Sign {typs, funs, preds}) ts stream =
+fun serialize attributes ({typs, funs, preds} : T.sign) ts stream =
   let
     fun wr_decl (n, Ts) = wr_line (sep (par (wr n #> fold wr1 Ts)))
   in
@@ -135,14 +135,14 @@
 
 (* SMT solver interface using the SMT-LIB input format *)
 
-val basic_builtins = T.Builtins {
+val basic_builtins = {
   builtin_typ = builtin_typ,
   builtin_num = builtin_num,
   builtin_fun = (fn true => builtin_funcs | false => builtin_preds) }
 
 val default_attributes = [":logic AUFLIRA", ":status unknown"]
 
-fun gen_interface builtins attributes = SMT_Solver.Interface {
+fun gen_interface builtins attributes = {
   normalize = [
     SMT_Normalize.RewriteTrivialLets,
     SMT_Normalize.RewriteNegativeNumerals,
@@ -150,14 +150,14 @@
     SMT_Normalize.AddAbsMinMaxRules,
     SMT_Normalize.AddPairRules,
     SMT_Normalize.AddFunUpdRules ],
-  translate = T.Config {
+  translate = {
     strict = true,
-    prefixes = T.Prefixes {
+    prefixes = {
       var_prefix = "x",
       typ_prefix = "T",
       fun_prefix = "uf_",
       pred_prefix = "up_" },
-    markers = T.Markers {
+    markers = {
       term_marker = term_marker,
       formula_marker = formula_marker },
     builtins = builtins,
--- a/src/HOL/SMT/Tools/yices_solver.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/SMT/Tools/yices_solver.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -26,7 +26,7 @@
 
 structure S = SMT_Solver
 
-fun core_oracle (SMT_Solver.ProofData {context, output, recon, ...}) =
+fun core_oracle ({context, output, recon, ...} : SMT_Solver.proof_data) =
   let
     val empty_line = (fn "" => true | _ => false)
     val split_first = (fn [] => ("", []) | l :: ls => (l, ls))
@@ -38,12 +38,11 @@
     else error (solver_name ^ " failed")
   end
 
-fun smtlib_solver oracle _ =
-  SMT_Solver.SolverConfig {
-    command = {env_var=env_var, remote_name=solver_name},
-    arguments = options,
-    interface = SMTLIB_Interface.interface,
-    reconstruct = oracle }
+fun smtlib_solver oracle _ = {
+  command = {env_var=env_var, remote_name=solver_name},
+  arguments = options,
+  interface = SMTLIB_Interface.interface,
+  reconstruct = oracle }
 
 val setup =
   Thm.add_oracle (Binding.name solver_name, core_oracle) #-> (fn (_, oracle) =>
--- a/src/HOL/SMT/Tools/z3_interface.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/SMT/Tools/z3_interface.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -88,7 +88,7 @@
   (@{term word_sless}, "bvslt"),
   (@{term word_sle}, "bvsle")]
 
-val builtins = T.Builtins {
+val builtins = {
   builtin_typ = builtin_typ,
   builtin_num = builtin_num,
   builtin_fun = (fn true => builtin_funcs | false => builtin_preds) }
--- a/src/HOL/SMT/Tools/z3_model.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/SMT/Tools/z3_model.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -47,7 +47,7 @@
 fun make_context (ttab, nctxt, vtab) =
   Context {ttab=ttab, nctxt=nctxt, vtab=vtab}
 
-fun empty_context (SMT_Translate.Recon {terms, ...}) =
+fun empty_context ({terms, ...} : SMT_Translate.recon) =
   let
     val ns = Symtab.fold (Term.add_free_names o snd) terms []
     val nctxt = Name.make_context ns
--- a/src/HOL/SMT/Tools/z3_proof.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/SMT/Tools/z3_proof.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -56,7 +56,7 @@
 fun make_context (Ttab, ttab, etab, ptab, nctxt) =
   Context {Ttab=Ttab, ttab=ttab, etab=etab, ptab=ptab, nctxt=nctxt}
 
-fun empty_context thy (SMT_Translate.Recon {typs, terms=ttab}) =
+fun empty_context thy ({typs, terms=ttab} : SMT_Translate.recon) =
   let
     val ttab' = Symtab.map (fn @{term True} => @{term "~False"} | t => t) ttab
     val ns = Symtab.fold (Term.add_free_names o snd) ttab' []
--- a/src/HOL/SMT/Tools/z3_proof_rules.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/SMT/Tools/z3_proof_rules.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -811,12 +811,12 @@
 
 (** |- ((ALL x. P x) | Q) = (ALL x. P x | Q) **)
 val pull_quant =
-  Thm o by_tac (Tactic.rtac @{thm refl} ORELSE' Simplifier.simp_tac HOL_ss)
+  Thm o by_tac (Tactic.rtac @{thm refl} ORELSE' Classical.best_tac HOL_cs)
 
 
 (** |- (ALL x. P x & Q x) = ((ALL x. P x) & (ALL x. Q x)) **)
 val push_quant =
-  Thm o by_tac (Tactic.rtac @{thm refl} ORELSE' Simplifier.simp_tac HOL_ss)
+  Thm o by_tac (Tactic.rtac @{thm refl} ORELSE' Classical.best_tac HOL_cs)
 
 
 (**
@@ -1131,15 +1131,12 @@
     iff_intro (under_assumption (contra_left conj) ct) (contra_right ct)
 
   fun conj_disj ct =
-    let
-      val cp as (cl, _) = Thm.dest_binop (Thm.dest_arg ct)
-      val (lhs, rhs) = pairself Thm.term_of cp
+    let val cp as (cl, cr) = Thm.dest_binop (Thm.dest_arg ct)
     in
-      if is_conj lhs andalso rhs = @{term False}
-      then contradiction true cl
-      else if is_disj lhs andalso rhs = @{term "~False"}
-      then contrapos2 (contradiction false o fst) cp
-      else prove_conj_disj_eq (Thm.dest_arg ct)
+      (case Thm.term_of cr of
+        @{term False} => contradiction true cl
+      | @{term "~False"} => contrapos2 (contradiction false o fst) cp
+      | _ => prove_conj_disj_eq (Thm.dest_arg ct))
     end
 
   val distinct =
--- a/src/HOL/SMT/Tools/z3_proof_terms.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/SMT/Tools/z3_proof_terms.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -89,7 +89,7 @@
 local
 fun mk_inst nctxt cert vs =
   let
-    val max = fold (curry Int.max o fst) vs 0
+    val max = fold (Integer.max o fst) vs 0
     val names = fst (Name.variants (replicate (max + 1) var_prefix) nctxt)
     fun mk (i, v) = (v, cert (Free (nth names i, #T (Thm.rep_cterm v))))
   in map mk vs end
--- a/src/HOL/SMT/Tools/z3_solver.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/SMT/Tools/z3_solver.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -52,22 +52,21 @@
     else error (solver_name ^ " failed")
   end
 
-fun core_oracle (SMT_Solver.ProofData {output, recon, ...}) =
+fun core_oracle ({output, recon, ...} : SMT_Solver.proof_data) =
   check_unsat recon output
   |> K @{cprop False}
 
-fun prover (SMT_Solver.ProofData {context, output, recon, assms}) =
+fun prover ({context, output, recon, assms} : SMT_Solver.proof_data) =
   check_unsat recon output
   |> Z3_Proof.reconstruct context assms recon
 
 fun solver oracle ctxt =
   let val with_proof = Config.get ctxt proofs
   in
-    SMT_Solver.SolverConfig {
-      command = {env_var=env_var, remote_name=solver_name},
-      arguments = cmdline_options ctxt,
-      interface = Z3_Interface.interface,
-      reconstruct = if with_proof then prover else oracle }
+    {command = {env_var=env_var, remote_name=solver_name},
+    arguments = cmdline_options ctxt,
+    interface = Z3_Interface.interface,
+    reconstruct = if with_proof then prover else oracle}
   end
 
 val setup =
--- a/src/HOL/SMT/Z3.thy	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/SMT/Z3.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -19,6 +19,30 @@
 
 lemmas [z3_rewrite] =
   refl eq_commute conj_commute disj_commute simp_thms nnf_simps
-  ring_distribs field_eq_simps
+  ring_distribs field_eq_simps if_True if_False
+
+lemma [z3_rewrite]: "(\<not>False \<longrightarrow> P) = P" by fast
+lemma [z3_rewrite]: "(P \<longrightarrow> Q) = (Q \<or> \<not>P)" by fast
+lemma [z3_rewrite]: "(\<not>P \<longrightarrow> Q) = (P \<or> Q)" by fast
+lemma [z3_rewrite]: "(\<not>P \<longrightarrow> Q) = (Q \<or> P)" by fast
+
+lemma [z3_rewrite]: "(if P then True else False) = P" by simp
+lemma [z3_rewrite]: "(if P then False else True) = (\<not>P)" by simp
+
+lemma [z3_rewrite]:
+  "0 + (x::int) = x"
+  "x + 0 = x"
+  "0 * x = 0"
+  "1 * x = x"
+  "x + y = y + x"
+  by auto
+
+lemma [z3_rewrite]:
+  "0 + (x::real) = x"
+  "x + 0 = x"
+  "0 * x = 0"
+  "1 * x = x"
+  "x + y = y + x"
+  by auto
 
 end
--- a/src/HOL/Set.thy	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Set.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -445,7 +445,7 @@
 
 subsubsection {* Subsets *}
 
-lemma subsetI [atp, intro!]: "(\<And>x. x \<in> A \<Longrightarrow> x \<in> B) \<Longrightarrow> A \<subseteq> B"
+lemma subsetI [intro!]: "(\<And>x. x \<in> A \<Longrightarrow> x \<in> B) \<Longrightarrow> A \<subseteq> B"
   unfolding mem_def by (rule le_funI, rule le_boolI)
 
 text {*
@@ -476,7 +476,7 @@
 lemma contra_subsetD: "A \<subseteq> B ==> c \<notin> B ==> c \<notin> A"
   by blast
 
-lemma subset_refl [simp,atp]: "A \<subseteq> A"
+lemma subset_refl [simp]: "A \<subseteq> A"
   by (fact order_refl)
 
 lemma subset_trans: "A \<subseteq> B ==> B \<subseteq> C ==> A \<subseteq> C"
--- a/src/HOL/Tools/Datatype/datatype_codegen.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Tools/Datatype/datatype_codegen.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -173,7 +173,7 @@
                  [Pretty.block [str "(case", Pretty.brk 1,
                    str "i", Pretty.brk 1, str "of",
                    Pretty.brk 1, str "0 =>", Pretty.brk 1,
-                   mk_constr "0" true (cname, valOf (AList.lookup (op =) cs cname)),
+                   mk_constr "0" true (cname, the (AList.lookup (op =) cs cname)),
                    Pretty.brk 1, str "| _ =>", Pretty.brk 1,
                    mk_choice cs1, str ")"]]
                else [mk_choice cs2])) ::
--- a/src/HOL/Tools/Datatype/datatype_realizer.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Tools/Datatype/datatype_realizer.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -136,7 +136,7 @@
     val rvs = rev (Thm.fold_terms Term.add_vars thm' []);
     val ivs1 = map Var (filter_out (fn (_, T) =>
       tname_of (body_type T) mem ["set", "bool"]) ivs);
-    val ivs2 = map (fn (ixn, _) => Var (ixn, valOf (AList.lookup (op =) rvs ixn))) ivs;
+    val ivs2 = map (fn (ixn, _) => Var (ixn, the (AList.lookup (op =) rvs ixn))) ivs;
 
     val prf = List.foldr forall_intr_prf
      (List.foldr (fn ((f, p), prf) =>
--- a/src/HOL/Tools/Function/scnp_solve.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Tools/Function/scnp_solve.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -79,7 +79,7 @@
 fun var_constrs (gp as GP (arities, gl)) =
   let
     val n = Int.max (num_graphs gp, num_prog_pts gp)
-    val k = List.foldl Int.max 1 arities
+    val k = fold Integer.max arities 1
 
     (* Injective, provided  a < 8, x < n, and i < k. *)
     fun prod a x i j = ((j * k + i) * n + x) * 8 + a + 1
--- a/src/HOL/Tools/Groebner_Basis/groebner.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Tools/Groebner_Basis/groebner.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -380,7 +380,7 @@
         val pol' = augment pol
         val allpols = (grob_sub (grob_mul grob_z pol') grob_1)::pols'
         val (l,cert) = grobner_weak vars' allpols
-        val d = fold_rev (fold_rev (curry Int.max o hd o snd) o snd) cert 0
+        val d = fold (fold (Integer.max o hd o snd) o snd) cert 0
         fun transform_monomial (c,m) =
             grob_cmul (c,tl m) (grob_pow vars pol (d - hd m))
         fun transform_polynomial q = fold_rev (grob_add o transform_monomial) q []
@@ -505,8 +505,8 @@
   val (l,r) = Thm.dest_equals eq
   val ctabl = mk_conj_tab (assume (Thm.capply @{cterm Trueprop} l))
   val ctabr = mk_conj_tab (assume (Thm.capply @{cterm Trueprop} r))
-  fun tabl c = valOf (Termtab.lookup ctabl (term_of c))
-  fun tabr c = valOf (Termtab.lookup ctabr (term_of c))
+  fun tabl c = the (Termtab.lookup ctabl (term_of c))
+  fun tabr c = the (Termtab.lookup ctabr (term_of c))
   val thl  = prove_conj tabl (conjuncts r) |> implies_intr_hyps
   val thr  = prove_conj tabr (conjuncts l) |> implies_intr_hyps
   val eqI = instantiate' [] [SOME l, SOME r] @{thm iffI}
@@ -839,7 +839,7 @@
    val (v,th1) = 
    case find_first(fn v=> is_defined v (Thm.dest_arg1 (Thm.rhs_of th))) vars of
     SOME v => (v,th')
-   | NONE => (valOf (find_first 
+   | NONE => (the (find_first 
           (fn v => is_defined v (Thm.dest_arg1 (Thm.rhs_of th'))) vars) ,th)
    val th2 = transitive th1 
         (instantiate' []  [(SOME o Thm.dest_arg1 o Thm.rhs_of) th1, SOME v] 
@@ -851,7 +851,7 @@
  let 
   val (vars,bod) = strip_exists tm
   val cjs = striplist (dest_binary @{cterm "op &"}) bod
-  val th1 = (valOf (get_first (try (isolate_variable vars)) cjs) 
+  val th1 = (the (get_first (try (isolate_variable vars)) cjs) 
              handle Option => raise CTERM ("unwind_polys_conv",[tm]))
   val eq = Thm.lhs_of th1
   val bod' = list_mk_binop @{cterm "op &"} (eq::(remove op aconvc eq cjs))
@@ -906,7 +906,7 @@
  fun solve_idealism evs ps eqs =
   if null evs then [] else
   let 
-   val (eq,cfs) = get_first (try (isolate_variables evs ps)) eqs |> valOf
+   val (eq,cfs) = get_first (try (isolate_variables evs ps)) eqs |> the
    val evs' = subtract op aconvc evs (map snd cfs)
    val eqs' = map (subst_in_poly cfs) (remove op aconvc eq eqs)
   in cfs @ solve_idealism evs' ps eqs'
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -615,7 +615,7 @@
 fun guess_nparams T =
   let
     val argTs = binder_types T
-    val nparams = fold (curry Int.max)
+    val nparams = fold Integer.max
       (map (fn x => x + 1) (find_indexes is_predT argTs)) 0
   in nparams end;
 
--- a/src/HOL/Tools/Qelim/cooper.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Tools/Qelim/cooper.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -325,7 +325,7 @@
     let val tab = fold Inttab.update
           (ds ~~ (map (fn x => nzprop (l div x)) ds)) Inttab.empty
     in
-      fn ct => valOf (Inttab.lookup tab (ct |> term_of |> dest_numeral))
+      fn ct => the (Inttab.lookup tab (ct |> term_of |> dest_numeral))
         handle Option =>
           (writeln ("noz: Theorems-Table contains no entry for " ^
               Syntax.string_of_term ctxt (Thm.term_of ct)); raise Option)
@@ -416,7 +416,7 @@
    in equal_elim (Thm.symmetric th) TrueI end;
  val dvd =
    let val tab = fold Inttab.update (ds ~~ (map divprop ds)) Inttab.empty in
-     fn ct => valOf (Inttab.lookup tab (term_of ct |> dest_numeral))
+     fn ct => the (Inttab.lookup tab (term_of ct |> dest_numeral))
        handle Option =>
         (writeln ("dvd: Theorems-Table contains no entry for" ^
             Syntax.string_of_term ctxt (Thm.term_of ct)); raise Option)
@@ -483,13 +483,13 @@
         (map (fn eq =>
                 let val (s,t) = cprop_of eq |> Thm.dest_arg |> Thm.dest_binop
                     val th = if term_of s = term_of t
-                             then valOf(Termtab.lookup inStab (term_of s))
+                             then the (Termtab.lookup inStab (term_of s))
                              else FWD (instantiate' [] [SOME s, SOME t] eqelem_th)
-                                [eq, valOf(Termtab.lookup inStab (term_of s))]
+                                [eq, the (Termtab.lookup inStab (term_of s))]
                  in (term_of t, th) end)
                   sths) Termtab.empty
         in
-          fn ct => valOf (Termtab.lookup tab (term_of ct))
+          fn ct => the (Termtab.lookup tab (term_of ct))
             handle Option =>
               (writeln ("inS: No theorem for " ^ Syntax.string_of_term ctxt (Thm.term_of ct));
                 raise Option)
--- a/src/HOL/Tools/Qelim/langford.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Tools/Qelim/langford.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -101,8 +101,8 @@
 fun conj_aci_rule eq = 
  let 
   val (l,r) = Thm.dest_equals eq
-  fun tabl c = valOf (Termtab.lookup (mk_conj_tab (assume l)) (term_of c))
-  fun tabr c = valOf (Termtab.lookup (mk_conj_tab (assume r)) (term_of c))
+  fun tabl c = the (Termtab.lookup (mk_conj_tab (assume l)) (term_of c))
+  fun tabr c = the (Termtab.lookup (mk_conj_tab (assume r)) (term_of c))
   val ll = Thm.dest_arg l
   val rr = Thm.dest_arg r
   
--- a/src/HOL/Tools/TFL/rules.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Tools/TFL/rules.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -658,7 +658,7 @@
   in (ants,get_lhs eq)
   end;
 
-fun restricted t = isSome (S.find_term
+fun restricted t = is_some (S.find_term
                             (fn (Const(@{const_name Recdef.cut},_)) =>true | _ => false)
                             t)
 
@@ -784,7 +784,7 @@
               val dummy = print_cterm "func:" (cterm_of thy func)
               val dummy = print_cterm "TC:" (cterm_of thy (HOLogic.mk_Trueprop TC))
               val dummy = tc_list := (TC :: !tc_list)
-              val nestedp = isSome (S.find_term is_func TC)
+              val nestedp = is_some (S.find_term is_func TC)
               val dummy = if nestedp then say "nested" else say "not_nested"
               val dummy = term_ref := ([func,TC]@(!term_ref))
               val th' = if nestedp then raise RULES_ERR "solver" "nested function"
--- a/src/HOL/Tools/TFL/tfl.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Tools/TFL/tfl.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -724,7 +724,7 @@
 in
 fun build_ih f P (pat,TCs) =
  let val globals = S.free_vars_lr pat
-     fun nested tm = isSome (S.find_term (curry (op aconv) f) tm)
+     fun nested tm = is_some (S.find_term (curry (op aconv) f) tm)
      fun dest_TC tm =
          let val (cntxt,R_y_pat) = S.strip_imp(#2(S.strip_forall tm))
              val (R,y,_) = S.dest_relation R_y_pat
@@ -753,7 +753,7 @@
 fun build_ih f (P,SV) (pat,TCs) =
  let val pat_vars = S.free_vars_lr pat
      val globals = pat_vars@SV
-     fun nested tm = isSome (S.find_term (curry (op aconv) f) tm)
+     fun nested tm = is_some (S.find_term (curry (op aconv) f) tm)
      fun dest_TC tm =
          let val (cntxt,R_y_pat) = S.strip_imp(#2(S.strip_forall tm))
              val (R,y,_) = S.dest_relation R_y_pat
@@ -786,7 +786,7 @@
  let val tych = Thry.typecheck thy
      val antc = tych(#ant(S.dest_imp tm))
      val thm' = R.SPEC_ALL thm
-     fun nested tm = isSome (S.find_term (curry (op aconv) f) tm)
+     fun nested tm = is_some (S.find_term (curry (op aconv) f) tm)
      fun get_cntxt TC = tych(#ant(S.dest_imp(#2(S.strip_forall(concl TC)))))
      fun mk_ih ((TC,locals),th2,nested) =
          R.GENL (map tych locals)
--- a/src/HOL/Tools/cnf_funcs.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Tools/cnf_funcs.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -489,7 +489,7 @@
 				else
 					NONE
 		in
-			Int.max (max, getOpt (idx, 0))
+			Int.max (max, the_default 0 idx)
 		end) (OldTerm.term_frees simp) 0)
 	(* finally, convert to definitional CNF (this should preserve the simplification) *)
 	val cnfx_thm = make_cnfx_thm_from_nnf simp
--- a/src/HOL/Tools/lin_arith.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Tools/lin_arith.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -672,7 +672,7 @@
       (case split_once_items ctxt subgoal of
         SOME new_subgoals => split_loop (new_subgoals @ subgoals)
       | NONE              => subgoal :: split_loop subgoals)
-  fun is_relevant t  = isSome (decomp ctxt t)
+  fun is_relevant t  = is_some (decomp ctxt t)
   (* filter_prems_tac is_relevant: *)
   val relevant_terms = filter_prems_tac_items is_relevant terms
   (* split_tac, NNF normalization: *)
@@ -746,7 +746,7 @@
 fun pre_tac ctxt i =
 let
   val split_thms = filter is_split_thm (#splits (get_arith_data ctxt))
-  fun is_relevant t = isSome (decomp ctxt t)
+  fun is_relevant t = is_some (decomp ctxt t)
 in
   DETERM (
     TRY (filter_prems_tac is_relevant i)
@@ -878,7 +878,7 @@
 local
 
 fun raw_tac ctxt ex =
-  (* FIXME: K true should be replaced by a sensible test (perhaps "isSome o
+  (* FIXME: K true should be replaced by a sensible test (perhaps "is_some o
      decomp sg"? -- but note that the test is applied to terms already before
      they are split/normalized) to speed things up in case there are lots of
      irrelevant terms involved; elimination of min/max can be optimized:
--- a/src/HOL/Tools/metis_tools.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Tools/metis_tools.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -123,7 +123,7 @@
   | metis_of_typeLit pos (ResClause.LTFree (s,x)) = metis_lit pos s [Metis.Term.Fn(x,[])];
 
 fun default_sort _ (TVar _) = false
-  | default_sort ctxt (TFree (x, s)) = (s = Option.getOpt (Variable.def_sort ctxt (x, ~1), []));
+  | default_sort ctxt (TFree (x, s)) = (s = the_default [] (Variable.def_sort ctxt (x, ~1)));
 
 fun metis_of_tfree tf =
   Metis.Thm.axiom (Metis.LiteralSet.singleton (metis_of_typeLit true tf));
@@ -204,7 +204,7 @@
 (*include the default sort, if available*)
 fun mk_tfree ctxt w =
   let val ww = "'" ^ w
-  in  TFree(ww, getOpt (Variable.def_sort ctxt (ww,~1), HOLogic.typeS))  end;
+  in  TFree(ww, the_default HOLogic.typeS (Variable.def_sort ctxt (ww, ~1)))  end;
 
 (*Remove the "apply" operator from an HO term*)
 fun strip_happ args (Metis.Term.Fn(".",[t,u])) = strip_happ (u::args) t
@@ -342,7 +342,7 @@
    trace_msg (fn () => "Isabelle: " ^ Display.string_of_thm_without_context th));
 
 fun lookth thpairs (fth : Metis.Thm.thm) =
-  valOf (AList.lookup (uncurry Metis.Thm.equal) thpairs fth)
+  the (AList.lookup (uncurry Metis.Thm.equal) thpairs fth)
   handle Option => error ("Failed to find a Metis theorem " ^ Metis.Thm.toString fth);
 
 fun is_TrueI th = Thm.eq_thm(TrueI,th);
@@ -372,7 +372,7 @@
   let val thy = ProofContext.theory_of ctxt
       val i_th   = lookth thpairs th
       val i_th_vars = Term.add_vars (prop_of i_th) []
-      fun find_var x = valOf (List.find (fn ((a,_),_) => a=x) i_th_vars)
+      fun find_var x = the (List.find (fn ((a,_),_) => a=x) i_th_vars)
       fun subst_translation (x,y) =
             let val v = find_var x
                 val t = fol_term_to_hol ctxt mode y (*we call infer_types below*)
--- a/src/HOL/Tools/prop_logic.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Tools/prop_logic.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -387,7 +387,7 @@
 			if !next_idx_is_valid then
 				Unsynchronized.inc next_idx
 			else (
-				next_idx          := Termtab.fold (curry Int.max o snd) table 0;
+				next_idx := Termtab.fold (Integer.max o snd) table 0;
 				next_idx_is_valid := true;
 				Unsynchronized.inc next_idx
 			)
--- a/src/HOL/Tools/quickcheck_generators.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Tools/quickcheck_generators.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -256,7 +256,7 @@
         val T = termifyT simpleT;
         val tTs = (map o apsnd) termifyT simple_tTs;
         val is_rec = exists is_some ks;
-        val k = fold (fn NONE => I | SOME k => curry Int.max k) ks 0;
+        val k = fold (fn NONE => I | SOME k => Integer.max k) ks 0;
         val vs = Name.names Name.context "x" (map snd simple_tTs);
         val vs' = (map o apsnd) termifyT vs;
         val tc = HOLogic.mk_return T @{typ Random.seed}
--- a/src/HOL/Tools/record.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Tools/record.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -750,7 +750,7 @@
                     val types = map snd flds';
                     val (args, rest) = splitargs (map fst flds') fargs;
                     val argtypes = map (Sign.certify_typ thy o decode_type thy) args;
-                    val midx = fold (fn T => fn i => Int.max (maxidx_of_typ T, i)) argtypes 0;
+                    val midx = fold Term.maxidx_typ argtypes 0;
                     val varifyT = varifyT midx;
                     val vartypes = map varifyT types;
 
--- a/src/HOL/Tools/refute.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Tools/refute.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -1081,7 +1081,7 @@
       | next max len sum (x1::x2::xs) =
       if x1>0 andalso (x2<max orelse max<0) then
         (* we can shift *)
-        SOME (valOf (make_first max (len+1) (sum+x1-1)) @ (x2+1) :: xs)
+        SOME (the (make_first max (len+1) (sum+x1-1)) @ (x2+1) :: xs)
       else
         (* continue search *)
         next max (len+1) (sum+x1) (x2::xs)
--- a/src/HOL/Tools/res_atp.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Tools/res_atp.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -38,8 +38,6 @@
 val pass_mark = 0.5;
 val convergence = 3.2;    (*Higher numbers allow longer inference chains*)
 val follow_defs = false;  (*Follow definitions. Makes problems bigger.*)
-val include_all = true;
-val include_atpset = true;
   
 (***************************************************************)
 (* Relevance Filtering                                         *)
@@ -294,15 +292,14 @@
 
 (*** white list and black list of lemmas ***)
 
-(*The rule subsetI is frequently omitted by the relevance filter. This could be theory data
-  or identified with ATPset (which however is too big currently)*)
+(*The rule subsetI is frequently omitted by the relevance filter.*)
 val whitelist_fo = [subsetI];
 (* ext has been a 'helper clause', always given to the atps.
   As such it was not passed to metis, but metis does not include ext (in contrast
   to the other helper_clauses *)
 val whitelist_ho = [ResHolClause.ext];
 
-(*** retrieve lemmas from clasimpset and atpset, may filter them ***)
+(*** retrieve lemmas and filter them ***)
 
 (*Hashing to detect duplicate and variant clauses, e.g. from the [iff] attribute*)
 
@@ -356,7 +353,7 @@
   boost an ATP's performance (for some reason)*)
 fun subtract_cls c_clauses ax_clauses =
   let val ht = mk_clause_table 2200
-      fun known x = isSome (Polyhash.peek ht x)
+      fun known x = is_some (Polyhash.peek ht x)
   in
       app (ignore o Polyhash.peekInsert ht) ax_clauses;
       filter (not o known) c_clauses
@@ -395,7 +392,7 @@
 fun name_thm_pairs ctxt =
   let val (mults,singles) = List.partition is_multi (all_valid_thms ctxt)
       val ht = mk_clause_table 900   (*ht for blacklisted theorems*)
-      fun blacklisted x = run_blacklist_filter andalso isSome (Polyhash.peek ht x)
+      fun blacklisted x = run_blacklist_filter andalso is_some (Polyhash.peek ht x)
   in
       app (fn th => ignore (Polyhash.peekInsert ht (th,()))) (ResBlacklist.get ctxt);
       filter (not o blacklisted o #2)
@@ -406,18 +403,11 @@
       (warning ("No name for theorem " ^ Display.string_of_thm_without_context th); false)
   | check_named _ = true;
 
-(* get lemmas from claset, simpset, atpset and extra supplied rules *)
-fun get_clasimp_atp_lemmas ctxt =
+fun get_all_lemmas ctxt =
   let val included_thms =
-    if include_all
-    then (tap (fn ths => ResAxioms.trace_msg
-                 (fn () => ("Including all " ^ Int.toString (length ths) ^ " theorems")))
-              (name_thm_pairs ctxt))
-    else
-    let val atpset_thms =
-            if include_atpset then ResAxioms.atpset_rules_of ctxt
-            else []
-    in  atpset_thms  end
+        tap (fn ths => ResAxioms.trace_msg
+                     (fn () => ("Including all " ^ Int.toString (length ths) ^ " theorems")))
+            (name_thm_pairs ctxt)
   in
     filter check_named included_thms
   end;
@@ -525,7 +515,7 @@
   let
     val thy = ProofContext.theory_of ctxt
     val isFO = isFO thy goal_cls
-    val included_thms = get_clasimp_atp_lemmas ctxt
+    val included_thms = get_all_lemmas ctxt
     val included_cls = included_thms |> ResAxioms.cnf_rules_pairs thy |> make_unique
                                      |> restrict_to_logic thy isFO
                                      |> remove_unwanted_clauses
--- a/src/HOL/Tools/res_axioms.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Tools/res_axioms.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -17,7 +17,6 @@
   val expand_defs_tac: thm -> tactic
   val combinators: thm -> thm
   val neg_conjecture_clauses: Proof.context -> thm -> int -> thm list * (string * typ) list
-  val atpset_rules_of: Proof.context -> (string * thm) list
   val suppress_endtheory: bool Unsynchronized.ref
     (*for emergency use where endtheory causes problems*)
   val setup: theory -> theory
@@ -302,7 +301,7 @@
       then exists (excessive_lambdas_fm Ts) (#2 (strip_comb t))
       else excessive_lambdas (t, max_lambda_nesting);
 
-(*The max apply_depth of any metis call in MetisExamples (on 31-10-2007) was 11.*)
+(*The max apply_depth of any metis call in Metis_Examples (on 31-10-2007) was 11.*)
 val max_apply_depth = 15;
 
 fun apply_depth (f$t) = Int.max (apply_depth f, apply_depth t + 1)
@@ -378,8 +377,6 @@
 
 fun pairname th = (Thm.get_name_hint th, th);
 
-fun atpset_rules_of ctxt = map pairname (ResAtpset.get ctxt);
-
 
 (**** Translate a set of theorems into CNF ****)
 
--- a/src/HOL/Tools/res_hol_clause.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Tools/res_hol_clause.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -69,12 +69,13 @@
   use of the "apply" operator. Use of hBOOL is also minimized.*)
 val minimize_applies = true;
 
-fun min_arity_of const_min_arity c = getOpt (Symtab.lookup const_min_arity c, 0);
+fun min_arity_of const_min_arity c = the_default 0 (Symtab.lookup const_min_arity c);
 
 (*True if the constant ever appears outside of the top-level position in literals.
   If false, the constant always receives all of its arguments and is used as a predicate.*)
-fun needs_hBOOL const_needs_hBOOL c = not minimize_applies orelse
-                    getOpt (Symtab.lookup const_needs_hBOOL c, false);
+fun needs_hBOOL const_needs_hBOOL c =
+  not minimize_applies orelse
+    the_default false (Symtab.lookup const_needs_hBOOL c);
 
 
 (******************************************************)
@@ -412,7 +413,7 @@
         val axclauses = map #2 (make_axiom_clauses dfg thy axcls)
         val ct0 = List.foldl count_clause init_counters conjectures
         val ct = List.foldl (count_user_clause user_lemmas) ct0 axclauses
-        fun needed c = valOf (Symtab.lookup ct c) > 0
+        fun needed c = the (Symtab.lookup ct c) > 0
         val IK = if needed "c_COMBI" orelse needed "c_COMBK"
                  then cnf_helper_thms thy [comb_I,comb_K]
                  else []
@@ -437,7 +438,7 @@
       case head of
           CombConst (a,_,_) => (*predicate or function version of "equal"?*)
             let val a = if a="equal" andalso not toplev then "c_fequal" else a
-            val const_min_arity = Symtab.map_default (a,n) (curry Int.min n) const_min_arity
+            val const_min_arity = Symtab.map_default (a, n) (Integer.min n) const_min_arity
             in
               if toplev then (const_min_arity, const_needs_hBOOL)
               else (const_min_arity, Symtab.update (a,true) (const_needs_hBOOL))
--- a/src/HOL/Tools/res_reconstruct.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Tools/res_reconstruct.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -64,7 +64,7 @@
 
 (*Integer constants, typically proof line numbers*)
 fun is_digit s = Char.isDigit (String.sub(s,0));
-val integer = Scan.many1 is_digit >> (valOf o Int.fromString o implode);
+val integer = Scan.many1 is_digit >> (the o Int.fromString o implode);
 
 (*Generalized FO terms, which include filenames, numbers, etc.*)
 fun termlist x = (term ::: Scan.repeat ($$"," |-- term)) x
@@ -225,8 +225,8 @@
 (*Update TVars/TFrees with detected sort constraints.*)
 fun fix_sorts vt =
   let fun tysubst (Type (a, Ts)) = Type (a, map tysubst Ts)
-        | tysubst (TVar (xi, s)) = TVar (xi, getOpt (Vartab.lookup vt xi, s))
-        | tysubst (TFree (x, s)) = TFree (x, getOpt (Vartab.lookup vt (x,~1), s))
+        | tysubst (TVar (xi, s)) = TVar (xi, the_default s (Vartab.lookup vt xi))
+        | tysubst (TFree (x, s)) = TFree (x, the_default s (Vartab.lookup vt (x, ~1)))
       fun tmsubst (Const (a, T)) = Const (a, tysubst T)
         | tmsubst (Free (a, T)) = Free (a, tysubst T)
         | tmsubst (Var (xi, T)) = Var (xi, tysubst T)
--- a/src/HOL/Tools/sat_funcs.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Tools/sat_funcs.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -218,9 +218,9 @@
 	fun index_of_literal chyp = (
 		case (HOLogic.dest_Trueprop o Thm.term_of) chyp of
 		  (Const ("Not", _) $ atom) =>
-			SOME (~(valOf (Termtab.lookup atom_table atom)))
+			SOME (~ (the (Termtab.lookup atom_table atom)))
 		| atom =>
-			SOME (valOf (Termtab.lookup atom_table atom))
+			SOME (the (Termtab.lookup atom_table atom))
 	) handle TERM _ => NONE;  (* 'chyp' is not a literal *)
 
 	(* int -> Thm.thm * (int * Thm.cterm) list *)
@@ -244,7 +244,7 @@
 			(* prove the clause, using information from 'clause_table' *)
 			let
 				val _      = if !trace_sat then tracing ("Proving clause #" ^ string_of_int id ^ " ...") else ()
-				val ids    = valOf (Inttab.lookup clause_table id)
+				val ids    = the (Inttab.lookup clause_table id)
 				val clause = resolve_raw_clauses (map prove_clause ids)
 				val _      = Array.update (clauses, id, RAW_CLAUSE clause)
 				val _      = if !trace_sat then tracing ("Replay chain successful; clause stored at #" ^ string_of_int id) else ()
@@ -367,7 +367,7 @@
 				(* [[c_1 && ... && c_n] |- c_1, ..., [c_1 && ... && c_n] |- c_n] *)
 				val cnf_clauses = Conjunction.elim_balanced (length sorted_clauses) cnf_thm
 				(* initialize the clause array with the given clauses *)
-				val max_idx    = valOf (Inttab.max_key clause_table)
+				val max_idx    = the (Inttab.max_key clause_table)
 				val clause_arr = Array.array (max_idx + 1, NO_CLAUSE)
 				val _          = fold (fn thm => fn idx => (Array.update (clause_arr, idx, ORIG_CLAUSE thm); idx+1)) cnf_clauses 0
 				(* replay the proof to derive the empty clause *)
--- a/src/HOL/Tools/sat_solver.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/Tools/sat_solver.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -502,7 +502,7 @@
         | False => NONE
         | _     =>
           let
-            val x = valOf (fresh_var xs')  (* a fresh variable must exist since 'fm' did not evaluate to 'True'/'False' *)
+            val x = the (fresh_var xs')  (* a fresh variable must exist since 'fm' did not evaluate to 'True'/'False' *)
           in
             case dpll (x::xs') fm' of  (* passing fm' rather than fm should save some simplification work *)
               SOME xs'' => SOME xs''
@@ -893,7 +893,7 @@
               (* for its literals to obtain the empty clause                *)
               val vids         = map (fn l => l div 2) ls
               val rs           = cid :: map (fn v => !clause_offset + v) vids
-              val new_empty_id = getOpt (Inttab.max_key (!clause_table), !clause_offset) + 1
+              val new_empty_id = the_default (!clause_offset) (Inttab.max_key (!clause_table)) + 1
             in
               (* update clause table and conflict id *)
               clause_table := Inttab.update_new (new_empty_id, rs) (!clause_table)
--- a/src/HOL/ex/ThreeDivides.thy	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/ex/ThreeDivides.thy	Wed Oct 21 08:16:25 2009 +0200
@@ -1,5 +1,4 @@
-(*  Title:      HOL/Isar_examples/ThreeDivides.thy
-    ID:         $Id$
+(*  Title:      HOL/ex/ThreeDivides.thy
     Author:     Benjamin Porter, 2005
 *)
 
--- a/src/HOL/ex/document/root.bib	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/ex/document/root.bib	Wed Oct 21 08:16:25 2009 +0200
@@ -80,7 +80,7 @@
                   Higher-Order Logic},
   year =         2001,
   note =         {Part of the Isabelle distribution,
-                  \url{http://isabelle.in.tum.de/library/HOL/Isar_examples/document.pdf}}
+                  \url{http://isabelle.in.tum.de/library/HOL/Isar_Examples/document.pdf}}
 }
 
 @PhdThesis{Wenzel:2001:Thesis,
--- a/src/HOL/ex/predicate_compile.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/HOL/ex/predicate_compile.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -542,7 +542,7 @@
 fun guess_nparams T =
   let
     val argTs = binder_types T
-    val nparams = fold (curry Int.max)
+    val nparams = fold Integer.max
       (map (fn x => x + 1) (find_indexes is_predT argTs)) 0
   in nparams end;
 
--- a/src/Provers/splitter.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/Provers/splitter.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -159,9 +159,9 @@
 (* check if the innermost abstraction that needs to be removed
    has a body of type T; otherwise the expansion thm will fail later on
 *)
-fun type_test (T,lbnos,apsns) =
-  let val (_,U: typ,_) = List.nth(apsns, Library.foldl Int.min (hd lbnos, tl lbnos))
-  in T=U end;
+fun type_test (T, lbnos, apsns) =
+  let val (_, U: typ, _) = List.nth (apsns, foldl1 Int.min lbnos)
+  in T = U end;
 
 (*************************************************************************
    Create a "split_pack".
--- a/src/Pure/Concurrent/lazy.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/Pure/Concurrent/lazy.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -12,7 +12,6 @@
   val value: 'a -> 'a lazy
   val force_result: 'a lazy -> 'a Exn.result
   val force: 'a lazy -> 'a
-  val map_force: ('a -> 'b) -> 'a lazy -> 'b lazy
 end;
 
 structure Lazy: LAZY =
@@ -49,8 +48,6 @@
 
 fun force r = Exn.release (force_result r);
 
-fun map_force f = value o f o force;
-
 end;
 end;
 
--- a/src/Pure/Concurrent/lazy_sequential.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/Pure/Concurrent/lazy_sequential.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -34,8 +34,6 @@
 
 fun force r = Exn.release (force_result r);
 
-fun map_force f = value o f o force;
-
 end;
 end;
 
--- a/src/Pure/General/integer.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/Pure/General/integer.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -6,6 +6,8 @@
 
 signature INTEGER =
 sig
+  val min: int -> int -> int
+  val max: int -> int -> int
   val add: int -> int -> int
   val mult: int -> int -> int
   val sum: int list -> int
@@ -23,6 +25,9 @@
 structure Integer : INTEGER =
 struct
 
+fun min x y = Int.min (x, y);
+fun max x y = Int.max (x, y);
+
 fun add x y = x + y;
 fun mult x y = x * y;
 
--- a/src/Pure/Isar/proof.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/Pure/Isar/proof.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -7,7 +7,7 @@
 
 signature PROOF =
 sig
-  type context = Context.proof
+  type context = Proof.context
   type method = Method.method
   type state
   val init: context -> state
@@ -121,7 +121,7 @@
 structure Proof: PROOF =
 struct
 
-type context = Context.proof;
+type context = Proof.context;
 type method = Method.method;
 
 
--- a/src/Pure/Isar/proof_context.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/Pure/Isar/proof_context.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -142,8 +142,7 @@
 structure ProofContext: PROOF_CONTEXT =
 struct
 
-val theory_of = Context.theory_of_proof;
-val init = Context.init_proof;
+open ProofContext;
 
 
 (** inner syntax mode **)
@@ -259,7 +258,7 @@
       else (Consts.merge (local_consts, thy_consts), thy_consts)
     end);
 
-fun transfer thy = Context.transfer_proof thy #> transfer_syntax thy;
+fun transfer thy = Context.raw_transfer thy #> transfer_syntax thy;
 
 fun theory f ctxt = transfer (f (theory_of ctxt)) ctxt;
 
--- a/src/Pure/ProofGeneral/pgip_output.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/Pure/ProofGeneral/pgip_output.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -219,8 +219,8 @@
         fun fileurl_to_xml url = XML.Elem ("fileurl", attrs_of_pgipurl url, [])
     in
         XML.Elem ("setrefs",
-                  (Option.getOpt (Option.map attrs_of_pgipurl url,[])) @ 
-                  (Option.getOpt (Option.map attrs_of_objtype objtype,[])) @
+                  (the_default [] (Option.map attrs_of_pgipurl url)) @ 
+                  (the_default [] (Option.map attrs_of_objtype objtype)) @
                   (opt_attr "thyname" thyname) @
                   (opt_attr "name" name),
                   (map idtable_to_xml idtables) @ 
--- a/src/Pure/ProofGeneral/pgml.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/Pure/ProofGeneral/pgml.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -155,6 +155,6 @@
         XML.Elem("pgml",
                  opt_attr "version" version @
                  opt_attr "systemid" systemid @
-                 Option.getOpt(Option.map PgipTypes.attrs_of_displayarea area,[]),
+                 the_default [] (Option.map PgipTypes.attrs_of_displayarea area),
                  map pgmlterm_to_xml content)
 end
--- a/src/Pure/ProofGeneral/proof_general_pgip.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/Pure/ProofGeneral/proof_general_pgip.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -613,7 +613,7 @@
                 (* TODO: interim: this is probably not right.
                    What we want is mapping onto simple PGIP name/context model. *)
                 val ctx = Toplevel.context_of (Isar.state ()) (* NB: raises UNDEF *)
-                val thy = Context.theory_of_proof ctx
+                val thy = ProofContext.theory_of ctx
                 val ths = [PureThy.get_thm thy thmname]
                 val deps = #2 (thms_deps ths);
             in
--- a/src/Pure/Tools/find_theorems.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/Pure/Tools/find_theorems.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -27,6 +27,27 @@
   Name of string | Intro | IntroIff | Elim | Dest | Solves | Simp of 'term |
   Pattern of 'term;
 
+fun apply_dummies tm =
+  strip_abs tm
+  |> fst
+  |> map (Term.dummy_pattern o snd)
+  |> betapplys o pair tm
+  |> (fn x => Term.replace_dummy_patterns x 1)
+  |> fst;
+
+fun parse_pattern ctxt nm =
+  let
+    val nm' = case Syntax.parse_term ctxt nm of Const (n, _) => n | _ => nm;
+    val consts = ProofContext.consts_of ctxt;
+  in
+    nm'
+    |> Consts.intern consts
+    |> Consts.the_abbreviation consts
+    |> snd
+    |> apply_dummies
+    handle TYPE _ => ProofContext.read_term_pattern ctxt nm
+  end;
+
 fun read_criterion _ (Name name) = Name name
   | read_criterion _ Intro = Intro
   | read_criterion _ IntroIff = IntroIff
@@ -34,7 +55,7 @@
   | read_criterion _ Dest = Dest
   | read_criterion _ Solves = Solves
   | read_criterion ctxt (Simp str) = Simp (ProofContext.read_term_pattern ctxt str)
-  | read_criterion ctxt (Pattern str) = Pattern (ProofContext.read_term_pattern ctxt str);
+  | read_criterion ctxt (Pattern str) = Pattern (parse_pattern ctxt str);
 
 fun pretty_criterion ctxt (b, c) =
   let
@@ -109,7 +130,7 @@
       in Vartab.fold (fn (_, (_, t)) => fn n => size_of_term t + n) subst 0 end;
 
     fun bestmatch [] = NONE
-      | bestmatch xs = SOME (foldr1 Int.min xs);
+      | bestmatch xs = SOME (foldl1 Int.min xs);
 
     val match_thm = matches o refine_term;
   in
@@ -142,7 +163,7 @@
     (*dest rules always have assumptions, so a dest with one
       assumption is as good as an intro rule with none*)
     if not (null successful)
-    then SOME (Thm.nprems_of thm - 1, foldr1 Int.min successful) else NONE
+    then SOME (Thm.nprems_of thm - 1, foldl1 Int.min successful) else NONE
   end;
 
 fun filter_intro doiff ctxt goal (_, thm) =
@@ -173,7 +194,7 @@
         assumption is as good as an intro rule with none*)
       if is_nontrivial (ProofContext.theory_of ctxt) (Thm.major_prem_of thm)
         andalso not (null successful)
-      then SOME (Thm.nprems_of thm - 1, foldr1 Int.min successful) else NONE
+      then SOME (Thm.nprems_of thm - 1, foldl1 Int.min successful) else NONE
     end
   else NONE
 
--- a/src/Pure/context.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/Pure/context.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -4,6 +4,11 @@
 Generic theory contexts with unique identity, arbitrarily typed data,
 monotonic development graph and history support.  Generic proof
 contexts with arbitrarily typed data.
+
+Firm naming conventions:
+   thy, thy', thy1, thy2: theory
+   ctxt, ctxt', ctxt1, ctxt2: Proof.context
+   context: Context.generic
 *)
 
 signature BASIC_CONTEXT =
@@ -11,6 +16,12 @@
   type theory
   type theory_ref
   exception THEORY of string * theory list
+  structure Proof: sig type context end
+  structure ProofContext:
+  sig
+    val theory_of: Proof.context -> theory
+    val init: theory -> Proof.context
+  end
 end;
 
 signature CONTEXT =
@@ -41,25 +52,23 @@
   val finish_thy: theory -> theory
   val begin_thy: (theory -> Pretty.pp) -> string -> theory list -> theory
   (*proof context*)
-  type proof
-  val theory_of_proof: proof -> theory
-  val transfer_proof: theory -> proof -> proof
-  val init_proof: theory -> proof
+  val raw_transfer: theory -> Proof.context -> Proof.context
   (*generic context*)
-  datatype generic = Theory of theory | Proof of proof
-  val cases: (theory -> 'a) -> (proof -> 'a) -> generic -> 'a
-  val mapping: (theory -> theory) -> (proof -> proof) -> generic -> generic
-  val mapping_result: (theory -> 'a * theory) -> (proof -> 'a * proof) -> generic -> 'a * generic
+  datatype generic = Theory of theory | Proof of Proof.context
+  val cases: (theory -> 'a) -> (Proof.context -> 'a) -> generic -> 'a
+  val mapping: (theory -> theory) -> (Proof.context -> Proof.context) -> generic -> generic
+  val mapping_result: (theory -> 'a * theory) -> (Proof.context -> 'a * Proof.context) ->
+    generic -> 'a * generic
   val the_theory: generic -> theory
-  val the_proof: generic -> proof
+  val the_proof: generic -> Proof.context
   val map_theory: (theory -> theory) -> generic -> generic
-  val map_proof: (proof -> proof) -> generic -> generic
+  val map_proof: (Proof.context -> Proof.context) -> generic -> generic
   val map_theory_result: (theory -> 'a * theory) -> generic -> 'a * generic
-  val map_proof_result: (proof -> 'a * proof) -> generic -> 'a * generic
+  val map_proof_result: (Proof.context -> 'a * Proof.context) -> generic -> 'a * generic
   val theory_map: (generic -> generic) -> theory -> theory
-  val proof_map: (generic -> generic) -> proof -> proof
-  val theory_of: generic -> theory   (*total*)
-  val proof_of: generic -> proof     (*total*)
+  val proof_map: (generic -> generic) -> Proof.context -> Proof.context
+  val theory_of: generic -> theory  (*total*)
+  val proof_of: generic -> Proof.context  (*total*)
   (*thread data*)
   val thread_data: unit -> generic option
   val the_thread_data: unit -> generic
@@ -72,18 +81,18 @@
 signature PRIVATE_CONTEXT =
 sig
   include CONTEXT
-  structure TheoryData:
+  structure Theory_Data:
   sig
     val declare: Object.T -> (Object.T -> Object.T) -> (Object.T -> Object.T) ->
       (Pretty.pp -> Object.T * Object.T -> Object.T) -> serial
     val get: serial -> (Object.T -> 'a) -> theory -> 'a
     val put: serial -> ('a -> Object.T) -> 'a -> theory -> theory
   end
-  structure ProofData:
+  structure Proof_Data:
   sig
     val declare: (theory -> Object.T) -> serial
-    val get: serial -> (Object.T -> 'a) -> proof -> 'a
-    val put: serial -> ('a -> Object.T) -> 'a -> proof -> proof
+    val get: serial -> (Object.T -> 'a) -> Proof.context -> 'a
+    val put: serial -> ('a -> Object.T) -> 'a -> Proof.context -> Proof.context
   end
 end;
 
@@ -116,10 +125,10 @@
 
 in
 
-fun invoke_empty k   = invoke (K o #empty) k ();
-val invoke_copy      = invoke #copy;
-val invoke_extend    = invoke #extend;
-fun invoke_merge pp  = invoke (fn kind => #merge kind pp);
+fun invoke_empty k = invoke (K o #empty) k ();
+val invoke_copy = invoke #copy;
+val invoke_extend = invoke #extend;
+fun invoke_merge pp = invoke (fn kind => #merge kind pp);
 
 fun declare_theory_data empty copy extend merge =
   let
@@ -167,9 +176,9 @@
 fun rep_theory (Theory args) = args;
 
 val identity_of = #1 o rep_theory;
-val data_of     = #2 o rep_theory;
+val data_of = #2 o rep_theory;
 val ancestry_of = #3 o rep_theory;
-val history_of  = #4 o rep_theory;
+val history_of = #4 o rep_theory;
 
 fun make_identity self draft id ids = {self = self, draft = draft, id = id, ids = ids};
 fun make_ancestry parents ancestors = {parents = parents, ancestors = ancestors};
@@ -204,7 +213,8 @@
 val is_draft = #draft o identity_of;
 
 fun reject_draft thy =
-  if is_draft thy then raise THEORY ("Illegal draft theory -- stable checkpoint required", [thy])
+  if is_draft thy then
+    raise THEORY ("Illegal draft theory -- stable checkpoint required", [thy])
   else thy;
 
 
@@ -244,12 +254,12 @@
   theory in external data structures -- a plain theory value would
   become stale as the self reference moves on*)
 
-datatype theory_ref = TheoryRef of theory Unsynchronized.ref;
+datatype theory_ref = Theory_Ref of theory Unsynchronized.ref;
 
-fun deref (TheoryRef (Unsynchronized.ref thy)) = thy;
+fun deref (Theory_Ref (Unsynchronized.ref thy)) = thy;
 
 fun check_thy thy =  (*thread-safe version*)
-  let val thy_ref = TheoryRef (the_self thy) in
+  let val thy_ref = Theory_Ref (the_self thy) in
     if is_stale thy then error ("Stale theory encountered:\n" ^ string_of_thy thy)
     else thy_ref
   end;
@@ -284,7 +294,8 @@
 (* consistent ancestors *)
 
 fun extend_ancestors thy thys =
-  if member eq_thy thys thy then raise THEORY ("Duplicate theory node", thy :: thys)
+  if member eq_thy thys thy then
+    raise THEORY ("Duplicate theory node", thy :: thys)
   else thy :: thys;
 
 fun extend_ancestors_of thy = extend_ancestors thy (ancestors_of thy);
@@ -407,7 +418,7 @@
 
 (* theory data *)
 
-structure TheoryData =
+structure Theory_Data =
 struct
 
 val declare = declare_theory_data;
@@ -425,13 +436,16 @@
 
 (*** proof context ***)
 
-(* datatype proof *)
-
-datatype proof = Prf of Object.T Datatab.table * theory_ref;
+(* datatype Proof.context *)
 
-fun theory_of_proof (Prf (_, thy_ref)) = deref thy_ref;
-fun data_of_proof (Prf (data, _)) = data;
-fun map_prf f (Prf (data, thy_ref)) = Prf (f data, thy_ref);
+structure Proof =
+struct
+  datatype context = Context of Object.T Datatab.table * theory_ref;
+end;
+
+fun theory_of_proof (Proof.Context (_, thy_ref)) = deref thy_ref;
+fun data_of_proof (Proof.Context (data, _)) = data;
+fun map_prf f (Proof.Context (data, thy_ref)) = Proof.Context (f data, thy_ref);
 
 
 (* proof data kinds *)
@@ -453,19 +467,22 @@
 
 in
 
-fun init_proof thy = Prf (init_data thy, check_thy thy);
-
-fun transfer_proof thy' (Prf (data, thy_ref)) =
+fun raw_transfer thy' (Proof.Context (data, thy_ref)) =
   let
     val thy = deref thy_ref;
     val _ = subthy (thy, thy') orelse error "transfer proof context: not a super theory";
     val _ = check_thy thy;
     val data' = init_new_data data thy';
     val thy_ref' = check_thy thy';
-  in Prf (data', thy_ref') end;
+  in Proof.Context (data', thy_ref') end;
 
+structure ProofContext =
+struct
+  val theory_of = theory_of_proof;
+  fun init thy = Proof.Context (init_data thy, check_thy thy);
+end;
 
-structure ProofData =
+structure Proof_Data =
 struct
 
 fun declare init =
@@ -477,7 +494,7 @@
 fun get k dest prf =
   dest (case Datatab.lookup (data_of_proof prf) k of
     SOME x => x
-  | NONE => invoke_init k (theory_of_proof prf));   (*adhoc value*)
+  | NONE => invoke_init k (ProofContext.theory_of prf));   (*adhoc value*)
 
 fun put k mk x = map_prf (Datatab.update (k, mk x));
 
@@ -489,7 +506,7 @@
 
 (*** generic context ***)
 
-datatype generic = Theory of theory | Proof of proof;
+datatype generic = Theory of theory | Proof of Proof.context;
 
 fun cases f _ (Theory thy) = f thy
   | cases _ g (Proof prf) = g prf;
@@ -509,8 +526,8 @@
 fun theory_map f = the_theory o f o Theory;
 fun proof_map f = the_proof o f o Proof;
 
-val theory_of = cases I theory_of_proof;
-val proof_of = cases init_proof I;
+val theory_of = cases I ProofContext.theory_of;
+val proof_of = cases ProofContext.init I;
 
 
 
@@ -546,8 +563,8 @@
 
 end;
 
-structure BasicContext: BASIC_CONTEXT = Context;
-open BasicContext;
+structure Basic_Context: BASIC_CONTEXT = Context;
+open Basic_Context;
 
 
 
@@ -576,19 +593,17 @@
 functor TheoryDataFun(Data: THEORY_DATA_ARGS): THEORY_DATA =
 struct
 
-structure TheoryData = Context.TheoryData;
-
 type T = Data.T;
 exception Data of T;
 
-val kind = TheoryData.declare
+val kind = Context.Theory_Data.declare
   (Data Data.empty)
   (fn Data x => Data (Data.copy x))
   (fn Data x => Data (Data.extend x))
   (fn pp => fn (Data x1, Data x2) => Data (Data.merge pp (x1, x2)));
 
-val get = TheoryData.get kind (fn Data x => x);
-val put = TheoryData.put kind Data;
+val get = Context.Theory_Data.get kind (fn Data x => x);
+val put = Context.Theory_Data.put kind Data;
 fun map f thy = put (f (get thy)) thy;
 
 fun init thy = map I thy;
@@ -608,23 +623,21 @@
 signature PROOF_DATA =
 sig
   type T
-  val get: Context.proof -> T
-  val put: T -> Context.proof -> Context.proof
-  val map: (T -> T) -> Context.proof -> Context.proof
+  val get: Proof.context -> T
+  val put: T -> Proof.context -> Proof.context
+  val map: (T -> T) -> Proof.context -> Proof.context
 end;
 
 functor ProofDataFun(Data: PROOF_DATA_ARGS): PROOF_DATA =
 struct
 
-structure ProofData = Context.ProofData;
-
 type T = Data.T;
 exception Data of T;
 
-val kind = ProofData.declare (Data o Data.init);
+val kind = Context.Proof_Data.declare (Data o Data.init);
 
-val get = ProofData.get kind (fn Data x => x);
-val put = ProofData.put kind Data;
+val get = Context.Proof_Data.get kind (fn Data x => x);
+val put = Context.Proof_Data.put kind Data;
 fun map f prf = put (f (get prf)) prf;
 
 end;
@@ -652,16 +665,16 @@
 functor GenericDataFun(Data: GENERIC_DATA_ARGS): GENERIC_DATA =
 struct
 
-structure ThyData = TheoryDataFun(open Data val copy = I);
-structure PrfData = ProofDataFun(type T = Data.T val init = ThyData.get);
+structure Thy_Data = TheoryDataFun(open Data val copy = I);
+structure Prf_Data = ProofDataFun(type T = Data.T val init = Thy_Data.get);
 
 type T = Data.T;
 
-fun get (Context.Theory thy) = ThyData.get thy
-  | get (Context.Proof prf) = PrfData.get prf;
+fun get (Context.Theory thy) = Thy_Data.get thy
+  | get (Context.Proof prf) = Prf_Data.get prf;
 
-fun put x (Context.Theory thy) = Context.Theory (ThyData.put x thy)
-  | put x (Context.Proof prf) = Context.Proof (PrfData.put x prf);
+fun put x (Context.Theory thy) = Context.Theory (Thy_Data.put x thy)
+  | put x (Context.Proof prf) = Context.Proof (Prf_Data.put x prf);
 
 fun map f ctxt = put (f (get ctxt)) ctxt;
 
@@ -670,7 +683,3 @@
 (*hide private interface*)
 structure Context: CONTEXT = Context;
 
-(*fake predeclarations*)
-structure Proof = struct type context = Context.proof end;
-structure ProofContext =
-struct val theory_of = Context.theory_of_proof val init = Context.init_proof end;
--- a/src/Pure/meta_simplifier.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/Pure/meta_simplifier.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -346,7 +346,7 @@
 fun activate_context thy ss =
   let
     val ctxt = the_context ss;
-    val ctxt' = Context.transfer_proof (Theory.merge (thy, ProofContext.theory_of ctxt)) ctxt;
+    val ctxt' = Context.raw_transfer (Theory.merge (thy, ProofContext.theory_of ctxt)) ctxt;
   in context ctxt' ss end;
 
 
@@ -1158,8 +1158,8 @@
             let
               val prem' = Thm.rhs_of eqn;
               val tprems = map term_of prems;
-              val i = 1 + Library.foldl Int.max (~1, map (fn p =>
-                find_index (fn q => q aconv p) tprems) (#hyps (rep_thm eqn)));
+              val i = 1 + fold Integer.max (map (fn p =>
+                find_index (fn q => q aconv p) tprems) (#hyps (rep_thm eqn))) ~1;
               val (rrs', asm') = rules_of_prem ss prem'
             in mut_impc prems concl rrss asms (prem' :: prems')
               (rrs' :: rrss') (asm' :: asms') (SOME (fold_rev (disch true)
--- a/src/Pure/pure_setup.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/Pure/pure_setup.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -28,7 +28,7 @@
 toplevel_pp ["typ"] "ProofDisplay.pp_typ Pure.thy";
 toplevel_pp ["Context", "theory"] "Context.pretty_thy";
 toplevel_pp ["Context", "theory_ref"] "Context.pretty_thy o Theory.deref";
-toplevel_pp ["Context", "proof"] "ProofDisplay.pp_context";
+toplevel_pp ["Context", "Proof", "context"] "ProofDisplay.pp_context";
 toplevel_pp ["Syntax", "ast"] "Syntax.pretty_ast";
 toplevel_pp ["Path", "T"] "Pretty.str o quote o Path.implode";
 toplevel_pp ["File", "ident"] "Pretty.str o quote o File.rep_ident";
@@ -42,10 +42,10 @@
 
 (* ML toplevel use commands *)
 
-fun use name          = Toplevel.program (fn () => ThyInfo.use name);
-fun use_thys name     = Toplevel.program (fn () => ThyInfo.use_thys name);
-fun use_thy name      = Toplevel.program (fn () => ThyInfo.use_thy name);
-fun time_use name     = Toplevel.program (fn () => ThyInfo.time_use name);
+fun use name = Toplevel.program (fn () => ThyInfo.use name);
+fun use_thys name = Toplevel.program (fn () => ThyInfo.use_thys name);
+fun use_thy name = Toplevel.program (fn () => ThyInfo.use_thy name);
+fun time_use name = Toplevel.program (fn () => ThyInfo.time_use name);
 fun time_use_thy name = Toplevel.program (fn () => ThyInfo.time_use_thy name);
 
 
--- a/src/Tools/Code/code_thingol.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/Tools/Code/code_thingol.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -420,7 +420,7 @@
 fun same_arity thy thms =
   let
     val num_args_of = length o snd o strip_comb o fst o Logic.dest_equals;
-    val k = fold (curry Int.max o num_args_of o Thm.prop_of) thms 0;
+    val k = fold (Integer.max o num_args_of o Thm.prop_of) thms 0;
   in map (expand_eta thy k) thms end;
 
 fun mk_desymbolization pre post mk vs =
--- a/src/Tools/atomize_elim.ML	Wed Oct 21 08:14:38 2009 +0200
+++ b/src/Tools/atomize_elim.ML	Wed Oct 21 08:16:25 2009 +0200
@@ -34,7 +34,7 @@
 
 (* Compute inverse permutation *)
 fun invert_perm pi =
-      (pi @ ((0 upto (fold (curry Int.max) pi 0)) \\ pi))
+      (pi @ ((0 upto (fold Integer.max pi 0)) \\ pi))
            |> map_index I
            |> sort (int_ord o pairself snd)
            |> map fst