src/HOL/indrule.ML
author clasohm
Wed, 04 Oct 1995 13:10:03 +0100
changeset 1264 3eb91524b938
parent 1190 9d1bdce3a38e
child 1424 ccb3c5ff6707
permissions -rw-r--r--
added local simpsets; removed IOA from 'make test'
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     1
(*  Title: 	HOL/indrule.ML
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     3
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     6
Induction rule module -- for Inductive/Coinductive Definitions
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
Proves a strong induction rule and a mutual induction rule
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     9
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    11
signature INDRULE =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    12
  sig
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
  val induct        : thm			(*main induction rule*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    14
  val mutual_induct : thm			(*mutual induction rule*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    15
  end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    16
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    18
functor Indrule_Fun
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    19
    (structure Inductive: sig include INDUCTIVE_ARG INDUCTIVE_I end and
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
	 Intr_elim: INTR_ELIM) : INDRULE  =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    21
struct
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    22
open Logic Ind_Syntax Inductive Intr_elim;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    23
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
val sign = sign_of thy;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    25
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26
val (Const(_,recT),rec_params) = strip_comb (hd rec_tms);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    27
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    28
val elem_type = dest_setT (body_type recT);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    29
val big_rec_name = space_implode "_" rec_names;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    30
val big_rec_tm = list_comb(Const(big_rec_name,recT), rec_params);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    31
1190
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
    32
val _ = writeln "  Proving the induction rule...";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    33
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    34
(*** Prove the main induction rule ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    35
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    36
val pred_name = "P";		(*name for predicate variables*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    37
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    38
val big_rec_def::part_rec_defs = Intr_elim.defs;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    39
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    40
(*Used to express induction rules: adds induction hypotheses.
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    41
   ind_alist = [(rec_tm1,pred1),...]  -- associates predicates with rec ops
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    42
   prem is a premise of an intr rule*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    43
fun add_induct_prem ind_alist (prem as Const("Trueprop",_) $ 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    44
		 (Const("op :",_)$t$X), iprems) =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    45
     (case gen_assoc (op aconv) (ind_alist, X) of
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    46
	  Some pred => prem :: mk_Trueprop (pred $ t) :: iprems
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    47
	| None => (*possibly membership in M(rec_tm), for M monotone*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    48
	    let fun mk_sb (rec_tm,pred) = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    49
		 (case binder_types (fastype_of pred) of
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    50
		      [T] => (rec_tm, 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    51
			      Int_const T $ rec_tm $ (Collect_const T $ pred))
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    52
		    | _ => error 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    53
		      "Bug: add_induct_prem called with non-unary predicate")
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    54
	    in  subst_free (map mk_sb ind_alist) prem :: iprems  end)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
  | add_induct_prem ind_alist (prem,iprems) = prem :: iprems;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    56
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    57
(*Make a premise of the induction rule.*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    58
fun induct_prem ind_alist intr =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    59
  let val quantfrees = map dest_Free (term_frees intr \\ rec_params)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    60
      val iprems = foldr (add_induct_prem ind_alist)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    61
			 (strip_imp_prems intr,[])
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    62
      val (t,X) = rule_concl intr
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    63
      val (Some pred) = gen_assoc (op aconv) (ind_alist, X)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    64
      val concl = mk_Trueprop (pred $ t)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    65
  in list_all_free (quantfrees, list_implies (iprems,concl)) end
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    66
  handle Bind => error"Recursion term not found in conclusion";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    67
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    68
(*Avoids backtracking by delivering the correct premise to each goal*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    69
fun ind_tac [] 0 = all_tac
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    70
  | ind_tac(prem::prems) i = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    71
	DEPTH_SOLVE_1 (ares_tac [Part_eqI, prem, refl] i) THEN
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    72
	ind_tac prems (i-1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    73
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    74
val pred = Free(pred_name, elem_type --> boolT);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    75
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    76
val ind_prems = map (induct_prem (map (rpair pred) rec_tms)) intr_tms;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    77
1190
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
    78
(*Debugging code...
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
    79
val _ = writeln "ind_prems = ";
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
    80
val _ = seq (writeln o Sign.string_of_term sign) ind_prems;
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
    81
*)
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
    82
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    83
val quant_induct = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    84
    prove_goalw_cterm part_rec_defs 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    85
      (cterm_of sign (list_implies (ind_prems, 
1190
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
    86
				mk_Trueprop (mk_all_imp (big_rec_tm,pred)))))
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    87
      (fn prems =>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    88
       [rtac (impI RS allI) 1,
1190
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
    89
	DETERM (etac raw_induct 1),
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 1190
diff changeset
    90
	asm_full_simp_tac (!simpset addsimps [Part_Collect]) 1,
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    91
	REPEAT (FIRSTGOAL (eresolve_tac [IntE, CollectE, exE, conjE, disjE] 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    92
			   ORELSE' hyp_subst_tac)),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    93
	ind_tac (rev prems) (length prems)])
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    94
    handle e => print_sign_exn sign e;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    95
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    96
(*** Prove the simultaneous induction rule ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    97
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    98
(*Make distinct predicates for each inductive set.
1190
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
    99
  Splits cartesian products in elem_type, IF nested to the right! *)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   100
1190
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   101
(*Given a recursive set, return the "split" predicate
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   102
  and a conclusion for the simultaneous induction rule*)
1190
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   103
fun mk_predpair rec_tm = 
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   104
  let val rec_name = (#1 o dest_Const o head_of) rec_tm
1190
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   105
      val T = factors elem_type ---> boolT
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   106
      val pfree = Free(pred_name ^ "_" ^ rec_name, T)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   107
      val frees = mk_frees "za" (binder_types T)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   108
      val qconcl = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   109
	foldr mk_all (frees, 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   110
		      imp $ (mk_mem (foldr1 mk_Pair frees, rec_tm))
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   111
			  $ (list_comb (pfree,frees)))
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   112
  in  (ap_split boolT pfree (binder_types T), 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   113
      qconcl)  
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   114
  end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   115
1190
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   116
val (preds,qconcls) = split_list (map mk_predpair rec_tms);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   117
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   118
(*Used to form simultaneous induction lemma*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   119
fun mk_rec_imp (rec_tm,pred) = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   120
    imp $ (mk_mem (Bound 0, rec_tm)) $  (pred $ Bound 0);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   121
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   122
(*To instantiate the main induction rule*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   123
val induct_concl = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   124
 mk_Trueprop(mk_all_imp(big_rec_tm,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   125
		     Abs("z", elem_type, 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   126
			 fold_bal (app conj) 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   127
			          (map mk_rec_imp (rec_tms~~preds)))))
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   128
and mutual_induct_concl = mk_Trueprop(fold_bal (app conj) qconcls);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   129
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   130
val lemma = (*makes the link between the two induction rules*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   131
    prove_goalw_cterm part_rec_defs 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   132
	  (cterm_of sign (mk_implies (induct_concl,mutual_induct_concl)))
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   133
	  (fn prems =>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   134
	   [cut_facts_tac prems 1,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   135
	    REPEAT (eresolve_tac [asm_rl, conjE, PartE, mp] 1
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   136
	     ORELSE resolve_tac [allI, impI, conjI, Part_eqI, refl] 1
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   137
	     ORELSE dresolve_tac [spec, mp, splitD] 1)])
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   138
    handle e => print_sign_exn sign e;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   139
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   140
(*Mutual induction follows by freeness of Inl/Inr.*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   141
1190
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   142
(*Simplification largely reduces the mutual induction rule to the 
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   143
  standard rule*)
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 1190
diff changeset
   144
val mut_ss = simpset_of "Fun"
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 1190
diff changeset
   145
             addsimps [Inl_Inr_eq, Inr_Inl_eq, Inl_eq, Inr_eq];
1190
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   146
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   147
val all_defs = con_defs@part_rec_defs;
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   148
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   149
(*Removes Collects caused by M-operators in the intro rules*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   150
val cmonos = [subset_refl RS Int_Collect_mono] RL monos RLN (2,[rev_subsetD]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   151
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   152
(*Avoids backtracking by delivering the correct premise to each goal*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   153
fun mutual_ind_tac [] 0 = all_tac
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   154
  | mutual_ind_tac(prem::prems) i = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   155
      DETERM
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   156
       (SELECT_GOAL 
1190
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   157
	  (
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   158
	   (*Simplify the assumptions and goal by unfolding Part and
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   159
	     using freeness of the Sum constructors; proves all but one
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   160
             conjunct by contradiction*)
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   161
	   rewrite_goals_tac all_defs  THEN
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   162
	   simp_tac (mut_ss addsimps [Part_def]) 1  THEN
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   163
	   IF_UNSOLVED (*simp_tac may have finished it off!*)
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   164
	     ((*simplify assumptions, but don't accept new rewrite rules!*)
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   165
	      asm_full_simp_tac (mut_ss setmksimps K[]) 1  THEN
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   166
	      (*unpackage and use "prem" in the corresponding place*)
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   167
	      REPEAT (rtac impI 1)  THEN
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   168
	      rtac (rewrite_rule all_defs prem) 1  THEN
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   169
	      (*prem must not be REPEATed below: could loop!*)
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   170
	      DEPTH_SOLVE (FIRSTGOAL (ares_tac [impI] ORELSE' 
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   171
				      eresolve_tac (conjE::mp::cmonos))))
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   172
	  ) i)
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   173
       THEN mutual_ind_tac prems (i-1);
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   174
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   175
val _ = writeln "  Proving the mutual induction rule...";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   176
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   177
val mutual_induct_split = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   178
    prove_goalw_cterm []
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   179
	  (cterm_of sign
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   180
	   (list_implies (map (induct_prem (rec_tms~~preds)) intr_tms,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   181
			  mutual_induct_concl)))
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   182
	  (fn prems =>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   183
	   [rtac (quant_induct RS lemma) 1,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   184
	    mutual_ind_tac (rev prems) (length prems)])
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   185
    handle e => print_sign_exn sign e;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   186
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   187
(*Attempts to remove all occurrences of split*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   188
val split_tac =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   189
    REPEAT (SOMEGOAL (FIRST' [rtac splitI, 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   190
			      dtac splitD,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   191
			      etac splitE,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   192
			      bound_hyp_subst_tac]))
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   193
    THEN prune_params_tac;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   194
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   195
(*strip quantifier*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   196
val induct = standard (quant_induct RS spec RSN (2,rev_mp));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   197
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   198
val mutual_induct = rule_by_tactic split_tac mutual_induct_split;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   199
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   200
end;