src/HOL/indrule.ML
author paulson
Fri, 29 Nov 1996 18:03:21 +0100
changeset 2284 80ebd1a213fd
parent 2270 d7513875b2b8
child 2637 e9b203f854ae
permissions -rw-r--r--
Swapped arguments of Crypt (for clarity and because it is conventional)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
     1
(*  Title:      HOL/indrule.ML
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
923
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
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
    13
  val induct        : thm                       (*main induction rule*)
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
    14
  val mutual_induct : thm                       (*mutual induction rule*)
923
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
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
    20
         Intr_elim: sig include INTR_ELIM INTR_ELIM_AUX end) : INDRULE  =
1424
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
    21
let
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
    22
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
    23
val sign = sign_of Inductive.thy;
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
1424
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
    25
val (Const(_,recT),rec_params) = strip_comb (hd Inductive.rec_tms);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26
1424
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
    27
val elem_type = Ind_Syntax.dest_setT (body_type recT);
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
    28
val big_rec_name = space_implode "_" Intr_elim.rec_names;
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    29
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
    30
1190
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
    31
val _ = writeln "  Proving the induction rule...";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    32
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    33
(*** Prove the main induction rule ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    34
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
    35
val pred_name = "P";            (*name for predicate variables*)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    36
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    37
val big_rec_def::part_rec_defs = Intr_elim.defs;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    38
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    39
(*Used to express induction rules: adds induction hypotheses.
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    40
   ind_alist = [(rec_tm1,pred1),...]  -- associates predicates with rec ops
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    41
   prem is a premise of an intr rule*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    42
fun add_induct_prem ind_alist (prem as Const("Trueprop",_) $ 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
    43
                 (Const("op :",_)$t$X), iprems) =
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    44
     (case gen_assoc (op aconv) (ind_alist, X) of
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
    45
          Some pred => prem :: Ind_Syntax.mk_Trueprop (pred $ t) :: iprems
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
    46
        | None => (*possibly membership in M(rec_tm), for M monotone*)
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
    47
            let fun mk_sb (rec_tm,pred) = 
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
    48
                 (case binder_types (fastype_of pred) of
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
    49
                      [T] => (rec_tm, 
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
    50
                              Ind_Syntax.Int_const T $ rec_tm $ 
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
    51
                                (Ind_Syntax.Collect_const T $ pred))
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
    52
                    | _ => error 
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
    53
                      "Bug: add_induct_prem called with non-unary predicate")
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
    54
            in  subst_free (map mk_sb ind_alist) prem :: iprems  end)
923
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)
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
    61
                         (Logic.strip_imp_prems intr,[])
1424
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
    62
      val (t,X) = Ind_Syntax.rule_concl intr
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    63
      val (Some pred) = gen_assoc (op aconv) (ind_alist, X)
1424
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
    64
      val concl = Ind_Syntax.mk_Trueprop (pred $ t)
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
    65
  in list_all_free (quantfrees, Logic.list_implies (iprems,concl)) end
923
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 = 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
    71
        DEPTH_SOLVE_1 (ares_tac [Part_eqI, prem, refl] i) THEN
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
    72
        ind_tac prems (i-1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    73
1424
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
    74
val pred = Free(pred_name, elem_type --> Ind_Syntax.boolT);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    75
1424
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
    76
val ind_prems = map (induct_prem (map (rpair pred) Inductive.rec_tms)) 
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
    77
                    Inductive.intr_tms;
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    78
1190
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
    79
(*Debugging code...
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
    80
val _ = writeln "ind_prems = ";
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
    81
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
    82
*)
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
    83
1861
505b104f675a Uses minimal simpset (min_ss) and full_simp_tac
paulson
parents: 1746
diff changeset
    84
(*We use a MINIMAL simpset because others (such as HOL_ss) contain too many
505b104f675a Uses minimal simpset (min_ss) and full_simp_tac
paulson
parents: 1746
diff changeset
    85
  simplifications.  If the premises get simplified, then the proofs will 
505b104f675a Uses minimal simpset (min_ss) and full_simp_tac
paulson
parents: 1746
diff changeset
    86
  fail.  This arose with a premise of the form {(F n,G n)|n . True}, which
505b104f675a Uses minimal simpset (min_ss) and full_simp_tac
paulson
parents: 1746
diff changeset
    87
  expanded to something containing ...&True. *)
505b104f675a Uses minimal simpset (min_ss) and full_simp_tac
paulson
parents: 1746
diff changeset
    88
val min_ss = empty_ss
505b104f675a Uses minimal simpset (min_ss) and full_simp_tac
paulson
parents: 1746
diff changeset
    89
      setmksimps (mksimps mksimps_pairs)
505b104f675a Uses minimal simpset (min_ss) and full_simp_tac
paulson
parents: 1746
diff changeset
    90
      setsolver (fn prems => resolve_tac (TrueI::refl::prems) ORELSE' atac
505b104f675a Uses minimal simpset (min_ss) and full_simp_tac
paulson
parents: 1746
diff changeset
    91
                             ORELSE' etac FalseE);
505b104f675a Uses minimal simpset (min_ss) and full_simp_tac
paulson
parents: 1746
diff changeset
    92
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    93
val quant_induct = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    94
    prove_goalw_cterm part_rec_defs 
1424
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
    95
      (cterm_of sign 
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
    96
       (Logic.list_implies (ind_prems, 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
    97
                            Ind_Syntax.mk_Trueprop (Ind_Syntax.mk_all_imp 
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
    98
                                                    (big_rec_tm,pred)))))
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    99
      (fn prems =>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   100
       [rtac (impI RS allI) 1,
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   101
        DETERM (etac Intr_elim.raw_induct 1),
1861
505b104f675a Uses minimal simpset (min_ss) and full_simp_tac
paulson
parents: 1746
diff changeset
   102
        full_simp_tac (min_ss addsimps [Part_Collect]) 1,
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   103
        REPEAT (FIRSTGOAL (eresolve_tac [IntE, CollectE, exE, conjE, disjE] 
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   104
                           ORELSE' hyp_subst_tac)),
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   105
        ind_tac (rev prems) (length prems)])
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   106
    handle e => print_sign_exn sign e;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   107
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   108
(*** Prove the simultaneous induction rule ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   109
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   110
(*Make distinct predicates for each inductive set.
1728
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   111
  Splits cartesian products in elem_type, however nested*)
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   112
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   113
(*The components of the element type, several if it is a product*)
1746
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1728
diff changeset
   114
val elem_factors = Prod_Syntax.factors elem_type;
1728
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   115
val elem_frees = mk_frees "za" elem_factors;
1746
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1728
diff changeset
   116
val elem_tuple = Prod_Syntax.mk_tuple elem_type elem_frees;
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   117
1190
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   118
(*Given a recursive set, return the "split" predicate
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   119
  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
   120
fun mk_predpair rec_tm = 
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   121
  let val rec_name = (#1 o dest_Const o head_of) rec_tm
1728
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   122
      val pfree = Free(pred_name ^ "_" ^ rec_name, 
2031
03a843f0f447 Ran expandshort
paulson
parents: 1985
diff changeset
   123
                       elem_factors ---> Ind_Syntax.boolT)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   124
      val qconcl = 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   125
        foldr Ind_Syntax.mk_all 
1728
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   126
          (elem_frees, 
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   127
           Ind_Syntax.imp $ (Ind_Syntax.mk_mem (elem_tuple, rec_tm))
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   128
                $ (list_comb (pfree, elem_frees)))
1746
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1728
diff changeset
   129
  in  (Prod_Syntax.ap_split elem_type Ind_Syntax.boolT pfree, 
1728
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   130
       qconcl)  
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   131
  end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   132
1424
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
   133
val (preds,qconcls) = split_list (map mk_predpair Inductive.rec_tms);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   134
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   135
(*Used to form simultaneous induction lemma*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   136
fun mk_rec_imp (rec_tm,pred) = 
1424
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
   137
    Ind_Syntax.imp $ (Ind_Syntax.mk_mem (Bound 0, rec_tm)) $  (pred $ Bound 0);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   138
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   139
(*To instantiate the main induction rule*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   140
val induct_concl = 
1424
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
   141
    Ind_Syntax.mk_Trueprop
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
   142
      (Ind_Syntax.mk_all_imp
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
   143
       (big_rec_tm,
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   144
        Abs("z", elem_type, 
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   145
            fold_bal (app Ind_Syntax.conj) 
2270
d7513875b2b8 Replaced map...~~ by ListPair.map
paulson
parents: 2031
diff changeset
   146
            (ListPair.map mk_rec_imp (Inductive.rec_tms,preds)))))
1424
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
   147
and mutual_induct_concl = 
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
   148
    Ind_Syntax.mk_Trueprop (fold_bal (app Ind_Syntax.conj) qconcls);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   149
1728
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   150
val lemma_tac = FIRST' [eresolve_tac [asm_rl, conjE, PartE, mp],
2031
03a843f0f447 Ran expandshort
paulson
parents: 1985
diff changeset
   151
                        resolve_tac [allI, impI, conjI, Part_eqI, refl],
03a843f0f447 Ran expandshort
paulson
parents: 1985
diff changeset
   152
                        dresolve_tac [spec, mp, splitD]];
1728
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   153
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   154
val lemma = (*makes the link between the two induction rules*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   155
    prove_goalw_cterm part_rec_defs 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   156
          (cterm_of sign (Logic.mk_implies (induct_concl,
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   157
                                            mutual_induct_concl)))
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   158
          (fn prems =>
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   159
           [cut_facts_tac prems 1,
1728
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   160
            REPEAT (rewrite_goals_tac [split RS eq_reflection] THEN
2031
03a843f0f447 Ran expandshort
paulson
parents: 1985
diff changeset
   161
                    lemma_tac 1)])
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   162
    handle e => print_sign_exn sign e;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   163
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   164
(*Mutual induction follows by freeness of Inl/Inr.*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   165
1190
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   166
(*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
   167
  standard rule*)
1985
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents: 1867
diff changeset
   168
val mut_ss = min_ss addsimps [Inl_not_Inr, Inr_not_Inl, Inl_eq, Inr_eq, split];
1190
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   169
1728
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   170
val all_defs = [split RS eq_reflection] @ Inductive.con_defs @ part_rec_defs;
1190
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   171
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   172
(*Removes Collects caused by M-operators in the intro rules*)
1424
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
   173
val cmonos = [subset_refl RS Int_Collect_mono] RL Inductive.monos RLN
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
   174
             (2,[rev_subsetD]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   175
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   176
(*Avoids backtracking by delivering the correct premise to each goal*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   177
fun mutual_ind_tac [] 0 = all_tac
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   178
  | mutual_ind_tac(prem::prems) i = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   179
      DETERM
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   180
       (SELECT_GOAL 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   181
          (
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   182
           (*Simplify the assumptions and goal by unfolding Part and
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   183
             using freeness of the Sum constructors; proves all but one
1190
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   184
             conjunct by contradiction*)
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   185
           rewrite_goals_tac all_defs  THEN
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   186
           simp_tac (mut_ss addsimps [Part_def]) 1  THEN
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   187
           IF_UNSOLVED (*simp_tac may have finished it off!*)
1867
37615e73f2d8 corrected comment
paulson
parents: 1861
diff changeset
   188
             ((*simplify assumptions*)
1728
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   189
              full_simp_tac mut_ss 1  THEN
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   190
              (*unpackage and use "prem" in the corresponding place*)
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   191
              REPEAT (rtac impI 1)  THEN
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   192
              rtac (rewrite_rule all_defs prem) 1  THEN
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   193
              (*prem must not be REPEATed below: could loop!*)
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   194
              DEPTH_SOLVE (FIRSTGOAL (ares_tac [impI] ORELSE' 
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   195
                                      eresolve_tac (conjE::mp::cmonos))))
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   196
          ) i)
1190
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   197
       THEN mutual_ind_tac prems (i-1);
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   198
9d1bdce3a38e Old version of mutual induction never worked. Now ensures that
lcp
parents: 923
diff changeset
   199
val _ = writeln "  Proving the mutual induction rule...";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   200
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   201
val mutual_induct_split = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   202
    prove_goalw_cterm []
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   203
          (cterm_of sign
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   204
           (Logic.list_implies (map (induct_prem (Inductive.rec_tms ~~ preds)) 
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   205
                              Inductive.intr_tms,
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   206
                          mutual_induct_concl)))
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   207
          (fn prems =>
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   208
           [rtac (quant_induct RS lemma) 1,
5d7a7e439cec expanded tabs
clasohm
parents: 1424
diff changeset
   209
            mutual_ind_tac (rev prems) (length prems)])
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   210
    handle e => print_sign_exn sign e;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   211
1728
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   212
(** Uncurrying the predicate in the ordinary induction rule **)
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   213
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   214
(*The name "x.1" comes from the "RS spec" !*)
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   215
val xvar = cterm_of sign (Var(("x",1), elem_type));
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   216
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   217
(*strip quantifier and instantiate the variable to a tuple*)
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   218
val induct0 = quant_induct RS spec RSN (2,rev_mp) |>
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   219
              freezeT |>     (*Because elem_type contains TFrees not TVars*)
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   220
              instantiate ([], [(xvar, cterm_of sign elem_tuple)]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   221
1424
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
   222
in
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
   223
  struct
1728
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   224
  val induct = standard 
1746
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1728
diff changeset
   225
                  (Prod_Syntax.split_rule_var
2031
03a843f0f447 Ran expandshort
paulson
parents: 1985
diff changeset
   226
                    (Var((pred_name,2), elem_type --> Ind_Syntax.boolT),
03a843f0f447 Ran expandshort
paulson
parents: 1985
diff changeset
   227
                     induct0));
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   228
1728
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   229
  (*Just "True" unless there's true mutual recursion.  This saves storage.*)
1424
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
   230
  val mutual_induct = 
1728
01beef6262aa Unfolding of arbitrarily nested tuples in induction rules
paulson
parents: 1653
diff changeset
   231
      if length Intr_elim.rec_names > 1
1746
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1728
diff changeset
   232
      then Prod_Syntax.remove_split mutual_induct_split
1424
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
   233
      else TrueI;
ccb3c5ff6707 Now mutual_induct is simply "True" unless it is going to be
paulson
parents: 1264
diff changeset
   234
  end
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   235
end;