src/HOL/Isar_Examples/Nested_Datatype.thy
changeset 37597 a02ea93e88c6
parent 33026 8f35633c4922
child 37671 fa53d267dab3
--- a/src/HOL/Isar_Examples/Nested_Datatype.thy	Mon Jun 28 15:32:08 2010 +0200
+++ b/src/HOL/Isar_Examples/Nested_Datatype.thy	Mon Jun 28 15:32:13 2010 +0200
@@ -10,17 +10,14 @@
     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_term :: "('a => ('a, 'b) term) => ('a, 'b) term => ('a, 'b) term" and
+  subst_term_list :: "('a => ('a, 'b) term) => ('a, 'b) term list => ('a, 'b) term list" where
+  "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"
 
-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"
-
+lemmas subst_simps = subst_term_subst_term_list.simps
 
 text {*
  \medskip A simple lemma about composition of substitutions.
@@ -44,13 +41,13 @@
   next
     fix b ts assume "?Q ts"
     then show "?P (App b ts)"
-      by (simp only: subst.simps)
+      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)
+      by (simp only: subst_simps)
   qed
 qed
 
@@ -59,18 +56,18 @@
 
 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)"
+    and app: "!!b ts. (\<forall>t \<in> set ts. P t) ==> 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"
+  fix b t ts assume "\<forall>t \<in> set ts. P t"
   then show "P (App b ts)" by (rule app)
 next
-  show "list_all P []" by simp
+  show "\<forall>t \<in> set []. P t" by simp
 next
-  fix t ts assume "P t" "list_all P ts"
-  then show "list_all P (t # ts)" by simp
+  fix t ts assume "P t" "\<forall>t' \<in> set ts. P t'"
+  then show "\<forall>t' \<in> set (t # ts). P t'" by simp
 qed
 
 lemma