ProgTutorial/Package/simple_inductive_package.ML
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Sun, 15 Dec 2013 23:49:05 +0000
changeset 552 82c482467d75
parent 535 5734ab5dd86d
child 562 daf404920ab9
permissions -rw-r--r--
updated to latest isabelle
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     1
(* @chunk SIMPLE_INDUCTIVE_PACKAGE *)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     2
signature SIMPLE_INDUCTIVE_PACKAGE =
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     3
sig
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     4
  val add_inductive_i:
76
b99fa5fa63fc adapted to changes in binding.ML
Christian Urban <urbanc@in.tum.de>
parents: 55
diff changeset
     5
    ((Binding.binding * typ) * mixfix) list ->  (*{predicates}*)
b99fa5fa63fc adapted to changes in binding.ML
Christian Urban <urbanc@in.tum.de>
parents: 55
diff changeset
     6
    (Binding.binding * typ) list ->  (*{parameters}*)
121
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
     7
    (Attrib.binding * term) list ->  (*{rules}*)
110
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
     8
    local_theory -> local_theory
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 111
diff changeset
     9
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    10
  val add_inductive:
76
b99fa5fa63fc adapted to changes in binding.ML
Christian Urban <urbanc@in.tum.de>
parents: 55
diff changeset
    11
    (Binding.binding * string option * mixfix) list ->  (*{predicates}*)
b99fa5fa63fc adapted to changes in binding.ML
Christian Urban <urbanc@in.tum.de>
parents: 55
diff changeset
    12
    (Binding.binding * string option * mixfix) list ->  (*{parameters}*)
110
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    13
    (Attrib.binding * string) list ->  (*{rules}*)
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 91
diff changeset
    14
    local_theory -> local_theory
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    15
end;
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    16
(* @end *)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    17
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    18
structure SimpleInductivePackage: SIMPLE_INDUCTIVE_PACKAGE =
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    19
struct
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    20
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    21
(* @chunk make_definitions *) 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    22
fun make_defs ((binding, syn), trm) lthy =
111
3798baeee55f rearranged some functions
Christian Urban <urbanc@in.tum.de>
parents: 110
diff changeset
    23
let 
514
7e25716c3744 updated to outer syntax / parser changes
Christian Urban <urbanc@in.tum.de>
parents: 475
diff changeset
    24
  val arg = ((binding, syn), ((Binding.suffix_name "_def" binding, []), trm))
401
36d61044f9bf updated to new Isabelle and clarified Skip_Proof
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
    25
  val ((_, (_ , thm)), lthy) = Local_Theory.define arg lthy
111
3798baeee55f rearranged some functions
Christian Urban <urbanc@in.tum.de>
parents: 110
diff changeset
    26
in 
3798baeee55f rearranged some functions
Christian Urban <urbanc@in.tum.de>
parents: 110
diff changeset
    27
  (thm, lthy) 
3798baeee55f rearranged some functions
Christian Urban <urbanc@in.tum.de>
parents: 110
diff changeset
    28
end
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 116
diff changeset
    29
(* @end *)
111
3798baeee55f rearranged some functions
Christian Urban <urbanc@in.tum.de>
parents: 110
diff changeset
    30
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    31
(* @chunk definitions_aux *) 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    32
fun defs_aux lthy orules preds params (pred, arg_types) =
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    33
let 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    34
  fun mk_all x P = HOLogic.all_const (fastype_of x) $ lambda x P 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    35
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    36
  val fresh_args = 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    37
        arg_types 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    38
        |> map (pair "z")
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    39
        |> Variable.variant_frees lthy orules 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    40
        |> map Free
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    41
in
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    42
  list_comb (pred, fresh_args)
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    43
  |> fold_rev (curry HOLogic.mk_imp) orules
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    44
  |> fold_rev mk_all preds
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    45
  |> fold_rev lambda (params @ fresh_args) 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    46
end
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    47
(* @end *)
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    48
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 116
diff changeset
    49
