move some example files into new HOLCF/Tutorial directory
authorhuffman
Wed, 19 May 2010 17:01:07 -0700
changeset 37000 41a22e7c1145
parent 36999 5d9091ba3128
child 37010 8096a4c755eb
move some example files into new HOLCF/Tutorial directory
src/HOLCF/IsaMakefile
src/HOLCF/Tutorial/Domain_ex.thy
src/HOLCF/Tutorial/Fixrec_ex.thy
src/HOLCF/Tutorial/New_Domain.thy
src/HOLCF/Tutorial/ROOT.ML
src/HOLCF/Tutorial/document/root.tex
src/HOLCF/ex/Domain_ex.thy
src/HOLCF/ex/Fixrec_ex.thy
src/HOLCF/ex/New_Domain.thy
src/HOLCF/ex/ROOT.ML
--- a/src/HOLCF/IsaMakefile	Wed May 19 16:28:24 2010 -0700
+++ b/src/HOLCF/IsaMakefile	Wed May 19 17:01:07 2010 -0700
@@ -7,6 +7,7 @@
 default: HOLCF
 images: HOLCF IOA
 test: HOLCF-IMP HOLCF-ex HOLCF-FOCUS \
+  HOLCF-Tutorial \
       IOA-ABP IOA-NTP IOA-Modelcheck IOA-Storage IOA-ex
 all: images test
 
@@ -78,6 +79,19 @@
 	@$(ISABELLE_TOOL) usedir -b -g true -r $(OUT)/HOL HOLCF
 
 
+## HOLCF-Tutorial
+
+HOLCF-Tutorial: HOLCF $(LOG)/HOLCF-Tutorial.gz
+
+$(LOG)/HOLCF-Tutorial.gz: $(OUT)/HOLCF \
+  Tutorial/Domain_ex.thy \
+  Tutorial/Fixrec_ex.thy \
+  Tutorial/New_Domain.thy \
+  Tutorial/document/root.tex \
+  Tutorial/ROOT.ML
+	@$(ISABELLE_TOOL) usedir $(OUT)/HOLCF Tutorial
+
+
 ## HOLCF-IMP
 
 HOLCF-IMP: HOLCF $(LOG)/HOLCF-IMP.gz
@@ -95,15 +109,12 @@
   ../HOL/Library/Nat_Infinity.thy \
   ex/Dagstuhl.thy \
   ex/Dnat.thy \
-  ex/Domain_ex.thy \
   ex/Domain_Proofs.thy \
   ex/Fix2.thy \
-  ex/Fixrec_ex.thy \
   ex/Focus_ex.thy \
   ex/Hoare.thy \
   ex/Letrec.thy \
   ex/Loop.thy \
-  ex/New_Domain.thy \
   ex/Powerdomain_ex.thy \
   ex/Stream.thy \
   ex/Strict_Fun.thy \
@@ -201,4 +212,4 @@
 	  $(LOG)/HOLCF-ex.gz $(LOG)/HOLCF-FOCUS.gz $(OUT)/IOA	\
 	  $(LOG)/IOA.gz $(LOG)/IOA-ABP.gz $(LOG)/IOA-NTP.gz	\
 	  $(LOG)/IOA-Modelcheck.gz $(LOG)/IOA-Storage.gz	\
