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