src/HOL/indrule.ML
changeset 1465 5d7a7e439cec
parent 1424 ccb3c5ff6707
child 1653 1a2ffa2fbf7d
--- a/src/HOL/indrule.ML	Tue Jan 30 15:19:20 1996 +0100
+++ b/src/HOL/indrule.ML	Tue Jan 30 15:24:36 1996 +0100
@@ -1,6 +1,6 @@
-(*  Title: 	HOL/indrule.ML
+(*  Title:      HOL/indrule.ML
     ID:         $Id$
-    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1994  University of Cambridge
 
 Induction rule module -- for Inductive/Coinductive Definitions
@@ -10,14 +10,14 @@
 
 signature INDRULE =
   sig
-  val induct        : thm			(*main induction rule*)
-  val mutual_induct : thm			(*mutual induction rule*)
+  val induct        : thm                       (*main induction rule*)
+  val mutual_induct : thm                       (*mutual induction rule*)
   end;
 
 
 functor Indrule_Fun
     (structure Inductive: sig include INDUCTIVE_ARG INDUCTIVE_I end and
-	 Intr_elim: sig include INTR_ELIM INTR_ELIM_AUX end) : INDRULE  =
+         Intr_elim: sig include INTR_ELIM INTR_ELIM_AUX end) : INDRULE  =
 let
 
 val sign = sign_of Inductive.thy;
@@ -32,7 +32,7 @@
 
 (*** Prove the main induction rule ***)
 
-val pred_name = "P";		(*name for predicate variables*)
+val pred_name = "P";            (*name for predicate variables*)
 
 val big_rec_def::part_rec_defs = Intr_elim.defs;
 
@@ -40,25 +40,25 @@
    ind_alist = [(rec_tm1,pred1),...]  -- associates predicates with rec ops
    prem is a premise of an intr rule*)
 fun add_induct_prem ind_alist (prem as Const("Trueprop",_) $ 
-		 (Const("op :",_)$t$X), iprems) =
+                 (Const("op :",_)$t$X), iprems) =
      (case gen_assoc (op aconv) (ind_alist, X) of
-	  Some pred => prem :: Ind_Syntax.mk_Trueprop (pred $ t) :: iprems
-	| None => (*possibly membership in M(rec_tm), for M monotone*)
-	    let fun mk_sb (rec_tm,pred) = 
-		 (case binder_types (fastype_of pred) of
-		      [T] => (rec_tm, 
-			      Ind_Syntax.Int_const T $ rec_tm $ 
-			        (Ind_Syntax.Collect_const T $ pred))
-		    | _ => error 
-		      "Bug: add_induct_prem called with non-unary predicate")
-	    in  subst_free (map mk_sb ind_alist) prem :: iprems  end)
+          Some pred => prem :: Ind_Syntax.mk_Trueprop (pred $ t) :: iprems
+        | None => (*possibly membership in M(rec_tm), for M monotone*)
+            let fun mk_sb (rec_tm,pred) = 
+                 (case binder_types (fastype_of pred) of
+                      [T] => (rec_tm, 
+                              Ind_Syntax.Int_const T $ rec_tm $ 
+                                (Ind_Syntax.Collect_const T $ pred))
+                    | _ => error 
+                      "Bug: add_induct_prem called with non-unary predicate")
+            in  subst_free (map mk_sb ind_alist) prem :: iprems  end)
   | add_induct_prem ind_alist (prem,iprems) = prem :: iprems;
 
 (*Make a premise of the induction rule.*)
 fun induct_prem ind_alist intr =
   let val quantfrees = map dest_Free (term_frees intr \\ rec_params)
       val iprems = foldr (add_induct_prem ind_alist)
-			 (Logic.strip_imp_prems intr,[])
+                         (Logic.strip_imp_prems intr,[])
       val (t,X) = Ind_Syntax.rule_concl intr
       val (Some pred) = gen_assoc (op aconv) (ind_alist, X)
       val concl = Ind_Syntax.mk_Trueprop (pred $ t)
@@ -68,8 +68,8 @@
 (*Avoids backtracking by delivering the correct premise to each goal*)
 fun ind_tac [] 0 = all_tac
   | ind_tac(prem::prems) i = 
