Nominal/nominal_inductive.ML
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Fri, 19 Aug 2011 10:56:12 +0900
changeset 2994 4ee772b12032
parent 2987 27aab7a105eb
child 3045 d0ad264f8c4f
permissions -rw-r--r--
Update to new Isabelle
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
(*  Title:      nominal_inductive.ML
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
    Author:     Christian Urban
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
    Infrastructure for proving strong induction theorems
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
    for inductive predicates involving nominal datatypes.
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
    Code based on an earlier version by Stefan Berghofer.
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
*)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
signature NOMINAL_INDUCTIVE =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
sig
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
  val prove_strong_inductive: string list -> string list -> term list list -> thm -> thm list -> 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
    Proof.context -> Proof.state
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
  val prove_strong_inductive_cmd: xstring * (string * string list) list -> Proof.context -> Proof.state
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
structure Nominal_Inductive : NOMINAL_INDUCTIVE =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
struct
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
fun mk_cplus p q = Thm.capply (Thm.capply @{cterm "plus :: perm => perm => perm"} p) q 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
fun mk_cminus p = Thm.capply @{cterm "uminus :: perm => perm"} p 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
fun minus_permute_intro_tac p = 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
  rtac (Drule.instantiate' [] [SOME (mk_cminus p)] @{thm permute_boolE})
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
fun minus_permute_elim p thm = 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
  thm RS (Drule.instantiate' [] [NONE, SOME (mk_cminus p)] @{thm permute_boolI})
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
2680
cd5614027c53 removed diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2645
diff changeset
    33
(* fixme: move to nominal_library *)
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
fun real_head_of (@{term Trueprop} $ t) = real_head_of t
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
  | real_head_of (Const ("==>", _) $ _ $ t) = real_head_of t
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
  | real_head_of (Const (@{const_name all}, _) $ Abs (_, _, t)) = real_head_of t
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
  | real_head_of (Const (@{const_name All}, _) $ Abs (_, _, t)) = real_head_of t
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
  | real_head_of (Const ("HOL.induct_forall", _) $ Abs (_, _, t)) = real_head_of t
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
  | real_head_of t = head_of t  
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
fun mk_vc_compat (avoid, avoid_trm) prems concl_args params = 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
  let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
    val vc_goal = concl_args
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
      |> HOLogic.mk_tuple
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
      |> mk_fresh_star avoid_trm 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
      |> HOLogic.mk_Trueprop
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
      |> (curry Logic.list_implies) prems
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
      |> (curry list_all_free) params
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
    val finite_goal = avoid_trm
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
      |> mk_finite
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
      |> HOLogic.mk_Trueprop
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
      |> (curry Logic.list_implies) prems
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
      |> (curry list_all_free) params
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
  in 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
    if null avoid then [] else [vc_goal, finite_goal]
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
  end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
2680
cd5614027c53 removed diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2645
diff changeset
    59
(* fixme: move to nominal_library *)
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
fun map_term prop f trm =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
  if prop trm 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
  then f trm
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
  else case trm of
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
    (t1 $ t2) => map_term prop f t1 $ map_term prop f t2
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
  | Abs (x, T, t) => Abs (x, T, map_term prop f t)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
  | _ => trm
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
fun add_p_c p (c, c_ty) trm =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
  let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
    val (P, args) = strip_comb trm
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
    val (P_name, P_ty) = dest_Free P
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
    val (ty_args, bool) = strip_type P_ty
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
    val args' = map (mk_perm p) args
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
  in
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
    list_comb (Free (P_name, (c_ty :: ty_args) ---> bool),  c :: args')
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
    |> (fn t => HOLogic.all_const c_ty $ lambda c t )
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
    |> (fn t => HOLogic.all_const @{typ perm} $  lambda p t)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
  end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
fun induct_forall_const T = Const ("HOL.induct_forall", (T --> @{typ bool}) --> @{typ bool})
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
fun mk_induct_forall (a, T) t =  induct_forall_const T $ Abs (a, T, t)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
fun add_c_prop qnt Ps (c, c_name, c_ty) trm =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
  let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
    fun add t = 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
      let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
        val (P, args) = strip_comb t
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
        val (P_name, P_ty) = dest_Free P
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
        val (ty_args, bool) = strip_type P_ty
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
        val args' = args
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
          |> qnt ? map (incr_boundvars 1)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
      in
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
        list_comb (Free (P_name, (c_ty :: ty_args) ---> bool), c :: args')
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
        |> qnt ? mk_induct_forall (c_name, c_ty)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
      end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
  in
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
    map_term (member (op =) Ps o head_of) add trm
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
  end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
fun prep_prem Ps c_name c_ty (avoid, avoid_trm) (params, prems, concl) =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
  let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
    val prems' = prems
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
      |> map (incr_boundvars 1) 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
      |> map (add_c_prop true Ps (Bound 0, c_name, c_ty))
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
    val avoid_trm' = avoid_trm
2994
4ee772b12032 Update to new Isabelle
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2987
diff changeset
   107
      |> fold_rev absfree (params @ [(c_name, c_ty)])
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
      |> strip_abs_body
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
      |> (fn t => mk_fresh_star_ty c_ty t (Bound 0))
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
      |> HOLogic.mk_Trueprop
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
    val prems'' = 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
      if null avoid 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
      then prems' 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
      else avoid_trm' :: prems'
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
    val concl' = concl
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
      |> incr_boundvars 1 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
      |> add_c_prop false Ps (Bound 0, c_name, c_ty)  
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
  in
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
    mk_full_horn (params @ [(c_name, c_ty)]) prems'' concl'
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
  end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
2680
cd5614027c53 removed diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2645
diff changeset
   124
(* fixme: move to nominal_library *)
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
fun same_name (Free (a1, _), Free (a2, _)) = (a1 = a2)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
  | same_name (Var (a1, _), Var (a2, _)) = (a1 = a2)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
  | same_name (Const (a1, _), Const (a2, _)) = (a1 = a2)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
  | same_name _ = false
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
2680
cd5614027c53 removed diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2645
diff changeset
   130
(* fixme: move to nominal_library *)
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
fun map7 _ [] [] [] [] [] [] [] = []
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
  | map7 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (r :: rs) (s :: ss) = 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
      f x y z u v r s :: map7 f xs ys zs us vs rs ss
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
(* local abbreviations *)
2765
7ac5e5c86c7d introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
parents: 2680
diff changeset
   136
7ac5e5c86c7d introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
parents: 2680
diff changeset
   137
local
7ac5e5c86c7d introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
parents: 2680
diff changeset
   138
  open Nominal_Permeq
7ac5e5c86c7d introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
parents: 2680
diff changeset
   139
in
7ac5e5c86c7d introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
parents: 2680
diff changeset
   140
(* by default eqvt_strict_config contains unwanted @{thm permute_pure} *) 
7ac5e5c86c7d introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
parents: 2680
diff changeset
   141
2768
639979b7fa6e added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
parents: 2765
diff changeset
   142
val eqvt_sconfig = eqvt_strict_config addpres @{thms permute_minus_cancel}
2765
7ac5e5c86c7d introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
parents: 2680
diff changeset
   143
7ac5e5c86c7d introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
parents: 2680
diff changeset
   144
fun eqvt_stac ctxt = eqvt_tac ctxt eqvt_sconfig
2768
639979b7fa6e added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
parents: 2765
diff changeset
   145
fun eqvt_srule ctxt = eqvt_rule ctxt eqvt_sconfig
2765
7ac5e5c86c7d introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
parents: 2680
diff changeset
   146
7ac5e5c86c7d introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
parents: 2680
diff changeset
   147
end
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
val all_elims = 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
  let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
     fun spec' ct = Drule.instantiate' [SOME (ctyp_of_term ct)] [NONE, SOME ct] @{thm spec}
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
  in
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
    fold (fn ct => fn th => th RS spec' ct)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
  end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
fun helper_tac flag prm p ctxt =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
  Subgoal.SUBPROOF (fn {context, prems, ...} =>
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
    let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
      val prems' = prems
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
        |> map (minus_permute_elim p)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
        |> map (eqvt_srule context)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
      val prm' = (prems' MRS prm)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
        |> flag ? (all_elims [p])
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
        |> flag ? (eqvt_srule context)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
    in
2680
cd5614027c53 removed diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2645
diff changeset
   167
      asm_full_simp_tac (HOL_ss addsimps (prm' :: @{thms induct_forall_def})) 1
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
    end) ctxt
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
fun non_binder_tac prem intr_cvars Ps ctxt = 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
  Subgoal.SUBPROOF (fn {context, params, prems, ...} =>
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
    let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
      val thy = ProofContext.theory_of context
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
      val (prms, p, _) = split_last2 (map snd params)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
      val prm_tys = map (fastype_of o term_of) prms
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
      val cperms = map (cterm_of thy o perm_const) prm_tys
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
      val p_prms = map2 (fn ct1 => fn ct2 => Thm.mk_binop ct1 p ct2) cperms prms 
2768
639979b7fa6e added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
parents: 2765
diff changeset
   178
      val prem' = prem
639979b7fa6e added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
parents: 2765
diff changeset
   179
        |> cterm_instantiate (intr_cvars ~~ p_prms) 
639979b7fa6e added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
parents: 2765
diff changeset
   180
        |> eqvt_srule ctxt
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
      (* for inductive-premises*)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
      fun tac1 prm = helper_tac true prm p context 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
      (* for non-inductive premises *)   
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
      fun tac2 prm =  
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
        EVERY' [ minus_permute_intro_tac p, 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
                 eqvt_stac context, 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
                 helper_tac false prm p context ]
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
      fun select prm (t, i) =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
        (if member same_name Ps (real_head_of t) then tac1 prm else tac2 prm) i
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
    in
2768
639979b7fa6e added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
parents: 2765
diff changeset
   194
      EVERY1 [ eqvt_stac context, 
639979b7fa6e added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
parents: 2765
diff changeset
   195
               rtac prem', 
639979b7fa6e added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
parents: 2765
diff changeset
   196
               RANGE (map (SUBGOAL o select) prems) ]
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
    end) ctxt
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
fun fresh_thm ctxt user_thm p c concl_args avoid_trm =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
  let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
    val conj1 = 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
      mk_fresh_star (mk_perm (Bound 0) (mk_perm p avoid_trm)) c
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
    val conj2 =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
      mk_fresh_star_ty @{typ perm} (mk_supp (HOLogic.mk_tuple (map (mk_perm p) concl_args))) (Bound 0)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
    val fresh_goal = mk_exists ("q", @{typ perm}) (HOLogic.mk_conj (conj1, conj2))
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
      |> HOLogic.mk_Trueprop
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
    val ss = @{thms finite_supp supp_Pair finite_Un permute_finite} @ 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
             @{thms fresh_star_Pair fresh_star_permute_iff}
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
    val simp = asm_full_simp_tac (HOL_ss addsimps ss)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
  in 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
    Goal.prove ctxt [] [] fresh_goal
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
      (K (HEADGOAL (rtac @{thm at_set_avoiding2} 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
          THEN_ALL_NEW EVERY' [cut_facts_tac user_thm, REPEAT o etac @{thm conjE}, simp])))
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
  end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
val supp_perm_eq' = @{lemma "fresh_star (supp (permute p x)) q ==> permute p x == permute (q + p) x" 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
  by (simp add: supp_perm_eq)}
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
val fresh_star_plus = @{lemma "fresh_star (permute q (permute p x)) c ==> fresh_star (permute (q + p) x) c" 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
  by (simp add: permute_plus)}
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
2645
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   223
fun binder_tac prem intr_cvars param_trms Ps user_thm avoid_trm concl_args ctxt = 
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
  Subgoal.FOCUS (fn {context = ctxt, params, prems, concl, ...} =>
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
    let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
      val thy = ProofContext.theory_of ctxt
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
      val (prms, p, c) = split_last2 (map snd params)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
      val prm_trms = map term_of prms
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
      val prm_tys = map fastype_of prm_trms
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
      val avoid_trm' = subst_free (param_trms ~~ prm_trms) avoid_trm 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
      val concl_args' = map (subst_free (param_trms ~~ prm_trms)) concl_args 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
      
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
      val user_thm' = map (cterm_instantiate (intr_cvars ~~ prms)) user_thm
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
        |> map (full_simplify (HOL_ss addsimps (@{thm fresh_star_Pair}::prems)))
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
      
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
      val fthm = fresh_thm ctxt user_thm' (term_of p) (term_of c) concl_args' avoid_trm'
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
      val (([(_, q)], fprop :: fresh_eqs), ctxt') = Obtain.result
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
              (K (EVERY1 [etac @{thm exE}, 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
                          full_simp_tac (HOL_basic_ss addsimps @{thms supp_Pair fresh_star_Un}), 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
                          REPEAT o etac @{thm conjE},
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
                          dtac fresh_star_plus,
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
                          REPEAT o dtac supp_perm_eq'])) [fthm] ctxt 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
      val expand_conv = Conv.try_conv (Conv.rewrs_conv fresh_eqs)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
      fun expand_conv_bot ctxt = Conv.bottom_conv (K expand_conv) ctxt
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
      val cperms = map (cterm_of thy o perm_const) prm_tys
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
      val qp_prms = map2 (fn ct1 => fn ct2 => Thm.mk_binop ct1 (mk_cplus q p) ct2) cperms prms 
2768
639979b7fa6e added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
parents: 2765
diff changeset
   251
      val prem' = prem
639979b7fa6e added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
parents: 2765
diff changeset
   252
        |> cterm_instantiate (intr_cvars ~~ qp_prms)
639979b7fa6e added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
parents: 2765
diff changeset
   253
        |> eqvt_srule ctxt
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
      val fprop' = eqvt_srule ctxt' fprop 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
      val tac_fresh = simp_tac (HOL_basic_ss addsimps [fprop'])
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
      (* for inductive-premises*)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
      fun tac1 prm = helper_tac true prm (mk_cplus q p) ctxt' 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   260
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   261
      (* for non-inductive premises *)   
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   262
      fun tac2 prm =  
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
        EVERY' [ minus_permute_intro_tac (mk_cplus q p), 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
                 eqvt_stac ctxt, 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
                 helper_tac false prm (mk_cplus q p) ctxt' ]
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
      fun select prm (t, i) =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
        (if member same_name Ps (real_head_of t) then tac1 prm else tac2 prm) i
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   269
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   270
      val side_thm = Goal.prove ctxt' [] [] (term_of concl)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   271
        (fn {context, ...} => 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
           EVERY1 [ CONVERSION (expand_conv_bot context),
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
                    eqvt_stac context,
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
                    rtac prem',
2680
cd5614027c53 removed diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2645
diff changeset
   275
                    RANGE (tac_fresh :: map (SUBGOAL o select) prems) ])
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   276
        |> singleton (ProofContext.export ctxt' ctxt)        
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   277
    in
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   278
      rtac side_thm 1
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
    end) ctxt
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   280
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   281
fun case_tac ctxt Ps avoid avoid_trm intr_cvars param_trms prem user_thm concl_args =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
  let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   283
    val tac1 = non_binder_tac prem intr_cvars Ps ctxt
2645
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   284
    val tac2 = binder_tac prem intr_cvars param_trms Ps user_thm avoid_trm concl_args ctxt
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   285
  in 
2768
639979b7fa6e added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
parents: 2765
diff changeset
   286
    EVERY' [ rtac @{thm allI}, rtac @{thm allI}, 
639979b7fa6e added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
parents: 2765
diff changeset
   287
             if null avoid then tac1 else tac2 ]
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   288
  end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   289
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   290
fun prove_sinduct_tac raw_induct user_thms Ps avoids avoid_trms intr_cvars param_trms concl_args 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   291
  {prems, context} =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   292
  let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   293
    val cases_tac = 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   294
      map7 (case_tac context Ps) avoids avoid_trms intr_cvars param_trms prems user_thms concl_args
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   295
  in 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   296
    EVERY1 [ DETERM o rtac raw_induct, RANGE cases_tac ]
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   297
  end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   298
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   299
val normalise = @{lemma "(Q --> (!p c. P p c)) ==> (!!c. Q ==> P (0::perm) c)" by simp}
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   300
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   301
fun prove_strong_inductive pred_names rule_names avoids raw_induct intrs ctxt =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   302
  let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   303
    val thy = ProofContext.theory_of ctxt
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   304
    val ((_, [raw_induct']), ctxt') = Variable.import true [raw_induct] ctxt
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   305
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   306
    val (ind_prems, ind_concl) = raw_induct'
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   307
      |> prop_of
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   308
      |> Logic.strip_horn
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   309
      |>> map strip_full_horn
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   310
    val params = map (fn (x, _, _) => x) ind_prems
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   311
    val param_trms = (map o map) Free params  
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   312
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   313
    val intr_vars_tys = map (fn t => rev (Term.add_vars (prop_of t) [])) intrs
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   314
    val intr_vars = (map o map) fst intr_vars_tys
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   315
    val intr_vars_substs = map2 (curry (op ~~)) intr_vars param_trms
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   316
    val intr_cvars = (map o map) (cterm_of thy o Var) intr_vars_tys      
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   317
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   318
    val (intr_prems, intr_concls) = intrs
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   319
      |> map prop_of
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   320
      |> map2 subst_Vars intr_vars_substs
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   321
      |> map Logic.strip_horn
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   322
      |> split_list
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   323
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   324
    val intr_concls_args = map (snd o strip_comb o HOLogic.dest_Trueprop) intr_concls 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   325
      
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   326
    val avoid_trms = avoids
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
      |> (map o map) (setify ctxt') 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
      |> map fold_union
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   329
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   330
    val vc_compat_goals = 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   331
      map4 mk_vc_compat (avoids ~~ avoid_trms) intr_prems intr_concls_args params
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   332
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   333
    val ([c_name, a, p], ctxt'') = Variable.variant_fixes ["c", "'a", "p"] ctxt'
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   334
    val c_ty = TFree (a, @{sort fs})
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   335
    val c = Free (c_name, c_ty)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   336
    val p = Free (p, @{typ perm})
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   337
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   338
    val (preconds, ind_concls) = ind_concl
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   339
      |> HOLogic.dest_Trueprop
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   340
      |> HOLogic.dest_conj 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   341
      |> map HOLogic.dest_imp
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   342
      |> split_list
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   343
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   344
    val Ps = map (fst o strip_comb) ind_concls
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   345
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   346
    val ind_concl' = ind_concls
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   347
      |> map (add_p_c p (c, c_ty))
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   348
      |> (curry (op ~~)) preconds  
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   349
      |> map HOLogic.mk_imp
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   350
      |> fold_conj
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   351
      |> HOLogic.mk_Trueprop
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   352
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   353
    val ind_prems' = ind_prems
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   354
      |> map2 (prep_prem Ps c_name c_ty) (avoids ~~ avoid_trms)   
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   355
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   356
    fun after_qed ctxt_outside user_thms ctxt = 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   357
      let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   358
        val strong_ind_thms = Goal.prove ctxt [] ind_prems' ind_concl' 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   359
        (prove_sinduct_tac raw_induct user_thms Ps avoids avoid_trms intr_cvars param_trms intr_concls_args) 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   360
          |> singleton (ProofContext.export ctxt ctxt_outside)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   361
          |> Datatype_Aux.split_conj_thm
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   362
          |> map (fn thm => thm RS normalise)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   363
          |> map (asm_full_simplify (HOL_basic_ss addsimps @{thms permute_zero induct_rulify})) 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   364
          |> map (Drule.rotate_prems (length ind_prems'))
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   365
          |> map zero_var_indexes
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   366
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   367
        val qualified_thm_name = pred_names
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   368
          |> map Long_Name.base_name
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   369
          |> space_implode "_"
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   370
          |> (fn s => Binding.qualify false s (Binding.name "strong_induct"))
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   371
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   372
        val attrs = 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   373
          [ Attrib.internal (K (Rule_Cases.consumes 1)),
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   374
            Attrib.internal (K (Rule_Cases.case_names rule_names)) ]
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   375
      in
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   376
        ctxt
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   377
        |> Local_Theory.note ((qualified_thm_name, attrs), strong_ind_thms)    
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   378
        |> snd   
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   379
      end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   380
  in
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   381
    Proof.theorem NONE (after_qed ctxt) ((map o map) (rpair []) vc_compat_goals) ctxt''
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   382
  end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   383
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   384
fun prove_strong_inductive_cmd (pred_name, avoids) ctxt =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   385
  let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   386
    val thy = ProofContext.theory_of ctxt;
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   387
    val ({names, ...}, {raw_induct, intrs, ...}) =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   388
      Inductive.the_inductive ctxt (Sign.intern_const thy pred_name);
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   389
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   390
    val rule_names = 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   391
      hd names
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   392
      |> the o Induct.lookup_inductP ctxt
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   393
      |> fst o Rule_Cases.get
2987
27aab7a105eb updated for new Isabelle (11. Aug.)
Christian Urban <urbanc@in.tum.de>
parents: 2768
diff changeset
   394
      |> map (fst o fst)
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   395
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   396
    val _ = (case duplicates (op = o pairself fst) avoids of
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   397
        [] => ()
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   398
      | xs => error ("Duplicate case names: " ^ commas_quote (map fst xs)))
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   399
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   400
    val _ = (case subtract (op =) rule_names (map fst avoids) of
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   401
        [] => ()
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   402
      | xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs))
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   403
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   404
    val avoids_ordered = order_default (op =) [] rule_names avoids
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   405
      
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   406
    fun read_avoids avoid_trms intr =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   407
      let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   408
        (* fixme hack *)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   409
        val (((_, ctrms), _), ctxt') = Variable.import true [intr] ctxt
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   410
        val trms = map (term_of o snd) ctrms
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   411
        val ctxt'' = fold Variable.declare_term trms ctxt' 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   412
      in
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   413
        map (Syntax.read_term ctxt'') avoid_trms 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   414
      end 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   415
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   416
    val avoid_trms = map2 read_avoids avoids_ordered intrs
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   417
  in
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   418
    prove_strong_inductive names rule_names avoid_trms raw_induct intrs ctxt
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   419
  end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   420
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   421
(* outer syntax *)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   422
local
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   423
  structure P = Parse;
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   424
  structure S = Scan
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   425
  
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   426
  val _ = Keyword.keyword "avoids"
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   427
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   428
  val single_avoid_parser = 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   429
    P.name -- (P.$$$ ":" |-- P.and_list1 P.term)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   430
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   431
  val avoids_parser = 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   432
    S.optional (P.$$$ "avoids" |-- P.enum1 "|" single_avoid_parser) []
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   433
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   434
  val main_parser = P.xname -- avoids_parser
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   435
in
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   436
  val _ =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   437
  Outer_Syntax.local_theory_to_proof "nominal_inductive"
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   438
    "prove strong induction theorem for inductive predicate involving nominal datatypes"
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   439
      Keyword.thy_goal (main_parser >> prove_strong_inductive_cmd)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   440
end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   441
2994
4ee772b12032 Update to new Isabelle
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2987
diff changeset
   442
end