(* @chunk definitions *) 
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    50
fun definitions rules params preds prednames syns arg_typess lthy =
111
3798baeee55f rearranged some functions
Christian Urban <urbanc@in.tum.de>
parents: 110
diff changeset
    51
let
475
25371f74c768 updated to post-2011-1 Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 426
diff changeset
    52
  val thy = Proof_Context.theory_of lthy
418
1d1e4cda8c54 updated to new isabelle
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
    53
  val orules = map (Object_Logic.atomize_term thy) rules
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    54
  val defs = 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    55
    map (defs_aux lthy orules preds params) (preds ~~ arg_typess) 
111
3798baeee55f rearranged some functions
Christian Urban <urbanc@in.tum.de>
parents: 110
diff changeset
    56
in
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    57
  fold_map make_defs (prednames ~~ syns ~~ defs) lthy
111
3798baeee55f rearranged some functions
Christian Urban <urbanc@in.tum.de>
parents: 110
diff changeset
    58
end
3798baeee55f rearranged some functions
Christian Urban <urbanc@in.tum.de>
parents: 110
diff changeset
    59
(* @end *)
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 91
diff changeset
    60
110
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    61
fun inst_spec ct = 
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    62
  Drule.instantiate' [SOME (ctyp_of_term ct)] [NONE, SOME ct] @{thm spec};
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 91
diff changeset
    63
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    64
(* @chunk induction_tac *)
552
82c482467d75 updated to latest isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 535
diff changeset
    65