-	DEPTH_SOLVE_1 (ares_tac [Part_eqI, prem, refl] i) THEN
-	ind_tac prems (i-1);
+        DEPTH_SOLVE_1 (ares_tac [Part_eqI, prem, refl] i) THEN
+        ind_tac prems (i-1);
 
 val pred = Free(pred_name, elem_type --> Ind_Syntax.boolT);
 
@@ -85,15 +85,15 @@
     prove_goalw_cterm part_rec_defs 
       (cterm_of sign 
        (Logic.list_implies (ind_prems, 
-			    Ind_Syntax.mk_Trueprop (Ind_Syntax.mk_all_imp 
-						    (big_rec_tm,pred)))))
+                            Ind_Syntax.mk_Trueprop (Ind_Syntax.mk_all_imp 
+                                                    (big_rec_tm,pred)))))
       (fn prems =>
        [rtac (impI RS allI) 1,
-	DETERM (etac Intr_elim.raw_induct 1),
-	asm_full_simp_tac (!simpset addsimps [Part_Collect]) 1,
-	REPEAT (FIRSTGOAL (eresolve_tac [IntE, CollectE, exE, conjE, disjE] 
-			   ORELSE' hyp_subst_tac)),
-	ind_tac (rev prems) (length prems)])
+        DETERM (etac Intr_elim.raw_induct 1),
+        asm_full_simp_tac (!simpset addsimps [Part_Collect]) 1,
+        REPEAT (FIRSTGOAL (eresolve_tac [IntE, CollectE, exE, conjE, disjE] 
+                           ORELSE' hyp_subst_tac)),
+        ind_tac (rev prems) (length prems)])
     handle e => print_sign_exn sign e;
 
 (*** Prove the simultaneous induction rule ***)
@@ -109,11 +109,11 @@
       val pfree = Free(pred_name ^ "_" ^ rec_name, T)
       val frees = mk_frees "za" (binder_types T)
       val qconcl = 
-	foldr Ind_Syntax.mk_all 
-	  (frees, 
-	   Ind_Syntax.imp $ (Ind_Syntax.mk_mem 
-			     (foldr1 Ind_Syntax.mk_Pair frees, rec_tm))
-	        $ (list_comb (pfree,frees)))
+        foldr Ind_Syntax.mk_all 
+          (frees, 
+           Ind_Syntax.imp $ (Ind_Syntax.mk_mem 
+                             (foldr1 Ind_Syntax.mk_Pair frees, rec_tm))
+                $ (list_comb (pfree,frees)))
   in  (Ind_Syntax.ap_split Ind_Syntax.boolT pfree (binder_types T), 
       qconcl)  
   end;
@@ -129,21 +129,21 @@
     Ind_Syntax.mk_Trueprop
       (Ind_Syntax.mk_all_imp
        (big_rec_tm,
-	Abs("z", elem_type, 
-	    fold_bal (app Ind_Syntax.conj) 
-	    (map mk_rec_imp (Inductive.rec_tms~~preds)))))
+        Abs("z", elem_type, 
+            fold_bal (app Ind_Syntax.conj) 
+            (map mk_rec_imp (Inductive.rec_tms~~preds)))))
 and mutual_induct_concl = 
     Ind_Syntax.mk_Trueprop (fold_bal (app Ind_Syntax.conj) qconcls);
 
 val lemma = (*makes the link between the two induction rules*)
     prove_goalw_cterm part_rec_defs 
-	  (cterm_of sign (Logic.mk_implies (induct_concl,
-					    mutual_induct_concl)))
-	  (fn prems =>
-	   [cut_facts_tac prems 1,
-	    REPEAT (eresolve_tac [asm_rl, conjE, PartE, mp] 1
-	     ORELSE resolve_tac [allI, impI, conjI, Part_eqI, refl] 1
-	     ORELSE dresolve_tac [spec, mp, splitD] 1)])
+          (cterm_of sign (Logic.mk_implies (induct_concl,
+                                            mutual_induct_concl)))
+          (fn prems =>
+           [cut_facts_tac prems 1,
+            REPEAT (eresolve_tac [asm_rl, conjE, PartE, mp] 1
+             ORELSE resolve_tac [allI, impI, conjI, Part_eqI, refl] 1
+             ORELSE dresolve_tac [spec, mp, splitD] 1)])
     handle e => print_sign_exn sign e;
 
 (*Mutual induction follows by freeness of Inl/Inr.*)
