src/HOL/intr_elim.ML
author oheimb
Wed, 12 Nov 1997 12:34:43 +0100
changeset 4206 688050e83d89
parent 4089 96fba19bcbe2
permissions -rw-r--r--
restored last version
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
     1
(*  Title:      HOL/intr_elim.ML
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1425
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
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
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
     9
signature INDUCTIVE_ARG =       (** Description of a (co)inductive def **)
923
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*)
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
    12
  val monos      : thm list             (*monotonicity of each M operator*)
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
    13
  val con_defs   : thm list             (*definitions of the constructors*)
923
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
1425
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
    16
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
    17
signature INDUCTIVE_I = (** Terms read from the theory section **)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    18
  sig
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
    19
  val rec_tms    : term list            (*the recursive sets*)
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
    20
  val intr_tms   : term list            (*terms for the introduction rules*)
923
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*)
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
    26
  val defs       : thm list             (*definitions made in thy*)
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
    27
  val mono       : thm                  (*monotonicity for the lfp definition*)
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
    28
  val intrs      : thm list             (*introduction rules*)
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
    29
  val elim       : thm                  (*case analysis theorem*)
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
    30
  val mk_cases   : thm list -> string -> thm    (*generates case theorems*)
1425
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
    31
  end;
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
    32
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
    33
