Misc minor updates
authorlcp
Tue, 26 Jul 1994 14:02:16 +0200
changeset 486 6b58082796f6
parent 485 5e00a676a211
child 487 af83700cb771
Misc minor updates
src/ZF/ex/Comb.ML
src/ZF/ex/Contract0.ML
src/ZF/ex/Data.ML
src/ZF/ex/LList.ML
src/ZF/ex/Ntree.ML
src/ZF/ex/TF.ML
src/ZF/ex/Term.ML
--- a/src/ZF/ex/Comb.ML	Tue Jul 26 13:44:42 1994 +0200
+++ b/src/ZF/ex/Comb.ML	Tue Jul 26 14:02:16 1994 +0200
@@ -17,8 +17,8 @@
   val thy_name = "Comb";
   val rec_specs = 
       [("comb", "univ(0)",
-	  [(["K","S"],	"i", NoSyn),
-	   (["#"],	"[i,i]=>i", Infixl 90)])];
+	  [(["K","S"],	"i", 		NoSyn),
+	   (["#"],	"[i,i]=>i", 	Infixl 90)])];
   val rec_styp = "i";
   val sintrs = 
 	  ["K : comb",
--- a/src/ZF/ex/Contract0.ML	Tue Jul 26 13:44:42 1994 +0200
+++ b/src/ZF/ex/Contract0.ML	Tue Jul 26 14:02:16 1994 +0200
@@ -1,4 +1,4 @@
-(*  Title: 	ZF/ex/contract.ML
+(*  Title: 	ZF/ex/Contract0.ML
     ID:         $Id$
     Author: 	Lawrence C Paulson
     Copyright   1993  University of Cambridge
--- a/src/ZF/ex/Data.ML	Tue Jul 26 13:44:42 1994 +0200
+++ b/src/ZF/ex/Data.ML	Tue Jul 26 14:02:16 1994 +0200
@@ -11,10 +11,10 @@
  (val thy        = Univ.thy
   val thy_name = "Data"
   val rec_specs  = [("data", "univ(A Un B)",
-                       [(["Con0"],   "i",NoSyn),
-                        (["Con1"],   "i=>i",NoSyn),
-                        (["Con2"],   "[i,i]=>i",NoSyn),
-                        (["Con3"],   "[i,i,i]=>i",NoSyn)])]
+                       [(["Con0"],   "i",	    NoSyn),
+                        (["Con1"],   "i=>i",	    NoSyn),
+                        (["Con2"],   "[i,i]=>i",    NoSyn),
+                        (["Con3"],   "[i,i,i]=>i",  NoSyn)])]
   val rec_styp   = "[i,i]=>i"
   val sintrs     = 
           ["Con0 : data(A,B)",
--- a/src/ZF/ex/LList.ML	Tue Jul 26 13:44:42 1994 +0200
+++ b/src/ZF/ex/LList.ML	Tue Jul 26 14:02:16 1994 +0200
@@ -10,8 +10,8 @@
  (val thy        = QUniv.thy
   val thy_name   = "LList"
   val rec_specs  = [("llist", "quniv(A)",
-                      [(["LNil"],   "i",NoSyn), 
-                       (["LCons"],  "[i,i]=>i",NoSyn)])]
+                      [(["LNil"],   "i",	NoSyn), 
+                       (["LCons"],  "[i,i]=>i",	NoSyn)])]
   val rec_styp   = "i=>i"
   val sintrs     = ["LNil : llist(A)",
                     "[| a: A;  l: llist(A) |] ==> LCons(a,l) : llist(A)"]
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/ZF/ex/Ntree.ML	Tue Jul 26 14:02:16 1994 +0200
@@ -0,0 +1,82 @@
+(*  Title: 	ZF/ex/Ntree.ML
+    ID:         $Id$
+    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1994  University of Cambridge
+
+Datatype definition n-ary branching trees
+Demonstrates a simple use of function space in a datatype definition
+
+Based upon ex/Term.ML
+*)
+
+structure Ntree = Datatype_Fun
+ (val thy	= Univ.thy;
+  val thy_name 	= "Ntree";
+  val rec_specs = 
+      [("ntree", "univ(A)",
+	  [(["Branch"], "[i,i]=>i", NoSyn)])];
+  val rec_styp 	= "i=>i";
+  val sintrs 	= 
+	["[| a: A;  n: nat;  h: n -> ntree(A) |] ==> Branch(a,h) : ntree(A)"];
+  val monos 	= [Pi_mono];
+  val type_intrs = (nat_fun_univ RS subsetD) :: datatype_intrs;
+  val type_elims = []);
+
+val [BranchI] = Ntree.intrs;
+
+goal Ntree.thy "ntree(A) = A * (UN n: nat. n -> ntree(A))";
+by (rtac (Ntree.unfold RS trans) 1);
+bws Ntree.con_defs;
+by (fast_tac (sum_cs addIs ([equalityI] @ datatype_intrs)
+ 	             addDs [Ntree.dom_subset RS subsetD]
+ 	             addEs [A_into_univ, nat_fun_into_univ]) 1);
+val ntree_unfold = result();
+
+(*A nicer induction rule than the standard one*)
+val major::prems = goal Ntree.thy
+    "[| t: ntree(A);  \
+\       !!x n h. [| x: A;  n: nat;  h: n -> ntree(A);  ALL i:n. P(h`i)  \
+\                |] ==> P(Branch(x,h))  \
+\    |] ==> P(t)";
+by (rtac (major RS Ntree.induct) 1);
+by (REPEAT_SOME (ares_tac prems));
+by (fast_tac (ZF_cs addEs [fun_weaken_type]) 1);
+by (fast_tac (ZF_cs addDs [apply_type]) 1);
+val ntree_induct2 = result();
+
+(*Induction on ntree(A) to prove an equation*)
+val major::prems = goal Ntree.thy
+  "[| t: ntree(A);  f: ntree(A)->B;  g: ntree(A)->B;		\
+\     !!x n h. [| x: A;  n: nat;  h: n -> ntree(A);  f O h = g O h |] ==> \
+\              f ` Branch(x,h) = g ` Branch(x,h)  		\
+\  |] ==> f`t=g`t";
+by (rtac (major RS ntree_induct2) 1);
+by (REPEAT_SOME (ares_tac prems));
+by (cut_facts_tac prems 1);
+br fun_extension 1;
+by (REPEAT_SOME (ares_tac [comp_fun]));
+by (asm_simp_tac (ZF_ss addsimps [comp_fun_apply]) 1);
+val ntree_induct_eqn = result();
+
+(**  Lemmas to justify using "Ntree" in other recursive type definitions **)
+
+goalw Ntree.thy Ntree.defs "!!A B. A<=B ==> ntree(A) <= ntree(B)";
+by (rtac lfp_mono 1);
+by (REPEAT (rtac Ntree.bnd_mono 1));
+by (REPEAT (ares_tac (univ_mono::basic_monos) 1));
+val ntree_mono = result();
+
+(*Easily provable by induction also*)
+goalw Ntree.thy (Ntree.defs@Ntree.con_defs) "ntree(univ(A)) <= univ(A)";
+by (rtac lfp_lowerbound 1);
+by (rtac (A_subset_univ RS univ_mono) 2);
+by (safe_tac ZF_cs);
+by (REPEAT (ares_tac [Pair_in_univ, nat_fun_univ RS subsetD] 1));
+val ntree_univ = result();
+
+val ntree_subset_univ = 
+    [ntree_mono, ntree_univ] MRS subset_trans |> standard;
+
+goal Ntree.thy "!!t A B. [| t: ntree(A);  A <= univ(B) |] ==> t: univ(B)";
+by (REPEAT (ares_tac [ntree_subset_univ RS subsetD] 1));
+val ntree_into_univ = result();
--- a/src/ZF/ex/TF.ML	Tue Jul 26 13:44:42 1994 +0200
+++ b/src/ZF/ex/TF.ML	Tue Jul 26 14:02:16 1994 +0200
@@ -10,10 +10,10 @@
  (val thy        = Univ.thy
   val thy_name   = "TF"
   val rec_specs  = [("tree", "univ(A)",
-                       [(["Tcons"],  "[i,i]=>i",NoSyn)]),
+                       [(["Tcons"],  "[i,i]=>i",  NoSyn)]),
                     ("forest", "univ(A)",
-                       [(["Fnil"],   "i",NoSyn),
-                        (["Fcons"],  "[i,i]=>i",NoSyn)])]
+                       [(["Fnil"],   "i",         NoSyn),
+                        (["Fcons"],  "[i,i]=>i",  NoSyn)])]
   val rec_styp   = "i=>i"
   val sintrs     = 
           ["[| a:A;  f: forest(A) |] ==> Tcons(a,f) : tree(A)",
--- a/src/ZF/ex/Term.ML	Tue Jul 26 13:44:42 1994 +0200
+++ b/src/ZF/ex/Term.ML	Tue Jul 26 14:02:16 1994 +0200
@@ -1,27 +1,26 @@
-(*  Title: 	ZF/ex/term.ML
+(*  Title: 	ZF/ex/Term.ML
     ID:         $Id$
     Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1993  University of Cambridge
+    Copyright   1994  University of Cambridge
 
 Datatype definition of terms over an alphabet.
 Illustrates the list functor (essentially the same type as in Trees & Forests)
 *)
 
 structure Term = Datatype_Fun
- (val thy = List.thy;
-  val thy_name = "Term";
+ (val thy	= List.thy;
+  val thy_name	= "Term";
   val rec_specs = 
       [("term", "univ(A)",
 	  [(["Apply"], "[i,i]=>i", NoSyn)])];
-  val rec_styp = "i=>i";
-  val sintrs = ["[| a: A;  l: list(term(A)) |] ==> Apply(a,l) : term(A)"];
-  val monos = [list_mono];
-  val type_intrs = datatype_intrs;
-  val type_elims = [make_elim (list_univ RS subsetD)]);
+  val rec_styp	= "i=>i";
+  val sintrs	= ["[| a: A;  l: list(term(A)) |] ==> Apply(a,l) : term(A)"];
+  val monos	= [list_mono];
+  val type_intrs = (list_univ RS subsetD) :: datatype_intrs;
+  val type_elims = []);
 
 val [ApplyI] = Term.intrs;
 
-(*Note that Apply is the identity function*)
 goal Term.thy "term(A) = A * list(term(A))";
 by (rtac (Term.unfold RS trans) 1);
 bws Term.con_defs;