-	  $(LOG)/IOA-ex.gz
+	  $(LOG)/IOA-ex.gz $(LOG)/HOLCF-Tutorial.gz
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/Tutorial/Domain_ex.thy	Wed May 19 17:01:07 2010 -0700
@@ -0,0 +1,211 @@
+(*  Title:      HOLCF/ex/Domain_ex.thy
+    Author:     Brian Huffman
+*)
+
+header {* Domain package examples *}
+
+theory Domain_ex
+imports HOLCF
+begin
+
+text {* Domain constructors are strict by default. *}
+
+domain d1 = d1a | d1b "d1" "d1"
+
+lemma "d1b\<cdot>\<bottom>\<cdot>y = \<bottom>" by simp
+
+text {* Constructors can be made lazy using the @{text "lazy"} keyword. *}
+
+domain d2 = d2a | d2b (lazy "d2")
+
+lemma "d2b\<cdot>x \<noteq> \<bottom>" by simp
+
+text {* Strict and lazy arguments may be mixed arbitrarily. *}
+
+domain d3 = d3a | d3b (lazy "d2") "d2"
+
+lemma "P (d3b\<cdot>x\<cdot>y = \<bottom>) \<longleftrightarrow> P (y = \<bottom>)" by simp
+
+text {* Selectors can be used with strict or lazy constructor arguments. *}
+
+domain d4 = d4a | d4b (lazy d4b_left :: "d2") (d4b_right :: "d2")
+
+lemma "y \<noteq> \<bottom> \<Longrightarrow> d4b_left\<cdot>(d4b\<cdot>x\<cdot>y) = x" by simp
+
+text {* Mixfix declarations can be given for data constructors. *}
+
+domain d5 = d5a | d5b (lazy "d5") "d5" (infixl ":#:" 70)
+
+lemma "d5a \<noteq> x :#: y :#: z" by simp
+
+text {* Mixfix declarations can also be given for type constructors. *}
+
+domain ('a, 'b) lazypair (infixl ":*:" 25) =
+  lpair (lazy lfst :: 'a) (lazy lsnd :: 'b) (infixl ":*:" 75)
+
+lemma "\<forall>p::('a :*: 'b). p \<sqsubseteq> lfst\<cdot>p :*: lsnd\<cdot>p"
+by (rule allI, case_tac p, simp_all)
+
+text {* Non-recursive constructor arguments can have arbitrary types. *}
+
+domain ('a, 'b) d6 = d6 "int lift" "'a \<oplus> 'b u" (lazy "('a :*: 'b) \<times> ('b \<rightarrow> 'a)")
+
+text {*
+  Indirect recusion is allowed for sums, products, lifting, and the
+  continuous function space.  However, the domain package does not
+  generate an induction rule in terms of the constructors.
+*}
+
+domain 'a d7 = d7a "'a d7 \<oplus> int lift" | d7b "'a \<otimes> 'a d7" | d7c (lazy "'a d7 \<rightarrow> 'a")
+  -- "Indirect recursion detected, skipping proofs of (co)induction rules"
+
+text {* Note that @{text d7.induct} is absent. *}
+
+text {*
+  Indirect recursion is also allowed using previously-defined datatypes.
+*}
+
+domain 'a slist = SNil | SCons 'a "'a slist"
+
+domain 'a stree = STip | SBranch "'a stree slist"
+
+text {* Mutually-recursive datatypes can be defined using the @{text "and"} keyword. *}
+
+domain d8 = d8a | d8b "d9" and d9 = d9a | d9b (lazy "d8")
+
+text {* Non-regular recursion is not allowed. *}
+(*
+domain ('a, 'b) altlist = ANil | ACons 'a "('b, 'a) altlist"
+  -- "illegal direct recursion with different arguments"
+domain 'a nest = Nest1 'a | Nest2 "'a nest nest"
+  -- "illegal direct recursion with different arguments"
+*)
+
+text {*
+  Mutually-recursive datatypes must have all the same type arguments,
+  not necessarily in the same order.
+*}
+
+domain ('a, 'b) list1 = Nil1 | Cons1 'a "('b, 'a) list2"
+   and ('b, 'a) list2 = Nil2 | Cons2 'b "('a, 'b) list1"
+
+text {* Induction rules for flat datatypes have no admissibility side-condition. *}
+
+domain 'a flattree = Tip | Branch "'a flattree" "'a flattree"
+
+lemma "\<lbrakk>P \<bottom>; P Tip; \<And>x y. \<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>; P x; P y\<rbrakk> \<Longrightarrow> P (Branch\<cdot>x\<cdot>y)\<rbrakk> \<Longrightarrow> P x"
+by (rule flattree.induct) -- "no admissibility requirement"
+
+text {* Trivial datatypes will produce a warning message. *}
+
+domain triv = Triv triv triv
+  -- "domain @{text Domain_ex.triv} is empty!"
+
+lemma "(x::triv) = \<bottom>" by (induct x, simp_all)
+
+text {* Lazy constructor arguments may have unpointed types. *}
+
+domain natlist = nnil | ncons (lazy "nat discr") natlist
+
+text {* Class constraints may be given for type parameters on the LHS. *}
+
+domain ('a::cpo) box = Box (lazy 'a)
+
+
+subsection {* Generated constants and theorems *}
+
+domain 'a tree = Leaf (lazy 'a) | Node (left :: "'a tree") (right :: "'a tree")
+
+lemmas tree_abs_defined_iff =
+  iso.abs_defined_iff [OF iso.intro [OF tree.abs_iso tree.rep_iso]]
+
+text {* Rules about ismorphism *}
+term tree_rep
+term tree_abs
+thm tree.rep_iso
+thm tree.abs_iso
+thm tree.iso_rews
+
+text {* Rules about constructors *}
+term Leaf
+term Node
+thm Leaf_def Node_def
+thm tree.nchotomy
+thm tree.exhaust
+thm tree.compacts
+thm tree.con_rews
+thm tree.dist_les
+thm tree.dist_eqs
+thm tree.inverts
+thm tree.injects
+
+text {* Rules about case combinator *}
+term tree_when
+thm tree.tree_when_def
+thm tree.when_rews
+
+text {* Rules about selectors *}
+term left
+term right
+thm tree.sel_rews
+
+text {* Rules about discriminators *}
+term is_Leaf
+term is_Node
+thm tree.dis_rews
+
+text {* Rules about pattern match combinators *}
+term Leaf_pat
+term Node_pat
+thm tree.pat_rews
+
+text {* Rules about monadic pattern match combinators *}
+term match_Leaf
+term match_Node
+thm tree.match_rews
+
+text {* Rules about take function *}
+term tree_take
+thm tree.take_def
+thm tree.take_0
+thm tree.take_Suc
+thm tree.take_rews
+thm tree.chain_take
+thm tree.take_take
+thm tree.deflation_take
+thm tree.take_below
+thm tree.take_lemma
+thm tree.lub_take
+thm tree.reach
+thm tree.finite_induct
+
+text {* Rules about finiteness predicate *}
+term tree_finite
+thm tree.finite_def
+thm tree.finite (* only generated for flat datatypes *)
+
+text {* Rules about bisimulation predicate *}
+term tree_bisim
+thm tree.bisim_def
+thm tree.coinduct
+
+text {* Induction rule *}
+thm tree.induct
+
+
+subsection {* Known bugs *}
+
+text {* Declaring a mixfix with spaces causes some strange parse errors. *}
+(*
+domain xx = xx ("x y")
+  -- "Inner syntax error: unexpected end of input"
+*)
+
+text {*
+  Non-cpo type parameters currently do not work.
+*}
+(*
+domain ('a::type) stream = snil | scons (lazy "'a discr") "'a stream"
+*)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/Tutorial/Fixrec_ex.thy	Wed May 19 17:01:07 2010 -0700
@@ -0,0 +1,246 @@
+(*  Title:      HOLCF/ex/Fixrec_ex.thy
+    Author:     Brian Huffman
+*)
+
+header {* Fixrec package examples *}
+
+theory Fixrec_ex
+imports HOLCF
+begin
+
+subsection {* Basic @{text fixrec} examples *}
+
+text {*
+  Fixrec patterns can mention any constructor defined by the domain
+  package, as well as any of the following built-in constructors:
+  Pair, spair, sinl, sinr, up, ONE, TT, FF.
+*}
+
+text {* Typical usage is with lazy constructors. *}
+
+fixrec down :: "'a u \<rightarrow> 'a"
+where "down\<cdot>(up\<cdot>x) = x"
+
+text {* With strict constructors, rewrite rules may require side conditions. *}
+
+fixrec from_sinl :: "'a \<oplus> 'b \<rightarrow> 'a"
+where "x \<noteq> \<bottom> \<Longrightarrow> from_sinl\<cdot>(sinl\<cdot>x) = x"
+
+text {* Lifting can turn a strict constructor into a lazy one. *}
+
+fixrec from_sinl_up :: "'a u \<oplus> 'b \<rightarrow> 'a"
+where "from_sinl_up\<cdot>(sinl\<cdot>(up\<cdot>x)) = x"
+
+text {* Fixrec also works with the HOL pair constructor. *}
+
+fixrec down2 :: "'a u \<times> 'b u \<rightarrow> 'a \<times> 'b"
+where "down2\<cdot>(up\<cdot>x, up\<cdot>y) = (x, y)"
+
+
+subsection {* Examples using @{text fixrec_simp} *}
+
+text {* A type of lazy lists. *}
+
+domain 'a llist = lNil | lCons (lazy 'a) (lazy "'a llist")
+
+text {* A zip function for lazy lists. *}
+
+text {* Notice that the patterns are not exhaustive. *}
+
+fixrec
+  lzip :: "'a llist \<rightarrow> 'b llist \<rightarrow> ('a \<times> 'b) llist"
+where
+  "lzip\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>(lCons\<cdot>y\<cdot>ys) = lCons\<cdot>(x, y)\<cdot>(lzip\<cdot>xs\<cdot>ys)"
+| "lzip\<cdot>lNil\<cdot>lNil = lNil"
+
+text {* @{text fixrec_simp} is useful for producing strictness theorems. *}
+text {* Note that pattern matching is done in left-to-right order. *}
+
+lemma lzip_stricts [simp]:
+  "lzip\<cdot>\<bottom>\<cdot>ys = \<bottom>"
+  "lzip\<cdot>lNil\<cdot>\<bottom> = \<bottom>"
+  "lzip\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>\<bottom> = \<bottom>"
+by fixrec_simp+
+
+text {* @{text fixrec_simp} can also produce rules for missing cases. *}
+
+lemma lzip_undefs [simp]:
+  "lzip\<cdot>lNil\<cdot>(lCons\<cdot>y\<cdot>ys) = \<bottom>"
+  "lzip\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>lNil = \<bottom>"
+by fixrec_simp+
+
+
+subsection {* Pattern matching with bottoms *}
+
+text {*
+  As an alternative to using @{text fixrec_simp}, it is also possible
+  to use bottom as a constructor pattern.  When using a bottom
+  pattern, the right-hand-side must also be bottom; otherwise, @{text
+  fixrec} will not be able to prove the equation.
+*}
+
+fixrec
+  from_sinr_up :: "'a \<oplus> 'b\<^sub>\<bottom> \<rightarrow> 'b"
+where
+  "from_sinr_up\<cdot>\<bottom> = \<bottom>"
+| "from_sinr_up\<cdot>(sinr\<cdot>(up\<cdot>x)) = x"
+
+text {*
+  If the function is already strict in that argument, then the bottom
+  pattern does not change the meaning of the function.  For example,
+  in the definition of @{term from_sinr_up}, the first equation is
+  actually redundant, and could have been proven separately by
+  @{text fixrec_simp}.
+*}
+
+text {*
+  A bottom pattern can also be used to make a function strict in a
+  certain argument, similar to a bang-pattern in Haskell.
+*}
+
+fixrec
+  seq :: "'a \<rightarrow> 'b \<rightarrow> 'b"
+where
+  "seq\<cdot>\<bottom>\<cdot>y = \<bottom>"
+| "x \<noteq> \<bottom> \<Longrightarrow> seq\<cdot>x\<cdot>y = y"
+
+
+subsection {* Skipping proofs of rewrite rules *}
+
+text {* Another zip function for lazy lists. *}
+
+text {*
+  Notice that this version has overlapping patterns.
+  The second equation cannot be proved as a theorem
+  because it only applies when the first pattern fails.
+*}
+
+fixrec (permissive)
+  lzip2 :: "'a llist \<rightarrow> 'b llist \<rightarrow> ('a \<times> 'b) llist"
+where
+  "lzip2\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>(lCons\<cdot>y\<cdot>ys) = lCons\<cdot>(x, y)\<cdot>(lzip\<cdot>xs\<cdot>ys)"
+| "lzip2\<cdot>xs\<cdot>ys = lNil"
+
+text {*
+  Usually fixrec tries to prove all equations as theorems.
+  The "permissive" option overrides this behavior, so fixrec
+  does not produce any simp rules.
+*}
+
+text {* Simp rules can be generated later using @{text fixrec_simp}. *}
+
+lemma lzip2_simps [simp]:
+  "lzip2\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>(lCons\<cdot>y\<cdot>ys) = lCons\<cdot>(x, y)\<cdot>(lzip\<cdot>xs\<cdot>ys)"
+  "lzip2\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>lNil = lNil"
+  "lzip2\<cdot>lNil\<cdot>(lCons\<cdot>y\<cdot>ys) = lNil"
+  "lzip2\<cdot>lNil\<cdot>lNil = lNil"
+by fixrec_simp+
+
+lemma lzip2_stricts [simp]:
+  "lzip2\<cdot>\<bottom>\<cdot>ys = \<bottom>"
+  "lzip2\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>\<bottom> = \<bottom>"
+by fixrec_simp+
+
+
+subsection {* Mutual recursion with @{text fixrec} *}
+
+text {* Tree and forest types. *}
+
+domain 'a tree = Leaf (lazy 'a) | Branch (lazy "'a forest")
+and    'a forest = Empty | Trees (lazy "'a tree") "'a forest"
+
+text {*
+  To define mutually recursive functions, give multiple type signatures
+  separated by the keyword @{text "and"}.
+*}
+
+fixrec
+  map_tree :: "('a \<rightarrow> 'b) \<rightarrow> ('a tree \<rightarrow> 'b tree)"
+and
+  map_forest :: "('a \<rightarrow> 'b) \<rightarrow> ('a forest \<rightarrow> 'b forest)"
+where
+  "map_tree\<cdot>f\<cdot>(Leaf\<cdot>x) = Leaf\<cdot>(f\<cdot>x)"
+| "map_tree\<cdot>f\<cdot>(Branch\<cdot>ts) = Branch\<cdot>(map_forest\<cdot>f\<cdot>ts)"
+| "map_forest\<cdot>f\<cdot>Empty = Empty"
+| "ts \<noteq> \<bottom> \<Longrightarrow>
+    map_forest\<cdot>f\<cdot>(Trees\<cdot>t\<cdot>ts) = Trees\<cdot>(map_tree\<cdot>f\<cdot>t)\<cdot>(map_forest\<cdot>f\<cdot>ts)"
+
+lemma map_tree_strict [simp]: "map_tree\<cdot>f\<cdot>\<bottom> = \<bottom>"
+by fixrec_simp
+
+lemma map_forest_strict [simp]: "map_forest\<cdot>f\<cdot>\<bottom> = \<bottom>"
+by fixrec_simp
+
+(*
+  Theorems generated:
+  @{text map_tree_def}  @{thm map_tree_def}
+  @{text map_forest_def}  @{thm map_forest_def}
+  @{text map_tree.unfold}  @{thm map_tree.unfold}
+  @{text map_forest.unfold}  @{thm map_forest.unfold}
+  @{text map_tree.simps}  @{thm map_tree.simps}
+  @{text map_forest.simps}  @{thm map_forest.simps}
+  @{text map_tree_map_forest.induct}  @{thm map_tree_map_forest.induct}
+*)
+
+
+subsection {* Looping simp rules *}
+
+text {*
+  The defining equations of a fixrec definition are declared as simp
+  rules by default.  In some cases, especially for constants with no
+  arguments or functions with variable patterns, the defining
+  equations may cause the simplifier to loop.  In these cases it will
+  be necessary to use a @{text "[simp del]"} declaration.
+*}
+
+fixrec
+  repeat :: "'a \<rightarrow> 'a llist"
+where
+  [simp del]: "repeat\<cdot>x = lCons\<cdot>x\<cdot>(repeat\<cdot>x)"
+
+text {*
+  We can derive other non-looping simp rules for @{const repeat} by
+  using the @{text subst} method with the @{text repeat.simps} rule.
+*}
+
+lemma repeat_simps [simp]:
+  "repeat\<cdot>x \<noteq> \<bottom>"
+  "repeat\<cdot>x \<noteq> lNil"
+  "repeat\<cdot>x = lCons\<cdot>y\<cdot>ys \<longleftrightarrow> x = y \<and> repeat\<cdot>x = ys"
+by (subst repeat.simps, simp)+
+
+lemma llist_when_repeat [simp]:
+  "llist_when\<cdot>z\<cdot>f\<cdot>(repeat\<cdot>x) = f\<cdot>x\<cdot>(repeat\<cdot>x)"
+by (subst repeat.simps, simp)
+
+text {*
+  For mutually-recursive constants, looping might only occur if all
+  equations are in the simpset at the same time.  In such cases it may
+  only be necessary to declare @{text "[simp del]"} on one equation.
+*}
+
+fixrec
+  inf_tree :: "'a tree" and inf_forest :: "'a forest"
+where
+  [simp del]: "inf_tree = Branch\<cdot>inf_forest"
+| "inf_forest = Trees\<cdot>inf_tree\<cdot>(Trees\<cdot>inf_tree\<cdot>Empty)"
+
+
+subsection {* Using @{text fixrec} inside locales *}
+
+locale test =
+  fixes foo :: "'a \<rightarrow> 'a"
+  assumes foo_strict: "foo\<cdot>\<bottom> = \<bottom>"
+begin
+
+fixrec
+  bar :: "'a u \<rightarrow> 'a"
+where
+  "bar\<cdot>(up\<cdot>x) = foo\<cdot>x"
+
+lemma bar_strict: "bar\<cdot>\<bottom> = \<bottom>"
+by fixrec_simp
+
+end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/Tutorial/New_Domain.thy	Wed May 19 17:01:07 2010 -0700
@@ -0,0 +1,92 @@
+(*  Title:      HOLCF/ex/New_Domain.thy
+    Author:     Brian Huffman
+*)
+
+header {* Definitional domain package *}
+
+theory New_Domain
+imports HOLCF
+begin
+
+text {*
+  The definitional domain package only works with representable domains,
+  i.e. types in class @{text rep}.
+*}
+
+default_sort rep
+
+text {*
+  Provided that @{text rep} is the default sort, the @{text new_domain}
+  package should work with any type definition supported by the old
+  domain package.
+*}
+
+new_domain 'a llist = LNil | LCons (lazy 'a) (lazy "'a llist")
+
+text {*
+  The difference is that the new domain package is completely
+  definitional, and does not generate any axioms.  The following type
+  and constant definitions are not produced by the old domain package.
+*}
+
+thm type_definition_llist
+thm llist_abs_def llist_rep_def
+
+text {*
+  The new domain package also adds support for indirect recursion with
+  user-defined datatypes.  This definition of a tree datatype uses
+  indirect recursion through the lazy list type constructor.
+*}
+
+new_domain 'a ltree = Leaf (lazy 'a) | Branch (lazy "'a ltree llist")
+
+text {*
+  For indirect-recursive definitions, the domain package is not able to
+  generate a high-level induction rule.  (It produces a warning
+  message instead.)  The low-level reach lemma (now proved as a
+  theorem, no longer generated as an axiom) can be used to derive
+  other induction rules.
+*}
+
+thm ltree.reach
+
+text {*
+  The definition of the take function uses map functions associated with
+  each type constructor involved in the definition.  A map function
+  for the lazy list type has been generated by the new domain package.
+*}
+
+thm ltree.take_rews
+thm llist_map_def
+
+lemma ltree_induct:
+  fixes P :: "'a ltree \<Rightarrow> bool"
+  assumes adm: "adm P"
+  assumes bot: "P \<bottom>"
+  assumes Leaf: "\<And>x. P (Leaf\<cdot>x)"
+  assumes Branch: "\<And>f l. \<forall>x. P (f\<cdot>x) \<Longrightarrow> P (Branch\<cdot>(llist_map\<cdot>f\<cdot>l))"
+  shows "P x"
+proof -
+  have "P (\<Squnion>i. ltree_take i\<cdot>x)"
+  using adm
+  proof (rule admD)
+    fix i
+    show "P (ltree_take i\<cdot>x)"
+    proof (induct i arbitrary: x)
+      case (0 x)
+      show "P (ltree_take 0\<cdot>x)" by (simp add: bot)
+    next
+      case (Suc n x)
+      show "P (ltree_take (Suc n)\<cdot>x)"
+        apply (cases x)
+        apply (simp add: bot)
+        apply (simp add: Leaf)
+        apply (simp add: Branch Suc)
+        done
+    qed
+  qed (simp add: ltree.chain_take)
+  thus ?thesis
+    by (simp add: ltree.reach)
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/Tutorial/ROOT.ML	Wed May 19 17:01:07 2010 -0700
@@ -0,0 +1,1 @@
+use_thys ["Domain_ex", "Fixrec_ex", "New_Domain"];
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/Tutorial/document/root.tex	Wed May 19 17:01:07 2010 -0700
@@ -0,0 +1,29 @@
+
+% HOLCF/document/root.tex
+
+\documentclass[11pt,a4paper]{article}
+\usepackage{graphicx,isabelle,isabellesym,latexsym}
+\usepackage[only,bigsqcap]{stmaryrd}
+\usepackage[latin1]{inputenc}
+\usepackage{pdfsetup}
+
+\urlstyle{rm}
+%\isabellestyle{it}
+\pagestyle{myheadings}
+
+\begin{document}
+
+\title{Isabelle/HOLCF Tutorial}
+\maketitle
+
+\tableofcontents
+
+%\newpage
+
+%\renewcommand{\isamarkupheader}[1]%
+%{\section{\isabellecontext: #1}\markright{THEORY~``\isabellecontext''}}
+
+\parindent 0pt\parskip 0.5ex
+\input{session}
+
+\end{document}
--- a/src/HOLCF/ex/Domain_ex.thy	Wed May 19 16:28:24 2010 -0700
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,210 +0,0 @@
-(*  Title:      HOLCF/ex/Domain_ex.thy
-    Author:     Brian Huffman
-*)
-
-header {* Domain package examples *}
-
-theory Domain_ex
-imports HOLCF
-begin
-
-text {* Domain constructors are strict by default. *}
-
-domain d1 = d1a | d1b "d1" "d1"
-
-lemma "d1b\<cdot>\<bottom>\<cdot>y = \<bottom>" by simp
-
-text {* Constructors can be made lazy using the @{text "lazy"} keyword. *}
-
-domain d2 = d2a | d2b (lazy "d2")
-
-lemma "d2b\<cdot>x \<noteq> \<bottom>" by simp
-
-text {* Strict and lazy arguments may be mixed arbitrarily. *}
-
-domain d3 = d3a | d3b (lazy "d2") "d2"
-
-lemma "P (d3b\<cdot>x\<cdot>y = \<bottom>) \<longleftrightarrow> P (y = \<bottom>)" by simp
-
-text {* Selectors can be used with strict or lazy constructor arguments. *}
-
-domain d4 = d4a | d4b (lazy d4b_left :: "d2") (d4b_right :: "d2")
-
-lemma "y \<noteq> \<bottom> \<Longrightarrow> d4b_left\<cdot>(d4b\<cdot>x\<cdot>y) = x" by simp
-
-text {* Mixfix declarations can be given for data constructors. *}
-
-domain d5 = d5a | d5b (lazy "d5") "d5" (infixl ":#:" 70)
-
-lemma "d5a \<noteq> x :#: y :#: z" by simp
-
-text {* Mixfix declarations can also be given for type constructors. *}
-
-domain ('a, 'b) lazypair (infixl ":*:" 25) =
-  lpair (lazy lfst :: 'a) (lazy lsnd :: 'b) (infixl ":*:" 75)
-
-lemma "\<forall>p::('a :*: 'b). p \<sqsubseteq> lfst\<cdot>p :*: lsnd\<cdot>p"
-by (rule allI, case_tac p, simp_all)
-
-text {* Non-recursive constructor arguments can have arbitrary types. *}
-
-domain ('a, 'b) d6 = d6 "int lift" "'a \<oplus> 'b u" (lazy "('a :*: 'b) \<times> ('b \<rightarrow> 'a)")
-
-text {*
-  Indirect recusion is allowed for sums, products, lifting, and the
-  continuous function space.  However, the domain package does not
-  generate an induction rule in terms of the constructors.
-*}
-
-domain 'a d7 = d7a "'a d7 \<oplus> int lift" | d7b "'a \<otimes> 'a d7" | d7c (lazy "'a d7 \<rightarrow> 'a")
-  -- "Indirect recursion detected, skipping proofs of (co)induction rules"
-(* d7.induct is absent *)
-
-text {*
-  Indirect recursion is also allowed using previously-defined datatypes.
-*}
-
-domain 'a slist = SNil | SCons 'a "'a slist"
-
-domain 'a stree = STip | SBranch "'a stree slist"
-
-text {* Mutually-recursive datatypes can be defined using the @{text "and"} keyword. *}
-
-domain d8 = d8a | d8b "d9" and d9 = d9a | d9b (lazy "d8")
-
-text {* Non-regular recursion is not allowed. *}
-(*
-domain ('a, 'b) altlist = ANil | ACons 'a "('b, 'a) altlist"
-  -- "illegal direct recursion with different arguments"
-domain 'a nest = Nest1 'a | Nest2 "'a nest nest"
-  -- "illegal direct recursion with different arguments"
-*)
-
-text {*
-  Mutually-recursive datatypes must have all the same type arguments,
-  not necessarily in the same order.
-*}
-
-domain ('a, 'b) list1 = Nil1 | Cons1 'a "('b, 'a) list2"
-   and ('b, 'a) list2 = Nil2 | Cons2 'b "('a, 'b) list1"
-
-text {* Induction rules for flat datatypes have no admissibility side-condition. *}
-
-domain 'a flattree = Tip | Branch "'a flattree" "'a flattree"
-
-lemma "\<lbrakk>P \<bottom>; P Tip; \<And>x y. \<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>; P x; P y\<rbrakk> \<Longrightarrow> P (Branch\<cdot>x\<cdot>y)\<rbrakk> \<Longrightarrow> P x"
-by (rule flattree.induct) -- "no admissibility requirement"
-
-text {* Trivial datatypes will produce a warning message. *}
-
-domain triv = Triv triv triv
-  -- "domain Domain_ex.triv is empty!"
-
-lemma "(x::triv) = \<bottom>" by (induct x, simp_all)
-
-text {* Lazy constructor arguments may have unpointed types. *}
-
-domain natlist = nnil | ncons (lazy "nat discr") natlist
-
-text {* Class constraints may be given for type parameters on the LHS. *}
-
-domain ('a::cpo) box = Box (lazy 'a)
-
-
-subsection {* Generated constants and theorems *}
-
-domain 'a tree = Leaf (lazy 'a) | Node (left :: "'a tree") (right :: "'a tree")
-
-lemmas tree_abs_defined_iff =
-  iso.abs_defined_iff [OF iso.intro [OF tree.abs_iso tree.rep_iso]]
-
-text {* Rules about ismorphism *}
-term tree_rep
-term tree_abs
-thm tree.rep_iso
-thm tree.abs_iso
-thm tree.iso_rews
-
-text {* Rules about constructors *}
-term Leaf
-term Node
-thm Leaf_def Node_def
-thm tree.nchotomy
-thm tree.exhaust
-thm tree.compacts
-thm tree.con_rews
-thm tree.dist_les
-thm tree.dist_eqs
-thm tree.inverts
-thm tree.injects
-
-text {* Rules about case combinator *}
-term tree_when
-thm tree.tree_when_def
-thm tree.when_rews
-
-text {* Rules about selectors *}
-term left
-term right
-thm tree.sel_rews
-
-text {* Rules about discriminators *}
-term is_Leaf
-term is_Node
-thm tree.dis_rews
-
-text {* Rules about pattern match combinators *}
-term Leaf_pat
-term Node_pat
-thm tree.pat_rews
-
-text {* Rules about monadic pattern match combinators *}
-term match_Leaf
-term match_Node
-thm tree.match_rews
-
-text {* Rules about take function *}
-term tree_take
-thm tree.take_def
-thm tree.take_0
-thm tree.take_Suc
-thm tree.take_rews
-thm tree.chain_take
-thm tree.take_take
-thm tree.deflation_take
-thm tree.take_below
-thm tree.take_lemma
-thm tree.lub_take
-thm tree.reach
-thm tree.finite_induct
-
-text {* Rules about finiteness predicate *}
-term tree_finite
-thm tree.finite_def
-thm tree.finite (* only generated for flat datatypes *)
-
-text {* Rules about bisimulation predicate *}
-term tree_bisim
-thm tree.bisim_def
-thm tree.coinduct
-
-text {* Induction rule *}
-thm tree.induct
-
-
-subsection {* Known bugs *}
-
-text {* Declaring a mixfix with spaces causes some strange parse errors. *}
-(*
-domain xx = xx ("x y")
-  -- "Inner syntax error: unexpected end of input"
-*)
-
-text {*
-  Non-cpo type parameters currently do not work.
-*}
-(*
-domain ('a::type) stream = snil | scons (lazy "'a discr") "'a stream"
-*)
-
-end
--- a/src/HOLCF/ex/Fixrec_ex.thy	Wed May 19 16:28:24 2010 -0700
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,246 +0,0 @@
-(*  Title:      HOLCF/ex/Fixrec_ex.thy
-    Author:     Brian Huffman
-*)
-
-header {* Fixrec package examples *}
-
-theory Fixrec_ex
-imports HOLCF
-begin
-
-subsection {* Basic @{text fixrec} examples *}
-
-text {*
-  Fixrec patterns can mention any constructor defined by the domain
-  package, as well as any of the following built-in constructors:
-  Pair, spair, sinl, sinr, up, ONE, TT, FF.
-*}
-
-text {* Typical usage is with lazy constructors. *}
-
-fixrec down :: "'a u \<rightarrow> 'a"
-where "down\<cdot>(up\<cdot>x) = x"
-
-text {* With strict constructors, rewrite rules may require side conditions. *}
-
-fixrec from_sinl :: "'a \<oplus> 'b \<rightarrow> 'a"
-where "x \<noteq> \<bottom> \<Longrightarrow> from_sinl\<cdot>(sinl\<cdot>x) = x"
-
-text {* Lifting can turn a strict constructor into a lazy one. *}
-
-fixrec from_sinl_up :: "'a u \<oplus> 'b \<rightarrow> 'a"
-where "from_sinl_up\<cdot>(sinl\<cdot>(up\<cdot>x)) = x"
-
-text {* Fixrec also works with the HOL pair constructor. *}
-
-fixrec down2 :: "'a u \<times> 'b u \<rightarrow> 'a \<times> 'b"
-where "down2\<cdot>(up\<cdot>x, up\<cdot>y) = (x, y)"
-
-
-subsection {* Examples using @{text fixrec_simp} *}
-
-text {* A type of lazy lists. *}
-
-domain 'a llist = lNil | lCons (lazy 'a) (lazy "'a llist")
-
-text {* A zip function for lazy lists. *}
-
-text {* Notice that the patterns are not exhaustive. *}
-
-fixrec
-  lzip :: "'a llist \<rightarrow> 'b llist \<rightarrow> ('a \<times> 'b) llist"
-where
-  "lzip\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>(lCons\<cdot>y\<cdot>ys) = lCons\<cdot>(x, y)\<cdot>(lzip\<cdot>xs\<cdot>ys)"
-| "lzip\<cdot>lNil\<cdot>lNil = lNil"
-
-text {* @{text fixrec_simp} is useful for producing strictness theorems. *}
-text {* Note that pattern matching is done in left-to-right order. *}
-
-lemma lzip_stricts [simp]:
-  "lzip\<cdot>\<bottom>\<cdot>ys = \<bottom>"
-  "lzip\<cdot>lNil\<cdot>\<bottom> = \<bottom>"
-  "lzip\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>\<bottom> = \<bottom>"
-by fixrec_simp+
-
-text {* @{text fixrec_simp} can also produce rules for missing cases. *}
-
-lemma lzip_undefs [simp]:
-  "lzip\<cdot>lNil\<cdot>(lCons\<cdot>y\<cdot>ys) = \<bottom>"
-  "lzip\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>lNil = \<bottom>"
-by fixrec_simp+
-
-
-subsection {* Pattern matching with bottoms *}
-
-text {*
-  As an alternative to using @{text fixrec_simp}, it is also possible
-  to use bottom as a constructor pattern.  When using a bottom
-  pattern, the right-hand-side must also be bottom; otherwise, @{text
-  fixrec} will not be able to prove the equation.
-*}
-
-fixrec
-  from_sinr_up :: "'a \<oplus> 'b\<^sub>\<bottom> \<rightarrow> 'b"
-where
-  "from_sinr_up\<cdot>\<bottom> = \<bottom>"
-| "from_sinr_up\<cdot>(sinr\<cdot>(up\<cdot>x)) = x"
-
-text {*
-  If the function is already strict in that argument, then the bottom
-  pattern does not change the meaning of the function.  For example,
-  in the definition of @{term from_sinr_up}, the first equation is
-  actually redundant, and could have been proven separately by
-  @{text fixrec_simp}.
-*}
-
-text {*
-  A bottom pattern can also be used to make a function strict in a
-  certain argument, similar to a bang-pattern in Haskell.
-*}
-
-fixrec
-  seq :: "'a \<rightarrow> 'b \<rightarrow> 'b"
-where
-  "seq\<cdot>\<bottom>\<cdot>y = \<bottom>"
-| "x \<noteq> \<bottom> \<Longrightarrow> seq\<cdot>x\<cdot>y = y"
-
-
-subsection {* Skipping proofs of rewrite rules *}
-
-text {* Another zip function for lazy lists. *}
-
-text {*
-  Notice that this version has overlapping patterns.
-  The second equation cannot be proved as a theorem
-  because it only applies when the first pattern fails.
-*}
-
-fixrec (permissive)
-  lzip2 :: "'a llist \<rightarrow> 'b llist \<rightarrow> ('a \<times> 'b) llist"
-where
-  "lzip2\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>(lCons\<cdot>y\<cdot>ys) = lCons\<cdot>(x, y)\<cdot>(lzip\<cdot>xs\<cdot>ys)"
-| "lzip2\<cdot>xs\<cdot>ys = lNil"
-
-text {*
-  Usually fixrec tries to prove all equations as theorems.
-  The "permissive" option overrides this behavior, so fixrec
-  does not produce any simp rules.
-*}
-
-text {* Simp rules can be generated later using @{text fixrec_simp}. *}
-
-lemma lzip2_simps [simp]:
-  "lzip2\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>(lCons\<cdot>y\<cdot>ys) = lCons\<cdot>(x, y)\<cdot>(lzip\<cdot>xs\<cdot>ys)"
-  "lzip2\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>lNil = lNil"
-  "lzip2\<cdot>lNil\<cdot>(lCons\<cdot>y\<cdot>ys) = lNil"
-  "lzip2\<cdot>lNil\<cdot>lNil = lNil"
-by fixrec_simp+
-
-lemma lzip2_stricts [simp]:
-  "lzip2\<cdot>\<bottom>\<cdot>ys = \<bottom>"
-  "lzip2\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>\<bottom> = \<bottom>"
-by fixrec_simp+
-
-
-subsection {* Mutual recursion with @{text fixrec} *}
-
-text {* Tree and forest types. *}
-
-domain 'a tree = Leaf (lazy 'a) | Branch (lazy "'a forest")
-and    'a forest = Empty | Trees (lazy "'a tree") "'a forest"
-
-text {*
-  To define mutually recursive functions, give multiple type signatures
-  separated by the keyword @{text "and"}.
-*}
-
-fixrec
-  map_tree :: "('a \<rightarrow> 'b) \<rightarrow> ('a tree \<rightarrow> 'b tree)"
-and
-  map_forest :: "('a \<rightarrow> 'b) \<rightarrow> ('a forest \<rightarrow> 'b forest)"
-where
-  "map_tree\<cdot>f\<cdot>(Leaf\<cdot>x) = Leaf\<cdot>(f\<cdot>x)"
-| "map_tree\<cdot>f\<cdot>(Branch\<cdot>ts) = Branch\<cdot>(map_forest\<cdot>f\<cdot>ts)"
-| "map_forest\<cdot>f\<cdot>Empty = Empty"
-| "ts \<noteq> \<bottom> \<Longrightarrow>
-    map_forest\<cdot>f\<cdot>(Trees\<cdot>t\<cdot>ts) = Trees\<cdot>(map_tree\<cdot>f\<cdot>t)\<cdot>(map_forest\<cdot>f\<cdot>ts)"
-
-lemma map_tree_strict [simp]: "map_tree\<cdot>f\<cdot>\<bottom> = \<bottom>"
-by fixrec_simp
-
-lemma map_forest_strict [simp]: "map_forest\<cdot>f\<cdot>\<bottom> = \<bottom>"
-by fixrec_simp
-
-text {*
-  Theorems generated:
-  @{text map_tree_def}  @{thm map_tree_def}
-  @{text map_forest_def}  @{thm map_forest_def}
-  @{text map_tree.unfold}  @{thm map_tree.unfold}
-  @{text map_forest.unfold}  @{thm map_forest.unfold}
-  @{text map_tree.simps}  @{thm map_tree.simps}
-  @{text map_forest.simps}  @{thm map_forest.simps}
-  @{text map_tree_map_forest.induct}  @{thm map_tree_map_forest.induct}
-*}
-
-
-subsection {* Looping simp rules *}
-
-text {*
-  The defining equations of a fixrec definition are declared as simp
-  rules by default.  In some cases, especially for constants with no
-  arguments or functions with variable patterns, the defining
-  equations may cause the simplifier to loop.  In these cases it will
-  be necessary to use a @{text "[simp del]"} declaration.
-*}
-
-fixrec
-  repeat :: "'a \<rightarrow> 'a llist"
-where
-  [simp del]: "repeat\<cdot>x = lCons\<cdot>x\<cdot>(repeat\<cdot>x)"
-
-text {*
-  We can derive other non-looping simp rules for @{const repeat} by
-  using the @{text subst} method with the @{text repeat.simps} rule.
-*}
-
-lemma repeat_simps [simp]:
-  "repeat\<cdot>x \<noteq> \<bottom>"
-  "repeat\<cdot>x \<noteq> lNil"
-  "repeat\<cdot>x = lCons\<cdot>y\<cdot>ys \<longleftrightarrow> x = y \<and> repeat\<cdot>x = ys"
-by (subst repeat.simps, simp)+
-
-lemma llist_when_repeat [simp]:
-  "llist_when\<cdot>z\<cdot>f\<cdot>(repeat\<cdot>x) = f\<cdot>x\<cdot>(repeat\<cdot>x)"
-by (subst repeat.simps, simp)
-
-text {*
-  For mutually-recursive constants, looping might only occur if all
-  equations are in the simpset at the same time.  In such cases it may
-  only be necessary to declare @{text "[simp del]"} on one equation.
-*}
-
-fixrec
-  inf_tree :: "'a tree" and inf_forest :: "'a forest"
-where
-  [simp del]: "inf_tree = Branch\<cdot>inf_forest"
-| "inf_forest = Trees\<cdot>inf_tree\<cdot>(Trees\<cdot>inf_tree\<cdot>Empty)"
-
-
-subsection {* Using @{text fixrec} inside locales *}
-
-locale test =
-  fixes foo :: "'a \<rightarrow> 'a"
-  assumes foo_strict: "foo\<cdot>\<bottom> = \<bottom>"
-begin
-
-fixrec
-  bar :: "'a u \<rightarrow> 'a"
-where
-  "bar\<cdot>(up\<cdot>x) = foo\<cdot>x"
-
-lemma bar_strict: "bar\<cdot>\<bottom> = \<bottom>"
-by fixrec_simp
-
-end
-
-end
--- a/src/HOLCF/ex/New_Domain.thy	Wed May 19 16:28:24 2010 -0700
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,92 +0,0 @@
-(*  Title:      HOLCF/ex/New_Domain.thy
-    Author:     Brian Huffman
-*)
-
-header {* Definitional domain package *}
-
-theory New_Domain
-imports HOLCF
-begin
-
-text {*
-  The definitional domain package only works with representable domains,
-  i.e. types in class @{text rep}.
-*}
-
-default_sort rep
-
-text {*
-  Provided that @{text rep} is the default sort, the @{text new_domain}
-  package should work with any type definition supported by the old
-  domain package.
-*}
-
-new_domain 'a llist = LNil | LCons (lazy 'a) (lazy "'a llist")
-
-text {*
-  The difference is that the new domain package is completely
-  definitional, and does not generate any axioms.  The following type
-  and constant definitions are not produced by the old domain package.
-*}
-
-thm type_definition_llist
-thm llist_abs_def llist_rep_def
-
-text {*
-  The new domain package also adds support for indirect recursion with
-  user-defined datatypes.  This definition of a tree datatype uses
-  indirect recursion through the lazy list type constructor.
-*}
-
-new_domain 'a ltree = Leaf (lazy 'a) | Branch (lazy "'a ltree llist")
-
-text {*
-  For indirect-recursive definitions, the domain package is not able to
-  generate a high-level induction rule.  (It produces a warning
-  message instead.)  The low-level reach lemma (now proved as a
-  theorem, no longer generated as an axiom) can be used to derive
-  other induction rules.
-*}
-
-thm ltree.reach
-
-text {*
-  The definition of the take function uses map functions associated with
-  each type constructor involved in the definition.  A map function
-  for the lazy list type has been generated by the new domain package.
-*}
-
-thm ltree.take_rews
-thm llist_map_def
-
-lemma ltree_induct:
-  fixes P :: "'a ltree \<Rightarrow> bool"
-  assumes adm: "adm P"
-  assumes bot: "P \<bottom>"
-  assumes Leaf: "\<And>x. P (Leaf\<cdot>x)"
-  assumes Branch: "\<And>f l. \<forall>x. P (f\<cdot>x) \<Longrightarrow> P (Branch\<cdot>(llist_map\<cdot>f\<cdot>l))"
-  shows "P x"
-proof -
-  have "P (\<Squnion>i. ltree_take i\<cdot>x)"
-  using adm
-  proof (rule admD)
-    fix i
-    show "P (ltree_take i\<cdot>x)"
-    proof (induct i arbitrary: x)
-      case (0 x)
-      show "P (ltree_take 0\<cdot>x)" by (simp add: bot)
-    next
-      case (Suc n x)
-      show "P (ltree_take (Suc n)\<cdot>x)"
-        apply (cases x)
-        apply (simp add: bot)
-        apply (simp add: Leaf)
-        apply (simp add: Branch Suc)
-        done
-    qed
-  qed (simp add: ltree.chain_take)
-  thus ?thesis
-    by (simp add: ltree.reach)
-qed
-
-end
--- a/src/HOLCF/ex/ROOT.ML	Wed May 19 16:28:24 2010 -0700
+++ b/src/HOLCF/ex/ROOT.ML	Wed May 19 17:01:07 2010 -0700
@@ -4,7 +4,6 @@
 *)
 
 use_thys ["Dnat", "Stream", "Dagstuhl", "Focus_ex", "Fix2", "Hoare",
-  "Loop", "Fixrec_ex", "Powerdomain_ex", "Domain_ex", "Domain_Proofs",
+  "Loop", "Powerdomain_ex", "Domain_Proofs",
   "Letrec",
-  "Strict_Fun",
-  "New_Domain"];
+  "Strict_Fun"];