signature INTR_ELIM_AUX =       (** Used to make induction rules **)
1425
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
    34
  sig
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
    35
  val raw_induct : thm                  (*raw induction rule from Fp.induct*)
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
    36
  val rec_names  : string list          (*names of recursive sets*)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    37
  end;
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
(*prove intr/elim rules for a fixedpoint definition*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    40
functor Intr_elim_Fun
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    41
    (structure Inductive: sig include INDUCTIVE_ARG INDUCTIVE_I end  
1425
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
    42
     and Fp: FP) 
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
    43
    : sig include INTR_ELIM INTR_ELIM_AUX end =
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
    44
let
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
    45
val rec_names = map (#1 o dest_Const o head_of) Inductive.rec_tms;
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    46
val big_rec_name = space_implode "_" rec_names;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    47
3978
7e1cfed19d94 Sign.stamp_names_of;
wenzelm
parents: 2688
diff changeset
    48
val _ = deny (big_rec_name mem (Sign.stamp_names_of (sign_of Inductive.thy)))
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    49
             ("Definition " ^ big_rec_name ^ 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
    50
              " would clash with the theory of the same name!");
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    51
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    52
(*fetch fp definitions from the theory*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    53
val big_rec_def::part_rec_defs = 
1425
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
    54
  map (get_def Inductive.thy)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
      (case rec_names of [_] => rec_names | _ => big_rec_name::rec_names);
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
1425
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
    58
val sign = sign_of Inductive.thy;
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    59
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 _ = writeln "  Proving monotonicity...";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    62
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    63
val Const("==",_) $ _ $ (Const(_,fpT) $ fp_abs) =
1425
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
    64
    big_rec_def |> rep_thm |> #prop |> Logic.unvarify;
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    65
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    66
(*For the type of the argument of mono*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    67
val [monoT] = binder_types fpT;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    68
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    69
val mono = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    70
    prove_goalw_cterm [] 
1425
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
    71
      (cterm_of sign (Ind_Syntax.mk_Trueprop 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
    72
                      (Const("mono", monoT --> Ind_Syntax.boolT) $ fp_abs)))
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    73
      (fn _ =>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    74
       [rtac monoI 1,
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
    75
        REPEAT (ares_tac (basic_monos @ Inductive.monos) 1)]);
923
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 unfold = standard (mono RS (big_rec_def RS Fp.Tarski));
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
(********)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    80
val _ = writeln "  Proving the introduction rules...";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    81
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    82
fun intro_tacsf disjIn prems = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    83
  [(*insert prems and underlying sets*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    84
   cut_facts_tac prems 1,
2031
03a843f0f447 Ran expandshort
paulson
parents: 1465
diff changeset
    85
   stac unfold 1,
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    86
   REPEAT (resolve_tac [Part_eqI,CollectI] 1),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    87
   (*Now 1-2 subgoals: the disjunction, perhaps equality.*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    88
   rtac disjIn 1,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    89
   (*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
    90
     backtracking may occur if the premises have extra variables!*)
2414
13df7d6c5c3b intro_tacsf: replaced ORELSE by APPEND in order to stop
paulson
parents: 2270
diff changeset
    91
   DEPTH_SOLVE_1 (resolve_tac [refl,exI,conjI] 1 APPEND assume_tac 1),
1191
d7e0c280f261 Added two final lines to intro_tacsf for mutual recursion
lcp
parents: 923
diff changeset
    92
   (*Now solve the equations like Inl 0 = Inl ?b2*)
1425
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
    93
   rewrite_goals_tac Inductive.con_defs,
1191
d7e0c280f261 Added two final lines to intro_tacsf for mutual recursion
lcp
parents: 923
diff changeset
    94
   REPEAT (rtac refl 1)];
d7e0c280f261 Added two final lines to intro_tacsf for mutual recursion
lcp
parents: 923
diff changeset
    95
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    96
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    97
(*combines disjI1 and disjI2 to access the corresponding nested disjunct...*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    98
val mk_disj_rls = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    99
    let fun f rl = rl RS disjI1
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
   100
        and g rl = rl RS disjI2
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   101
    in  accesses_bal(f, g, asm_rl)  end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   102
2270
d7513875b2b8 Replaced map...~~ by ListPair.map
paulson
parents: 2031
diff changeset
   103
val intrs = ListPair.map (uncurry (prove_goalw_cterm part_rec_defs))
d7513875b2b8 Replaced map...~~ by ListPair.map
paulson
parents: 2031
diff changeset
   104
            (map (cterm_of sign) Inductive.intr_tms,
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
   105
             map intro_tacsf (mk_disj_rls(length Inductive.intr_tms)));
923
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
(********)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   108
val _ = writeln "  Proving the elimination 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
(*Breaks down logical connectives in the monotonic function*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   111
val basic_elim_tac =
1425
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
   112
    REPEAT (SOMEGOAL (eresolve_tac (Ind_Syntax.elim_rls @ 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
   113
                                    Ind_Syntax.sumprod_free_SEs)
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
   114
              ORELSE' bound_hyp_subst_tac))
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   115
    THEN prune_params_tac;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   116
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   117
(*Applies freeness of the given constructors, which *must* be unfolded by
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   118
  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
   119
  for inference systems.
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   120
fun con_elim_tac defs =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   121
    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
   122
 *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   123
fun con_elim_tac simps =
1425
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
   124
  let val elim_tac = REPEAT o (eresolve_tac (Ind_Syntax.elim_rls @ 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1425
diff changeset
   125
                                             Ind_Syntax.sumprod_free_SEs))
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   126
  in ALLGOALS(EVERY'[elim_tac,
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3978
diff changeset
   127
                     asm_full_simp_tac (simpset_of Nat.thy addsimps simps),
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   128
                     elim_tac,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   129
                     REPEAT o bound_hyp_subst_tac])
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   130
     THEN prune_params_tac
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
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   133
1425
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
   134
in
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
   135
  struct
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
   136
  val thy   = Inductive.thy
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
   137
  and defs  = big_rec_def :: part_rec_defs
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
   138
  and mono  = mono
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
   139
  and intrs = intrs;
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   140
1425
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
   141
  val elim = rule_by_tactic basic_elim_tac 
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
   142
                  (unfold RS Ind_Syntax.equals_CollectD);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   143
1425
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
   144
  (*String s should have the form t:Si where Si is an inductive set*)
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
   145
  fun mk_cases defs s = 
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
   146
      rule_by_tactic (con_elim_tac defs)
2688
889a1cbd1aca rule_by_tactic no longer standardizes its result
paulson
parents: 2414
diff changeset
   147
          (assume_read Inductive.thy s  RS  elim) 
889a1cbd1aca rule_by_tactic no longer standardizes its result
paulson
parents: 2414
diff changeset
   148
      |> standard;
1425
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
   149
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
   150
  val raw_induct = standard ([big_rec_def, mono] MRS Fp.induct)
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
   151
  and rec_names = rec_names
7b61bcfeaa95 Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents: 1264
diff changeset
   152
  end
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   153
end;