src/HOL/Tools/Function/induction_schema.ML
author haftmann
Tue, 24 Nov 2009 17:28:25 +0100
changeset 33955 fff6f11b1f09
parent 33855 cd8acf137c9c
child 34232 36a2a3029fd3
permissions -rw-r--r--
curried take/drop
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33471
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
     1
(*  Title:      HOL/Tools/Function/induction_schema.ML
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
     2
    Author:     Alexander Krauss, TU Muenchen
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
     3
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
     4
A method to prove induction schemas.
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
     5
*)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
     6
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
     7
signature INDUCTION_SCHEMA =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
     8
sig
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
     9
  val mk_ind_tac : (int -> tactic) -> (int -> tactic) -> (int -> tactic)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    10
                   -> Proof.context -> thm list -> tactic
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    11
  val induction_schema_tac : Proof.context -> thm list -> tactic
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    12
  val setup : theory -> theory
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    13
end
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    14
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    15
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    16
structure Induction_Schema : INDUCTION_SCHEMA =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    17
struct
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    18
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    19
open Function_Lib
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    20
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    21
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    22
type rec_call_info = int * (string * typ) list * term list * term list
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    23
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    24
datatype scheme_case =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    25
  SchemeCase of
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    26
  {
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    27
   bidx : int,
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    28
   qs: (string * typ) list,
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    29
   oqnames: string list,
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    30
   gs: term list,
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    31
   lhs: term list,
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    32
   rs: rec_call_info list
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    33
  }
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    34
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    35
datatype scheme_branch = 
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    36
  SchemeBranch of
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    37
  {
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    38
   P : term,
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    39
   xs: (string * typ) list,
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    40
   ws: (string * typ) list,
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    41
   Cs: term list
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    42
  }
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    43
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    44
datatype ind_scheme =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    45
  IndScheme of
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    46
  {
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    47
   T: typ, (* sum of products *)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    48
   branches: scheme_branch list,
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    49
   cases: scheme_case list
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    50
  }
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    51
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    52
val ind_atomize = MetaSimplifier.rewrite true @{thms induct_atomize}
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    53
val ind_rulify = MetaSimplifier.rewrite true @{thms induct_rulify}
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    54
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    55
fun meta thm = thm RS eq_reflection
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    56
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    57
val sum_prod_conv = MetaSimplifier.rewrite true 
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    58
                    (map meta (@{thm split_conv} :: @{thms sum.cases}))
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    59
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    60
fun term_conv thy cv t = 
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    61
    cv (cterm_of thy t)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    62
    |> prop_of |> Logic.dest_equals |> snd
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    63
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    64
fun mk_relT T = HOLogic.mk_setT (HOLogic.mk_prodT (T, T))
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    65
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    66
fun dest_hhf ctxt t = 
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    67
    let 
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    68
      val (ctxt', vars, imp) = dest_all_all_ctx ctxt t
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    69
    in
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    70
      (ctxt', vars, Logic.strip_imp_prems imp, Logic.strip_imp_concl imp)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    71
    end
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    72
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    73
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    74
fun mk_scheme' ctxt cases concl =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    75
    let
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    76
      fun mk_branch concl =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    77
          let
33855
cd8acf137c9c eliminated dead code and some unused bindings, reported by polyml
krauss
parents: 33697
diff changeset
    78
            val (_, ws, Cs, _ $ Pxs) = dest_hhf ctxt concl
33471
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    79
            val (P, xs) = strip_comb Pxs
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    80
          in
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    81
            SchemeBranch { P=P, xs=map dest_Free xs, ws=ws, Cs=Cs }
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    82
          end
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    83
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    84
      val (branches, cases') = (* correction *)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    85
          case Logic.dest_conjunction_list concl of
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    86
            [conc] => 
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    87
            let 
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    88
              val _ $ Pxs = Logic.strip_assums_concl conc
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    89
              val (P, _) = strip_comb Pxs
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    90
              val (cases', conds) = take_prefix (Term.exists_subterm (curry op aconv P)) cases
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    91
              val concl' = fold_rev (curry Logic.mk_implies) conds conc
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    92
            in
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    93
              ([mk_branch concl'], cases')
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    94
            end
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    95
          | concls => (map mk_branch concls, cases)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    96
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    97
      fun mk_case premise =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    98
          let
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    99
            val (ctxt', qs, prems, _ $ Plhs) = dest_hhf ctxt premise
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   100
            val (P, lhs) = strip_comb Plhs
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   101
                                
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   102
            fun bidx Q = find_index (fn SchemeBranch {P=P',...} => Q aconv P') branches
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   103
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   104
            fun mk_rcinfo pr =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   105
                let
33855
cd8acf137c9c eliminated dead code and some unused bindings, reported by polyml
krauss
parents: 33697
diff changeset
   106
                  val (_, Gvs, Gas, _ $ Phyp) = dest_hhf ctxt' pr
33471
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   107
                  val (P', rcs) = strip_comb Phyp
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   108
                in
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   109
                  (bidx P', Gvs, Gas, rcs)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   110
                end
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   111
                
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   112
            fun is_pred v = exists (fn SchemeBranch {P,...} => v aconv P) branches
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   113
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   114
            val (gs, rcprs) = 
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   115
                take_prefix (not o Term.exists_subterm is_pred) prems
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   116
          in
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   117
            SchemeCase {bidx=bidx P, qs=qs, oqnames=map fst qs(*FIXME*), gs=gs, lhs=lhs, rs=map mk_rcinfo rcprs}
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   118
          end
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   119
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   120
      fun PT_of (SchemeBranch { xs, ...}) =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   121
            foldr1 HOLogic.mk_prodT (map snd xs)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   122
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   123
      val ST = Balanced_Tree.make (uncurry SumTree.mk_sumT) (map PT_of branches)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   124
    in
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   125
      IndScheme {T=ST, cases=map mk_case cases', branches=branches }
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   126
    end
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   127
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   128
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   129
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   130
fun mk_completeness ctxt (IndScheme {cases, branches, ...}) bidx =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   131
    let
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   132
      val SchemeBranch { xs, ws, Cs, ... } = nth branches bidx
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   133
      val relevant_cases = filter (fn SchemeCase {bidx=bidx', ...} => bidx' = bidx) cases
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   134
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   135
      val allqnames = fold (fn SchemeCase {qs, ...} => fold (insert (op =) o Free) qs) relevant_cases []
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   136
      val (Pbool :: xs') = map Free (Variable.variant_frees ctxt allqnames (("P", HOLogic.boolT) :: xs))
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   137
      val Cs' = map (Pattern.rewrite_term (ProofContext.theory_of ctxt) (filter_out (op aconv) (map Free xs ~~ xs')) []) Cs
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   138
                       
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   139
      fun mk_case (SchemeCase {qs, oqnames, gs, lhs, ...}) =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   140
          HOLogic.mk_Trueprop Pbool
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   141
                     |> fold_rev (fn x_l => curry Logic.mk_implies (HOLogic.mk_Trueprop(HOLogic.mk_eq x_l)))
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   142
                                 (xs' ~~ lhs)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   143
                     |> fold_rev (curry Logic.mk_implies) gs
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   144
                     |> fold_rev mk_forall_rename (oqnames ~~ map Free qs)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   145
    in
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   146
      HOLogic.mk_Trueprop Pbool
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   147
       |> fold_rev (curry Logic.mk_implies o mk_case) relevant_cases
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   148
       |> fold_rev (curry Logic.mk_implies) Cs'
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   149
       |> fold_rev (Logic.all o Free) ws
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   150
       |> fold_rev mk_forall_rename (map fst xs ~~ xs')
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   151
       |> mk_forall_rename ("P", Pbool)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   152
    end
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   153
33855
cd8acf137c9c eliminated dead code and some unused bindings, reported by polyml
krauss
parents: 33697
diff changeset
   154
fun mk_wf R (IndScheme {T, ...}) =
33471
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   155
    HOLogic.Trueprop $ (Const (@{const_name wf}, mk_relT T --> HOLogic.boolT) $ R)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   156
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   157
fun mk_ineqs R (IndScheme {T, cases, branches}) =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   158
    let
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   159
      fun inject i ts =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   160
          SumTree.mk_inj T (length branches) (i + 1) (foldr1 HOLogic.mk_prod ts)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   161
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   162
      val thesis = Free ("thesis", HOLogic.boolT) (* FIXME *)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   163
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   164
      fun mk_pres bdx args = 
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   165
          let
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   166
            val SchemeBranch { xs, ws, Cs, ... } = nth branches bdx
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   167
            fun replace (x, v) t = betapply (lambda (Free x) t, v)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   168
            val Cs' = map (fold replace (xs ~~ args)) Cs
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   169
            val cse = 
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   170
                HOLogic.mk_Trueprop thesis
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   171
                |> fold_rev (curry Logic.mk_implies) Cs'
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   172
                |> fold_rev (Logic.all o Free) ws
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   173
          in
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   174
            Logic.mk_implies (cse, HOLogic.mk_Trueprop thesis)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   175
          end
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   176
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   177
      fun f (SchemeCase {bidx, qs, oqnames, gs, lhs, rs, ...}) = 
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   178
          let
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   179
            fun g (bidx', Gvs, Gas, rcarg) =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   180
                let val export = 
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   181
                         fold_rev (curry Logic.mk_implies) Gas
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   182
                         #> fold_rev (curry Logic.mk_implies) gs
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   183
                         #> fold_rev (Logic.all o Free) Gvs
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   184
                         #> fold_rev mk_forall_rename (oqnames ~~ map Free qs)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   185
                in
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   186
                (HOLogic.mk_mem (HOLogic.mk_prod (inject bidx' rcarg, inject bidx lhs), R)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   187
                 |> HOLogic.mk_Trueprop
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   188
                 |> export,
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   189
                 mk_pres bidx' rcarg
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   190
                 |> export
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   191
                 |> Logic.all thesis)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   192
                end
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   193
          in
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   194
            map g rs
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   195
          end
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   196
    in
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   197
      map f cases
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   198
    end
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   199
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   200
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   201
fun mk_ind_goal thy branches =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   202
    let
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   203
      fun brnch (SchemeBranch { P, xs, ws, Cs, ... }) =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   204
          HOLogic.mk_Trueprop (list_comb (P, map Free xs))
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   205
          |> fold_rev (curry Logic.mk_implies) Cs
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   206
          |> fold_rev (Logic.all o Free) ws
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   207
          |> term_conv thy ind_atomize
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   208
          |> ObjectLogic.drop_judgment thy
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   209
          |> tupled_lambda (foldr1 HOLogic.mk_prod (map Free xs))
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   210
    in
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   211
      SumTree.mk_sumcases HOLogic.boolT (map brnch branches)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   212
    end
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   213
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   214
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   215
fun mk_induct_rule ctxt R x complete_thms wf_thm ineqss (IndScheme {T, cases=scases, branches}) =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   216
    let
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   217
      val n = length branches
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   218
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   219
      val scases_idx = map_index I scases
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   220
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   221
      fun inject i ts =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   222
          SumTree.mk_inj T n (i + 1) (foldr1 HOLogic.mk_prod ts)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   223
      val P_of = nth (map (fn (SchemeBranch { P, ... }) => P) branches)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   224
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   225
      val thy = ProofContext.theory_of ctxt
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   226
      val cert = cterm_of thy 
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   227
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   228
      val P_comp = mk_ind_goal thy branches
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   229
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   230
      (* Inductive Hypothesis: !!z. (z,x):R ==> P z *)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   231
      val ihyp = Term.all T $ Abs ("z", T, 
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   232
               Logic.mk_implies
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   233
                 (HOLogic.mk_Trueprop (
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   234
                  Const ("op :", HOLogic.mk_prodT (T, T) --> mk_relT T --> HOLogic.boolT) 
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   235
                    $ (HOLogic.pair_const T T $ Bound 0 $ x) 
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   236
                    $ R),
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   237
                   HOLogic.mk_Trueprop (P_comp $ Bound 0)))
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   238
           |> cert
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   239
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   240
      val aihyp = assume ihyp
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   241
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   242
     (* Rule for case splitting along the sum types *)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   243
      val xss = map (fn (SchemeBranch { xs, ... }) => map Free xs) branches
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   244
      val pats = map_index (uncurry inject) xss
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   245
      val sum_split_rule = Pat_Completeness.prove_completeness thy [x] (P_comp $ x) xss (map single pats)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   246
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   247
      fun prove_branch (bidx, (SchemeBranch { P, xs, ws, Cs, ... }, (complete_thm, pat))) =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   248
          let
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   249
            val fxs = map Free xs
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   250
            val branch_hyp = assume (cert (HOLogic.mk_Trueprop (HOLogic.mk_eq (x, pat))))
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   251
                             
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   252
            val C_hyps = map (cert #> assume) Cs
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   253
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   254
            val (relevant_cases, ineqss') = filter (fn ((_, SchemeCase {bidx=bidx', ...}), _) => bidx' = bidx) (scases_idx ~~ ineqss)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   255
                                            |> split_list
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   256
                           
33855
cd8acf137c9c eliminated dead code and some unused bindings, reported by polyml
krauss
parents: 33697
diff changeset
   257
            fun prove_case (cidx, SchemeCase {qs, gs, lhs, rs, ...}) ineq_press =
33471
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   258
                let
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   259
                  val case_hyps = map (assume o cert o HOLogic.mk_Trueprop o HOLogic.mk_eq) (fxs ~~ lhs)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   260
                           
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   261
                  val cqs = map (cert o Free) qs
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   262
                  val ags = map (assume o cert) gs
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   263
                            
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   264
                  val replace_x_ss = HOL_basic_ss addsimps (branch_hyp :: case_hyps)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   265
                  val sih = full_simplify replace_x_ss aihyp
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   266
                            
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   267
                  fun mk_Prec (idx, Gvs, Gas, rcargs) (ineq, pres) =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   268
                      let
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   269
                        val cGas = map (assume o cert) Gas
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   270
                        val cGvs = map (cert o Free) Gvs
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   271
                        val import = fold forall_elim (cqs @ cGvs)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   272
                                     #> fold Thm.elim_implies (ags @ cGas)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   273
                        val ipres = pres
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   274
                                     |> forall_elim (cert (list_comb (P_of idx, rcargs)))
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   275
                                     |> import
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   276
                      in
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   277
                        sih |> forall_elim (cert (inject idx rcargs))
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   278
                            |> Thm.elim_implies (import ineq) (* Psum rcargs *)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   279
                            |> Conv.fconv_rule sum_prod_conv
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   280
                            |> Conv.fconv_rule ind_rulify
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   281
                            |> (fn th => th COMP ipres) (* P rs *)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   282
                            |> fold_rev (implies_intr o cprop_of) cGas
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   283
                            |> fold_rev forall_intr cGvs
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   284
                      end
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   285
                      
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   286
                  val P_recs = map2 mk_Prec rs ineq_press   (*  [P rec1, P rec2, ... ]  *)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   287
                               
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   288
                  val step = HOLogic.mk_Trueprop (list_comb (P, lhs))
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   289
                             |> fold_rev (curry Logic.mk_implies o prop_of) P_recs
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   290
                             |> fold_rev (curry Logic.mk_implies) gs
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   291
                             |> fold_rev (Logic.all o Free) qs
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   292
                             |> cert
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   293
                             
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   294
                  val Plhs_to_Pxs_conv = 
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   295
                      foldl1 (uncurry Conv.combination_conv) 
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   296
                      (Conv.all_conv :: map (fn ch => K (Thm.symmetric (ch RS eq_reflection))) case_hyps)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   297
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   298
                  val res = assume step
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   299
                                   |> fold forall_elim cqs
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   300
                                   |> fold Thm.elim_implies ags
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   301
                                   |> fold Thm.elim_implies P_recs (* P lhs *) 
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   302
                                   |> Conv.fconv_rule (Conv.arg_conv Plhs_to_Pxs_conv) (* P xs *)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   303
                                   |> fold_rev (implies_intr o cprop_of) (ags @ case_hyps)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   304
                                   |> fold_rev forall_intr cqs (* !!qs. Gas ==> xs = lhss ==> P xs *)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   305
                in
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   306
                  (res, (cidx, step))
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   307
                end
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   308
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   309
            val (cases, steps) = split_list (map2 prove_case relevant_cases ineqss')
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   310
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   311
            val bstep = complete_thm
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   312
                |> forall_elim (cert (list_comb (P, fxs)))
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   313
                |> fold (forall_elim o cert) (fxs @ map Free ws)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   314
                |> fold Thm.elim_implies C_hyps             (* FIXME: optimization using rotate_prems *)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   315
                |> fold Thm.elim_implies cases (* P xs *)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   316
                |> fold_rev (implies_intr o cprop_of) C_hyps
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   317
                |> fold_rev (forall_intr o cert o Free) ws
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   318
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   319
            val Pxs = cert (HOLogic.mk_Trueprop (P_comp $ x))
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   320
                     |> Goal.init
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   321
                     |> (MetaSimplifier.rewrite_goals_tac (map meta (branch_hyp :: @{thm split_conv} :: @{thms sum.cases}))
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   322
                         THEN CONVERSION ind_rulify 1)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   323
                     |> Seq.hd
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   324
                     |> Thm.elim_implies (Conv.fconv_rule Drule.beta_eta_conversion bstep)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   325
                     |> Goal.finish ctxt
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   326
                     |> implies_intr (cprop_of branch_hyp)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   327
                     |> fold_rev (forall_intr o cert) fxs
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   328
          in
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   329
            (Pxs, steps)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   330
          end
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   331
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   332
      val (branches, steps) = split_list (map_index prove_branch (branches ~~ (complete_thms ~~ pats)))
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   333
                              |> apsnd flat
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   334
                           
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   335
      val istep = sum_split_rule
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   336
                |> fold (fn b => fn th => Drule.compose_single (b, 1, th)) branches
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   337
                |> implies_intr ihyp
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   338
                |> forall_intr (cert x) (* "!!x. (!!y<x. P y) ==> P x" *)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   339
         
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   340
      val induct_rule =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   341
          @{thm "wf_induct_rule"}
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   342
            |> (curry op COMP) wf_thm 
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   343
            |> (curry op COMP) istep
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   344
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   345
      val steps_sorted = map snd (sort (int_ord o pairself fst) steps)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   346
    in
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   347
      (steps_sorted, induct_rule)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   348
    end
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   349
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   350
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   351
fun mk_ind_tac comp_tac pres_tac term_tac ctxt facts = (ALLGOALS (Method.insert_tac facts)) THEN HEADGOAL 
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   352
(SUBGOAL (fn (t, i) =>
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   353
  let
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   354
    val (ctxt', _, cases, concl) = dest_hhf ctxt t
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   355
    val scheme as IndScheme {T=ST, branches, ...} = mk_scheme' ctxt' cases concl
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   356
(*     val _ = tracing (makestring scheme)*)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   357
    val ([Rn,xn], ctxt'') = Variable.variant_fixes ["R","x"] ctxt'
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   358
    val R = Free (Rn, mk_relT ST)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   359
    val x = Free (xn, ST)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   360
    val cert = cterm_of (ProofContext.theory_of ctxt)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   361
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   362
    val ineqss = mk_ineqs R scheme
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   363
                   |> map (map (pairself (assume o cert)))
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   364
    val complete = map_range (mk_completeness ctxt scheme #> cert #> assume) (length branches)
33855
cd8acf137c9c eliminated dead code and some unused bindings, reported by polyml
krauss
parents: 33697
diff changeset
   365
    val wf_thm = mk_wf R scheme |> cert |> assume
33471
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   366
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   367
    val (descent, pres) = split_list (flat ineqss)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   368
    val newgoals = complete @ pres @ wf_thm :: descent 
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   369
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   370
    val (steps, indthm) = mk_induct_rule ctxt'' R x complete wf_thm ineqss scheme
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   371
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   372
    fun project (i, SchemeBranch {xs, ...}) =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   373
        let
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   374
          val inst = cert (SumTree.mk_inj ST (length branches) (i + 1) (foldr1 HOLogic.mk_prod (map Free xs)))
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   375
        in
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   376
          indthm |> Drule.instantiate' [] [SOME inst]
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   377
                 |> simplify SumTree.sumcase_split_ss
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   378
                 |> Conv.fconv_rule ind_rulify
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   379
(*                 |> (fn thm => (tracing (makestring thm); thm))*)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   380
        end                  
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   381
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   382
    val res = Conjunction.intr_balanced (map_index project branches)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   383
                 |> fold_rev implies_intr (map cprop_of newgoals @ steps)
33697
wenzelm
parents: 33471
diff changeset
   384
                 |> Drule.generalize ([], [Rn])
33471
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   385
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   386
    val nbranches = length branches
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   387
    val npres = length pres
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   388
  in
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   389
    Thm.compose_no_flatten false (res, length newgoals) i
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   390
    THEN term_tac (i + nbranches + npres)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   391
    THEN (EVERY (map (TRY o pres_tac) ((i + nbranches + npres - 1) downto (i + nbranches))))
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   392
    THEN (EVERY (map (TRY o comp_tac) ((i + nbranches - 1) downto i)))
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   393
  end))
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   394
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   395
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   396
fun induction_schema_tac ctxt =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   397
  mk_ind_tac (K all_tac) (assume_tac APPEND' Goal.assume_rule_tac ctxt) (K all_tac) ctxt;
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   398
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   399
val setup =
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   400
  Method.setup @{binding induction_schema} (Scan.succeed (RAW_METHOD o induction_schema_tac))
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   401
    "proves an induction principle"
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   402
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
   403
end