src/HOL/intr_elim.ML
author lcp
Tue, 25 Jul 1995 17:01:25 +0200
changeset 1191 d7e0c280f261
parent 923 ff1574a81019
child 1264 3eb91524b938
permissions -rw-r--r--
Added two final lines to intro_tacsf for mutual recursion (borrowed from ZF version)
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/intr_elim.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
Introduction/elimination 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
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     9
signature INDUCTIVE_ARG =	(** Description of a (co)inductive def **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
  sig
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    11
  val thy        : theory               (*new theory with inductive defs*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    12
  val monos      : thm list		(*monotonicity of each M operator*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
  val con_defs   : thm list		(*definitions of the constructors*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    14
  end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    15
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    16
(*internal items*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
signature INDUCTIVE_I =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    18
  sig
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    19
  val rec_tms    : term list		(*the recursive sets*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
  val intr_tms   : term list		(*terms for the introduction rules*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    21
  end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    22
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    23
signature INTR_ELIM =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
  sig
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    25
  val thy        : theory               (*copy of input theory*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26
  val defs	 : thm list		(*definitions made in thy*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    27
  val mono	 : thm			(*monotonicity for the lfp definition*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    28
  val unfold     : thm			(*fixed-point equation*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    29
  val intrs      : thm list		(*introduction rules*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    30
  val elim       : thm			(*case analysis theorem*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    31
  val raw_induct : thm			(*raw induction rule from Fp.induct*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    32
  val mk_cases : thm list -> string -> thm	(*generates case theorems*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    33
  val rec_names  : string list		(*names of recursive sets*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    34
  end;
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
(*prove intr/elim rules for a fixedpoint definition*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    37
functor Intr_elim_Fun
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    38
    (structure Inductive: sig include INDUCTIVE_ARG INDUCTIVE_I end  
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    39
     and Fp: FP) : INTR_ELIM =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    40
struct
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    41
open Logic Inductive Ind_Syntax;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    42
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    43
val rec_names = map (#1 o dest_Const o head_of) rec_tms;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    44
val big_rec_name = space_implode "_" rec_names;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    45
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    46
val _ = deny (big_rec_name  mem  map ! (stamps_of_thy thy))
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    47
             ("Definition " ^ big_rec_name ^ 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    48
	      " would clash with the theory of the same name!");
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    49
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    50
(*fetch fp definitions from the theory*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    51
val big_rec_def::part_rec_defs = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    52
  map (get_def thy)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    53
      (case rec_names of [_] => rec_names | _ => big_rec_name::rec_names);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    54
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    56
val sign = sign_of thy;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    57
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    58
(********)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    59
val _ = writeln "  Proving monotonicity...";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    60
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    61
val Const("==",_) $ _ $ (Const(_,fpT) $ fp_abs) =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    62
    big_rec_def |> rep_thm |> #prop |> unvarify;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    63
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    64
(*For the type of the argument of mono*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    65
val [monoT] = binder_types fpT;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    66
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    67
val mono = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    68
    prove_goalw_cterm [] 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    69
      (cterm_of sign (mk_Trueprop (Const("mono", monoT-->boolT) $ fp_abs)))
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    70
      (fn _ =>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    71
       [rtac monoI 1,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    72
	REPEAT (ares_tac (basic_monos @ monos) 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 unfold = standard (mono RS (big_rec_def RS Fp.Tarski));
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
(********)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    77
val _ = writeln "  Proving the introduction rules...";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    78
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    79
fun intro_tacsf disjIn prems = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    80
  [(*insert prems and underlying sets*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    81
   cut_facts_tac prems 1,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    82
   rtac (unfold RS ssubst) 1,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    83
   REPEAT (resolve_tac [Part_eqI,CollectI] 1),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    84
   (*Now 1-2 subgoals: the disjunction, perhaps equality.*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    85
   rtac disjIn 1,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    86
   (*Not ares_tac, since refl must be tried before any equality assumptions;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    87
     backtracking may occur if the premises have extra variables!*)
1191
d7e0c280f261 Added two final lines to intro_tacsf for mutual recursion
lcp
parents: 923
diff changeset
    88
   DEPTH_SOLVE_1 (resolve_tac [refl,exI,conjI] 1 ORELSE assume_tac 1),
d7e0c280f261 Added two final lines to intro_tacsf for mutual recursion
lcp
parents: 923
diff changeset
    89
   (*Now solve the equations like Inl 0 = Inl ?b2*)
d7e0c280f261 Added two final lines to intro_tacsf for mutual recursion
lcp
parents: 923
diff changeset
    90
   rewrite_goals_tac con_defs,
d7e0c280f261 Added two final lines to intro_tacsf for mutual recursion
lcp
parents: 923
diff changeset
    91
   REPEAT (rtac refl 1)];
d7e0c280f261 Added two final lines to intro_tacsf for mutual recursion
lcp
parents: 923
diff changeset
    92
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    93
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    94
(*combines disjI1 and disjI2 to access the corresponding nested disjunct...*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    95
val mk_disj_rls = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    96
    let fun f rl = rl RS disjI1
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    97
	and g rl = rl RS disjI2
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    98
    in  accesses_bal(f, g, asm_rl)  end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    99
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   100
val intrs = map (uncurry (prove_goalw_cterm part_rec_defs))
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   101
            (map (cterm_of sign) intr_tms ~~ 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   102
	     map intro_tacsf (mk_disj_rls(length intr_tms)));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   103
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   104
(********)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   105
val _ = writeln "  Proving the elimination rule...";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   106
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   107
(*Includes rules for Suc and Pair since they are common constructions*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   108
val elim_rls = [asm_rl, FalseE, Suc_neq_Zero, Zero_neq_Suc,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   109
		make_elim Suc_inject, 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   110
		refl_thin, conjE, exE, disjE];
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   111
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   112
(*Breaks down logical connectives in the monotonic function*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   113
val basic_elim_tac =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   114
    REPEAT (SOMEGOAL (eresolve_tac (elim_rls@sumprod_free_SEs)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   115
	      ORELSE' bound_hyp_subst_tac))
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   116
    THEN prune_params_tac;
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
val elim = rule_by_tactic basic_elim_tac (unfold RS equals_CollectD);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   119
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   120
(*Applies freeness of the given constructors, which *must* be unfolded by
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   121
  the given defs.  Cannot simply use the local con_defs because con_defs=[] 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   122
  for inference systems.
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   123
fun con_elim_tac defs =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   124
    rewrite_goals_tac defs THEN basic_elim_tac THEN fold_tac defs;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   125
 *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   126
fun con_elim_tac simps =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   127
  let val elim_tac = REPEAT o (eresolve_tac (elim_rls@sumprod_free_SEs))
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   128
  in ALLGOALS(EVERY'[elim_tac,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   129
                     asm_full_simp_tac (nat_ss addsimps simps),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   130
                     elim_tac,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   131
                     REPEAT o bound_hyp_subst_tac])
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   132
     THEN prune_params_tac
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   133
  end;
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
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   136
(*String s should have the form t:Si where Si is an inductive set*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   137
fun mk_cases defs s = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   138
    rule_by_tactic (con_elim_tac defs)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   139
      (assume_read thy s  RS  elim);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   140
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   141
val defs = big_rec_def::part_rec_defs;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   142
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   143
val raw_induct = standard ([big_rec_def, mono] MRS Fp.induct);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   144
end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   145