Nominal/nominal_inductive.ML
author Christian Urban <urbanc@in.tum.de>
Thu, 19 Apr 2018 13:58:22 +0100
branchNominal2-Isabelle2016-1
changeset 3246 66114fa3d2ee
parent 3245 017e33849f4d
permissions -rw-r--r--
updated to Isabelle 2016-1
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
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
     3
    Author:     Tjark Weber
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
    Infrastructure for proving strong induction theorems
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
    for inductive predicates involving nominal datatypes.
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
    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
     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
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    13
  val prove_strong_inductive: string list -> string list -> term list list -> thm -> thm list ->
2639
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
  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
    16
end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
structure Nominal_Inductive : NOMINAL_INDUCTIVE =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
struct
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    21
fun mk_cplus p q =
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    22
  Thm.apply (Thm.apply @{cterm "plus :: perm => perm => perm"} p) q
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    24
fun mk_cminus p =
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    25
  Thm.apply @{cterm "uminus :: perm => perm"} p
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
    27
fun minus_permute_intro_tac ctxt p =
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
    28
  resolve_tac ctxt [Thm.instantiate' [] [SOME (mk_cminus p)] @{thm permute_boolE}]
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    30
fun minus_permute_elim p thm =
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
    31
  thm RS (Thm.instantiate' [] [NONE, SOME (mk_cminus p)] @{thm permute_boolI})
2639
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
3231
188826f1ccdb updated to massive changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3218
diff changeset
    35
  | real_head_of (Const (@{const_name Pure.imp}, _) $ _ $ t) = real_head_of t
188826f1ccdb updated to massive changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3218
diff changeset
    36
  | real_head_of (Const (@{const_name Pure.all}, _) $ Abs (_, _, t)) = real_head_of t
2639
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
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
    38
  | real_head_of (Const (@{const_name HOL.induct_forall}, _) $ Abs (_, _, t)) = real_head_of t
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    39
  | real_head_of t = head_of t
2639
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
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    42
fun mk_vc_compat (avoid, avoid_trm) prems concl_args params =
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    43
  if null avoid then
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    44
    []
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    45
  else
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    46
    let
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    47
      val vc_goal = concl_args
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    48
        |> HOLogic.mk_tuple
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    49
        |> mk_fresh_star avoid_trm
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    50
        |> HOLogic.mk_Trueprop
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    51
        |> (curry Logic.list_implies) prems
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    52
        |> fold_rev (Logic.all o Free) params
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    53
      val finite_goal = avoid_trm
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    54
        |> mk_finite
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    55
        |> HOLogic.mk_Trueprop
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    56
        |> (curry Logic.list_implies) prems
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    57
        |> fold_rev (Logic.all o Free) params
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    58
    in
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    59
      [vc_goal, finite_goal]
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    60
    end
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
2680
cd5614027c53 removed diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2645
diff changeset
    62
(* fixme: move to nominal_library *)
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
fun map_term prop f trm =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
  if prop trm 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
  then f trm
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
  else case trm of
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
    (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
    68
  | 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
    69
  | _ => trm
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
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
    72
  let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
    val (P, args) = strip_comb trm
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
    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
    75
    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
    76
    val args' = map (mk_perm p) args
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
  in
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
    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
    79
    |> (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
    80
    |> (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
    81
  end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    83
fun induct_forall_const T =
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
    84
  Const (@{const_name HOL.induct_forall}, (T --> @{typ bool}) --> @{typ bool})
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    85
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    86
fun mk_induct_forall (a, T) t =
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    87
  induct_forall_const T $ Abs (a, T, t)
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
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
    90
  let
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
    91
    fun add t =
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
      let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
        val (P, args) = strip_comb t
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
        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
    95
        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
    96
        val args' = args
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
          |> qnt ? map (incr_boundvars 1)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
      in
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
        list_comb (Free (P_name, (c_ty :: ty_args) ---> bool), c :: args')
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   100
          |> qnt ? mk_induct_forall (c_name, c_ty)
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
      end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
  in
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
    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
   104
  end
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
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
   107
  let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
    val prems' = prems
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   109
      |> map (incr_boundvars 1)
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
      |> 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
   111
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
    val avoid_trm' = avoid_trm
2994
4ee772b12032 Update to new Isabelle
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2987
diff changeset
   113
      |> 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
   114
      |> strip_abs_body
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
      |> (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
   116
      |> HOLogic.mk_Trueprop
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   118
    val prems'' =
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   119
      if null avoid
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   120
      then prems'
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
      else avoid_trm' :: prems'
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
    val concl' = concl
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   124
      |> incr_boundvars 1
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   125
      |> add_c_prop false Ps (Bound 0, c_name, c_ty)
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
  in
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
    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
   128
  end
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 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
   132
  | 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
   133
  | 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
   134
  | same_name _ = false
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
2680
cd5614027c53 removed diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2645
diff changeset
   136
(* fixme: move to nominal_library *)
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
fun map7 _ [] [] [] [] [] [] [] = []
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   138
  | map7 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (r :: rs) (s :: ss) =
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
      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
   140
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
(* 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
   142
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
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
   144
  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
   145
in
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   146
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   147
  (* by default eqvt_strict_config contains unwanted @{thm permute_pure} *) 
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
   148
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   149
  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
   150
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   151
  fun eqvt_stac ctxt = eqvt_tac ctxt eqvt_sconfig
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   152
  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
   153
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
   154
end
2639
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
val all_elims = 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
  let
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   158
    fun spec' ct =
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   159
      Thm.instantiate' [SOME (Thm.ctyp_of_cterm ct)] [NONE, SOME ct] @{thm spec}
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
  in
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
    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
   162
  end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
fun helper_tac flag prm p ctxt =
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   165
  Subgoal.SUBPROOF (fn {context = ctxt', prems, ...} =>
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
    let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
      val prems' = prems
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
        |> map (minus_permute_elim p)
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   169
        |> map (eqvt_srule ctxt')
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
      val prm' = (prems' MRS prm)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
        |> flag ? (all_elims [p])
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   173
        |> flag ? (eqvt_srule ctxt')
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
    in
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   175
      asm_full_simp_tac (put_simpset HOL_ss ctxt' addsimps (prm' :: @{thms HOL.induct_forall_def})) 1
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
    end) ctxt
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
fun non_binder_tac prem intr_cvars Ps ctxt = 
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
   179
  Subgoal.SUBPROOF (fn {context = ctxt', params, prems, ...} =>
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
    let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
      val (prms, p, _) = split_last2 (map snd params)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
   182
      val prm_tys = map (fastype_of o Thm.term_of) prms
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
   183
      val cperms = map (Thm.cterm_of ctxt' o perm_const) prm_tys
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
      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
   185
      val prem' = prem
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   186
        |> infer_instantiate ctxt' (map (#1 o dest_Var o Thm.term_of) intr_cvars ~~ p_prms) 
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   187
        |> eqvt_srule ctxt'
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
      (* for inductive-premises*)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
   190
      fun tac1 prm = helper_tac true prm p ctxt'
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
      (* for non-inductive premises *)   
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
      fun tac2 prm =  
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   194
        EVERY' [ minus_permute_intro_tac ctxt' p, 
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
   195
                 eqvt_stac ctxt', 
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
   196
                 helper_tac false prm p ctxt' ]
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
      fun select prm (t, i) =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
        (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
   200
    in
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
   201
      EVERY1 [ eqvt_stac ctxt', 
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   202
               resolve_tac ctxt' [prem'], 
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
   203
               RANGE (map (SUBGOAL o select) prems) ]
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
    end) ctxt
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
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
   207
  let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
    val conj1 = 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
      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
   210
    val conj2 =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
      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
   212
    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
   213
      |> HOLogic.mk_Trueprop
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
    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
   216
             @{thms fresh_star_Pair fresh_star_permute_iff}
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3214
diff changeset
   217
    val simp = asm_full_simp_tac (put_simpset HOL_ss ctxt addsimps ss)
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
  in 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
    Goal.prove ctxt [] [] fresh_goal
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   220
      (K (HEADGOAL (resolve_tac ctxt @{thms at_set_avoiding2}
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   221
          THEN_ALL_NEW EVERY' [cut_facts_tac user_thm, REPEAT o
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   222
            eresolve_tac ctxt @{thms conjE}, simp])))
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
  end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
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
   226
  by (simp add: supp_perm_eq)}
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   227
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
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
   229
  by (simp add: permute_plus)}
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
2645
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   232
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
   233
  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
   234
    let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
      val (prms, p, c) = split_last2 (map snd params)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
   236
      val prm_trms = map Thm.term_of prms
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
      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
   238
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
      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
   240
      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
   241
      
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   242
      val user_thm' = map (infer_instantiate ctxt (map (#1 o dest_Var o Thm.term_of) intr_cvars ~~ prms)) user_thm
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3214
diff changeset
   243
        |> map (full_simplify (put_simpset HOL_ss ctxt addsimps (@{thm fresh_star_Pair}::prems)))
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
      
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
   245
      val fthm = fresh_thm ctxt user_thm' (Thm.term_of p) (Thm.term_of c) concl_args' avoid_trm'
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
      val (([(_, q)], fprop :: fresh_eqs), ctxt') = Obtain.result
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   248
              (K (EVERY1 [eresolve_tac ctxt @{thms exE}, 
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3214
diff changeset
   249
                          full_simp_tac (put_simpset HOL_basic_ss ctxt
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3214
diff changeset
   250
                            addsimps @{thms supp_Pair fresh_star_Un}),
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   251
                          REPEAT o eresolve_tac ctxt @{thms conjE},
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   252
                          dresolve_tac ctxt [fresh_star_plus],
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   253
                          REPEAT o dresolve_tac ctxt [supp_perm_eq']])) [fthm] 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 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
   256
      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
   257
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   258
      val cperms = map (Thm.cterm_of ctxt' o perm_const) prm_tys
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
      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
   260
      val prem' = prem
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   261
        |> infer_instantiate ctxt' (map (#1 o dest_Var o Thm.term_of) intr_cvars ~~ qp_prms)
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   262
        |> eqvt_srule ctxt'
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
      val fprop' = eqvt_srule ctxt' fprop 
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3214
diff changeset
   265
      val tac_fresh = simp_tac (put_simpset HOL_basic_ss ctxt' addsimps [fprop'])
2639
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
      (* for inductive-premises*)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
      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
   269
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   270
      (* for non-inductive premises *)   
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   271
      fun tac2 prm =  
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   272
        EVERY' [ minus_permute_intro_tac ctxt' (mk_cplus q p), 
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   273
                 eqvt_stac ctxt', 
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
                 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
   275
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   276
      fun select prm (t, i) =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   277
        (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
   278
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
   279
      val side_thm = Goal.prove ctxt' [] [] (Thm.term_of concl)
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
   280
        (fn {context = ctxt'', ...} => 
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
   281
           EVERY1 [ CONVERSION (expand_conv_bot ctxt''),
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
   282
                    eqvt_stac ctxt'',
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   283
                    resolve_tac ctxt'' [prem'],
2680
cd5614027c53 removed diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2645
diff changeset
   284
                    RANGE (tac_fresh :: map (SUBGOAL o select) prems) ])
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
   285
        |> singleton (Proof_Context.export ctxt' ctxt)
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   286
    in
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   287
      resolve_tac ctxt [side_thm] 1
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   288
    end) ctxt
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 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
   291
  let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   292
    val tac1 = non_binder_tac prem intr_cvars Ps ctxt
2645
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   293
    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
   294
  in 
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   295
    EVERY' [ resolve_tac ctxt @{thms allI}, resolve_tac ctxt @{thms allI}, 
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
   296
             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
   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
fun prove_sinduct_tac raw_induct user_thms Ps avoids avoid_trms intr_cvars param_trms concl_args 
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   300
  {prems, context = ctxt} =
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   301
  let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   302
    val cases_tac = 
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   303
      map7 (case_tac ctxt Ps) avoids avoid_trms intr_cvars param_trms prems user_thms concl_args
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   304
  in 
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   305
    EVERY1 [ DETERM o resolve_tac ctxt [raw_induct], RANGE cases_tac ]
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   306
  end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   307
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   308
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
   309
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   310
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
   311
  let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   312
    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
   313
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   314
    val (ind_prems, ind_concl) = raw_induct'
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
   315
      |> Thm.prop_of
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   316
      |> Logic.strip_horn
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   317
      |>> map strip_full_horn
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   318
    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
   319
    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
   320
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
   321
    val intr_vars_tys = map (fn t => rev (Term.add_vars (Thm.prop_of t) [])) intrs
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   322
    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
   323
    val intr_vars_substs = map2 (curry (op ~~)) intr_vars param_trms
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
   324
    val intr_cvars = (map o map) (Thm.cterm_of ctxt o Var) intr_vars_tys
2639
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 (intr_prems, intr_concls) = intrs
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
   327
      |> map Thm.prop_of
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
      |> map2 subst_Vars intr_vars_substs
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   329
      |> map Logic.strip_horn
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   330
      |> split_list
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   331
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   332
    val intr_concls_args =
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   333
      map (snd o fixed_nonfixed_args ctxt' o HOLogic.dest_Trueprop) intr_concls
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   334
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   335
    val avoid_trms = avoids
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   336
      |> (map o map) (setify ctxt') 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   337
      |> map fold_union
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   338
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   339
    val vc_compat_goals = 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   340
      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
   341
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   342
    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
   343
    val c_ty = TFree (a, @{sort fs})
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   344
    val c = Free (c_name, c_ty)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   345
    val p = Free (p, @{typ perm})
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   346
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   347
    val (preconds, ind_concls) = ind_concl
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   348
      |> HOLogic.dest_Trueprop
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   349
      |> HOLogic.dest_conj 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   350
      |> map HOLogic.dest_imp
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   351
      |> split_list
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 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
   354
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   355
    val ind_concl' = ind_concls
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   356
      |> 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
   357
      |> (curry (op ~~)) preconds  
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   358
      |> map HOLogic.mk_imp
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   359
      |> fold_conj
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   360
      |> HOLogic.mk_Trueprop
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   361
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   362
    val ind_prems' = ind_prems
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   363
      |> 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
   364
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   365
    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
   366
      let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   367
        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
   368
        (prove_sinduct_tac raw_induct user_thms Ps avoids avoid_trms intr_cvars param_trms intr_concls_args) 
3045
d0ad264f8c4f updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
parents: 2994
diff changeset
   369
          |> singleton (Proof_Context.export ctxt ctxt_outside)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
   370
          |> Old_Datatype_Aux.split_conj_thm
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   371
          |> map (fn thm => thm RS normalise)
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3214
diff changeset
   372
          |> map (asm_full_simplify (put_simpset HOL_basic_ss ctxt
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3214
diff changeset
   373
              addsimps @{thms permute_zero induct_rulify})) 
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   374
          |> map (Drule.rotate_prems (length ind_prems'))
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   375
          |> map zero_var_indexes
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   376
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   377
        val qualified_thm_name = pred_names
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   378
          |> map Long_Name.base_name
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   379
          |> space_implode "_"
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   380
          |> (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
   381
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   382
        val attrs = 
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   383
          [ Attrib.internal (K (Rule_Cases.consumes 1)),
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   384
            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
   385
      in
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   386
        ctxt
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   387
        |> 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
   388
        |> snd   
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   389
      end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   390
  in
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   391
    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
   392
  end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   393
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   394
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
   395
  let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   396
    val ({names, ...}, {raw_induct, intrs, ...}) =
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   397
      Inductive.the_inductive ctxt (long_name ctxt pred_name)
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   398
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   399
    val rule_names = hd names
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   400
      |> the o Induct.lookup_inductP ctxt
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   401
      |> fst o Rule_Cases.get
2987
27aab7a105eb updated for new Isabelle (11. Aug.)
Christian Urban <urbanc@in.tum.de>
parents: 2768
diff changeset
   402
      |> map (fst o fst)
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   403
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   404
    val case_names = map fst avoids
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   405
    val _ = case duplicates (op =) case_names of
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   406
        [] => ()
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   407
      | xs => error ("Duplicate case names: " ^ commas_quote xs)
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   408
    val _ = case subtract (op =) rule_names case_names of
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   409
        [] => ()
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   410
      | xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs)
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   411
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   412
    val avoids_ordered = order_default (op =) [] rule_names avoids
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   413
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   414
    fun read_avoids avoid_trms intr =
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   415
      let
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   416
        (* fixme hack *)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   417
        val (((_, ctrms), _), ctxt') = Variable.import true [intr] ctxt
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
   418
        val trms = map (Thm.term_of o snd) ctrms
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   419
        val ctxt'' = fold Variable.declare_term trms ctxt'
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   420
      in
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   421
        map (Syntax.read_term ctxt'') avoid_trms
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   422
      end
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   423
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   424
    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
   425
  in
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   426
    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
   427
  end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   428
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   429
(* outer syntax *)
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   430
local
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   431
  val single_avoid_parser =
3135
92b9b8d2888d updated to new Isabelle (20 March)
Christian Urban <urbanc@in.tum.de>
parents: 3123
diff changeset
   432
    Parse.name -- (@{keyword ":"} |-- Parse.and_list1 Parse.term)
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   433
3214
13ab4f0a0b0e Various changes to support Nominal2 commands in local contexts.
webertj
parents: 3190
diff changeset
   434
  val avoids_parser =
3135
92b9b8d2888d updated to new Isabelle (20 March)
Christian Urban <urbanc@in.tum.de>
parents: 3123
diff changeset
   435
    Scan.optional (@{keyword "avoids"} |-- Parse.enum1 "|" single_avoid_parser) []
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   436
3245
017e33849f4d updated to Isabelle 2016-1
Christian Urban <urbanc@in.tum.de>
parents: 3244
diff changeset
   437
  val main_parser = Parse.name -- avoids_parser
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   438
in
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   439
  val _ =
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3231
diff changeset
   440
    Outer_Syntax.local_theory_to_proof @{command_keyword nominal_inductive}
3135
92b9b8d2888d updated to new Isabelle (20 March)
Christian Urban <urbanc@in.tum.de>
parents: 3123
diff changeset
   441
      "prove strong induction theorem for inductive predicate involving nominal datatypes"
92b9b8d2888d updated to new Isabelle (20 March)
Christian Urban <urbanc@in.tum.de>
parents: 3123
diff changeset
   442
        (main_parser >> prove_strong_inductive_cmd)
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   443
end
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   444
2994
4ee772b12032 Update to new Isabelle
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2987
diff changeset
   445
end