@@ -164,43 +164,43 @@
   | mutual_ind_tac(prem::prems) i = 
       DETERM
        (SELECT_GOAL 
-	  (
-	   (*Simplify the assumptions and goal by unfolding Part and
-	     using freeness of the Sum constructors; proves all but one
+          (
+           (*Simplify the assumptions and goal by unfolding Part and
+             using freeness of the Sum constructors; proves all but one
              conjunct by contradiction*)
-	   rewrite_goals_tac all_defs  THEN
-	   simp_tac (mut_ss addsimps [Part_def]) 1  THEN
-	   IF_UNSOLVED (*simp_tac may have finished it off!*)
-	     ((*simplify assumptions, but don't accept new rewrite rules!*)
-	      asm_full_simp_tac (mut_ss setmksimps K[]) 1  THEN
-	      (*unpackage and use "prem" in the corresponding place*)
-	      REPEAT (rtac impI 1)  THEN
-	      rtac (rewrite_rule all_defs prem) 1  THEN
-	      (*prem must not be REPEATed below: could loop!*)
-	      DEPTH_SOLVE (FIRSTGOAL (ares_tac [impI] ORELSE' 
-				      eresolve_tac (conjE::mp::cmonos))))
-	  ) i)
+           rewrite_goals_tac all_defs  THEN
+           simp_tac (mut_ss addsimps [Part_def]) 1  THEN
+           IF_UNSOLVED (*simp_tac may have finished it off!*)
+             ((*simplify assumptions, but don't accept new rewrite rules!*)
+              asm_full_simp_tac (mut_ss setmksimps K[]) 1  THEN
+              (*unpackage and use "prem" in the corresponding place*)
+              REPEAT (rtac impI 1)  THEN
+              rtac (rewrite_rule all_defs prem) 1  THEN
+              (*prem must not be REPEATed below: could loop!*)
+              DEPTH_SOLVE (FIRSTGOAL (ares_tac [impI] ORELSE' 
+                                      eresolve_tac (conjE::mp::cmonos))))
+          ) i)
        THEN mutual_ind_tac prems (i-1);
 
 val _ = writeln "  Proving the mutual induction rule...";
 
 val mutual_induct_split = 
     prove_goalw_cterm []
-	  (cterm_of sign
-	   (Logic.list_implies (map (induct_prem (Inductive.rec_tms ~~ preds)) 
-			      Inductive.intr_tms,
-			  mutual_induct_concl)))
-	  (fn prems =>
-	   [rtac (quant_induct RS lemma) 1,
-	    mutual_ind_tac (rev prems) (length prems)])
+          (cterm_of sign
+           (Logic.list_implies (map (induct_prem (Inductive.rec_tms ~~ preds)) 
+                              Inductive.intr_tms,
+                          mutual_induct_concl)))
+          (fn prems =>
+           [rtac (quant_induct RS lemma) 1,
+            mutual_ind_tac (rev prems) (length prems)])
     handle e => print_sign_exn sign e;
 
 (*Attempts to remove all occurrences of split*)
 val split_tac =
     REPEAT (SOMEGOAL (FIRST' [rtac splitI, 
-			      dtac splitD,
-			      etac splitE,
-			      bound_hyp_subst_tac]))
+                              dtac splitD,
+                              etac splitE,
+                              bound_hyp_subst_tac]))
     THEN prune_params_tac;
 
 in
@@ -210,7 +210,7 @@
 
   val mutual_induct = 
       if length Intr_elim.rec_names > 1 orelse
-	 length (Ind_Syntax.factors elem_type) > 1
+         length (Ind_Syntax.factors elem_type) > 1
       then rule_by_tactic split_tac mutual_induct_split
       else TrueI;
   end