fun induction_tac ctxt defs prems insts =
82c482467d75 updated to latest isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 535
diff changeset
    66
  EVERY1 [Object_Logic.full_atomize_tac ctxt,
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    67
          cut_facts_tac prems,
552
82c482467d75 updated to latest isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 535
diff changeset
    68
          rewrite_goal_tac ctxt defs,
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    69
          EVERY' (map (dtac o inst_spec) insts),
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    70
          assume_tac]
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    71
(* @end *)
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    72
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    73
(* @chunk induction_rules *)
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    74
fun inductions rules defs parnames preds Tss lthy1  =
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    75
let
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    76
  val (_, lthy2) = Variable.add_fixes parnames lthy1
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    77
  
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    78
  val Ps = replicate (length preds) "P"
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    79
  val (newprednames, lthy3) = Variable.variant_fixes Ps lthy2
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    80
475
25371f74c768 updated to post-2011-1 Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 426
diff changeset
    81
  val thy = Proof_Context.theory_of lthy3			      
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    82
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    83
  val Tss' = map (fn Ts => Ts ---> HOLogic.boolT) Tss
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    84
  val newpreds = map Free (newprednames ~~ Tss')
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    85
  val cnewpreds = map (cterm_of thy) newpreds
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    86
  val rules' = map (subst_free (preds ~~ newpreds)) rules
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    87
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    88
  fun prove_induction ((pred, newpred), Ts)  =
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    89
  let
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    90
    val (newargnames, lthy4) = 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    91
          Variable.variant_fixes (replicate (length Ts) "z") lthy3;
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    92
    val newargs = map Free (newargnames ~~ Ts)
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    93
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    94
    val prem = HOLogic.mk_Trueprop (list_comb (pred, newargs))
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    95
    val goal = Logic.list_implies 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    96
         (rules', HOLogic.mk_Trueprop (list_comb (newpred, newargs)))
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    97
  in
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    98
    Goal.prove lthy4 [] [prem] goal
552
82c482467d75 updated to latest isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 535
diff changeset
    99
      (fn {prems, context, ...} => induction_tac context defs prems cnewpreds)
475
25371f74c768 updated to post-2011-1 Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 426
diff changeset
   100
      |> singleton (Proof_Context.export lthy4 lthy1)
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   101
  end 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   102
in
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   103
  map prove_induction (preds ~~ newpreds ~~ Tss)
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   104
end
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   105
(* @end *)
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   106
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 91
diff changeset
   107
val all_elims = fold (fn ct => fn th => th RS inst_spec ct);
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 91
diff changeset
   108
val imp_elims = fold (fn th => fn th' => [th', th] MRS @{thm mp});
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 91
diff changeset
   109
111
3798baeee55f rearranged some functions
Christian Urban <urbanc@in.tum.de>
parents: 110
diff changeset
   110
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   111
(* @chunk subproof1 *) 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   112
fun subproof2 prem params2 prems2 =  
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   113
 SUBPROOF (fn {prems, ...} =>
124
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   114
  let
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   115
    val prem' = prems MRS prem;
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   116
    val prem'' = 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   117
       case prop_of prem' of
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   118
           _ $ (Const (@{const_name All}, _) $ _) =>
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   119
             prem' |> all_elims params2 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   120
                   |> imp_elims prems2
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   121
         | _ => prem';
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   122
  in 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   123
    rtac prem'' 1 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   124
  end)
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   125
(* @end *)
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   126
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   127
(* @chunk subproof2 *) 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   128
fun subproof1 rules preds i = 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   129
 SUBPROOF (fn {params, prems, context = ctxt', ...} =>
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   130
  let
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   131
    val (prems1, prems2) = chop (length prems - length rules) prems;
294
ee9d53fbb56b made changes for SUBPROOF and sat_tac
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   132
    val (params1, params2) = chop (length params - length preds) (map snd params);
124
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   133
  in
552
82c482467d75 updated to latest isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 535
diff changeset
   134
    rtac (Object_Logic.rulify ctxt' (all_elims params1 (nth prems2 i))) 1 
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   135
    THEN
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   136
    EVERY1 (map (fn prem => subproof2 prem params2 prems2 ctxt') prems1)
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   137
  end)
110
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   138
(* @end *)
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   139
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   140
fun introductions_tac defs rules preds i ctxt =
552
82c482467d75 updated to latest isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 535
diff changeset
   141
  EVERY1 [Object_Logic.rulify_tac ctxt,
82c482467d75 updated to latest isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 535
diff changeset
   142
          rewrite_goal_tac ctxt defs,
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   143
          REPEAT o (resolve_tac [@{thm allI},@{thm impI}]),
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   144
          subproof1 rules preds i ctxt]
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   145
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   146
110
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   147
(* @chunk intro_rules *) 
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   148
fun introductions rules parnames preds defs lthy1 = 
110
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   149
let
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   150
  val (_, lthy2) = Variable.add_fixes parnames lthy1
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   151
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   152
  fun prove_intro (i, goal) =
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   153
    Goal.prove lthy2 [] [] goal
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   154
       (fn {context, ...} => introductions_tac defs rules preds i context)
475
25371f74c768 updated to post-2011-1 Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 426
diff changeset
   155
       |> singleton (Proof_Context.export lthy2 lthy1)
110
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   156
in
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   157
  map_index prove_intro rules
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   158
end
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   159
(* @end *)
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   160
110
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   161
(* @chunk add_inductive_i *)
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   162
fun add_inductive_i preds params specs lthy =
124
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   163
let
159
64fa844064fa updated to chages in binding module
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
   164
  val params' = map (fn (p, T) => Free (Binding.name_of p, T)) params;
64fa844064fa updated to chages in binding module
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
   165
  val preds' = map (fn ((R, T), _) => list_comb (Free (Binding.name_of R, T), params')) preds;
124
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   166
  val Tss = map (binder_types o fastype_of) preds';   
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   167
  val (ass, rules) = split_list specs;  (* FIXME: ass not used? *)  
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   168
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   169
  val prednames = map (fst o fst) preds
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   170
  val syns = map snd preds
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   171
  val parnames = map (Binding.name_of o fst) params
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   172
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   173
  val (defs, lthy1) = definitions rules params' preds' prednames syns Tss lthy;
110
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   174
      
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   175
  val inducts = inductions rules defs parnames preds' Tss lthy1 	
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   176
  
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   177
  val intros = introductions rules parnames preds' defs lthy1
110
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   178
159
64fa844064fa updated to chages in binding module
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
   179
  val mut_name = space_implode "_" (map (Binding.name_of o fst o fst) preds);
64fa844064fa updated to chages in binding module
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
   180
  val case_names = map (Binding.name_of o fst o fst) specs
124
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   181
in
110
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   182
    lthy1 
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 375
diff changeset
   183
    |> Local_Theory.notes (map (fn (((a, atts), _), th) =>
159
64fa844064fa updated to chages in binding module
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
   184
        ((Binding.qualify false mut_name a, atts), [([th], [])])) (specs ~~ intros)) 
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 375
diff changeset
   185
    |-> (fn intross => Local_Theory.note 
159
64fa844064fa updated to chages in binding module
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
   186
         ((Binding.qualify false mut_name (Binding.name "intros"), []), maps snd intross)) 
110
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   187
    |>> snd 
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 375
diff changeset
   188
    ||>> (Local_Theory.notes (map (fn (((R, _), _), th) =>
159
64fa844064fa updated to chages in binding module
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
   189
         ((Binding.qualify false (Binding.name_of R) (Binding.name "induct"),
375
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
   190
          [Attrib.internal (K (Rule_Cases.case_names case_names)),
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
   191
           Attrib.internal (K (Rule_Cases.consumes 1)),
110
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   192
           Attrib.internal (K (Induct.induct_pred ""))]), [([th], [])]))
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   193
          (preds ~~ inducts)) #>> maps snd) 
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   194
    |> snd
124
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   195
end
110
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   196
(* @end *)
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   197
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   198
(* @chunk read_specification *)
110
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   199
fun read_specification' vars specs lthy =
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   200
let 
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   201
  val specs' = map (fn (a, s) => (a, [s])) specs
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   202
  val ((varst, specst), _) = 
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   203
                   Specification.read_specification vars specs' lthy
110
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   204
  val specst' = map (apsnd the_single) specst
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   205
in   
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   206
  (varst, specst')
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   207
end 
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   208
(* @end *)
110
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   209
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   210
(* @chunk add_inductive *)
110
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   211
fun add_inductive preds params specs lthy =
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   212
let
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   213
  val (vars, specs') = read_specification' (preds @ params) specs lthy;
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   214
  val (preds', params') = chop (length preds) vars;
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   215
  val params'' = map fst params'
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   216
in
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   217
  add_inductive_i preds' params'' specs' lthy
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   218
end;
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   219
(* @end *)
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   220
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 111
diff changeset
   221
(* @chunk parser *)
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   222
val spec_parser = 
426
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 418
diff changeset
   223
   Parse.opt_target --
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 418
diff changeset
   224
   Parse.fixes -- 
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 418
diff changeset
   225
   Parse.for_fixes --
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 91
diff changeset
   226
   Scan.optional 
426
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 418
diff changeset
   227
     (Parse.$$$ "where" |--
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 418
diff changeset
   228
        Parse.!!! 
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 418
diff changeset
   229
          (Parse.enum1 "|" 
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 418
diff changeset
   230
             (Parse_Spec.opt_thm_name ":" -- Parse.prop))) []
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 111
diff changeset
   231
(* @end *)
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   232
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 111
diff changeset
   233
(* @chunk syntax *)
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   234
val specification =
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   235
  spec_parser >>
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 91
diff changeset
   236
    (fn (((loc, preds), params), specs) =>
110
12533bb49615 recovered old version of simple_induct; split the main function into small functions
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   237
      Toplevel.local_theory loc (add_inductive preds params specs))
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   238
535
5734ab5dd86d adapted to new build framework
Christian Urban <urbanc@in.tum.de>
parents: 514
diff changeset
   239
val _ = Outer_Syntax.command @{command_spec "simple_inductive"} "define inductive predicates"
514
7e25716c3744 updated to outer syntax / parser changes
Christian Urban <urbanc@in.tum.de>
parents: 475
diff changeset
   240
          specification
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   241
(* @end *)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   242
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   243
end;