Nominal/Ex/Typing.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 04 Jan 2011 13:47:38 +0000
changeset 2637 3890483c674f
parent 2636 0865caafbfe6
child 2638 e1e2ca92760b
permissions -rw-r--r--
final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Lambda
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports "../Nominal2" 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
atom_decl name
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
nominal_datatype lam =
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
  Var "name"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
| App "lam" "lam"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
| Lam x::"name" l::"lam"  bind x in l
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
thm lam.distinct
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
thm lam.induct
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
thm lam.exhaust lam.strong_exhaust
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
thm lam.fv_defs
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
thm lam.bn_defs
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
thm lam.perm_simps
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
thm lam.eq_iff
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
thm lam.fv_bn_eqvt
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
thm lam.size_eqvt
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
fun mk_cminus p = Thm.capply @{cterm "uminus::perm \<Rightarrow> perm"} p 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
fun minus_permute_intro_tac p = 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
  rtac (Drule.instantiate' [] [SOME (mk_cminus p)] @{thm permute_boolE})
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
fun minus_permute_elim p thm = 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
  thm RS (Drule.instantiate' [] [NONE, SOME (mk_cminus p)] @{thm permute_boolI})
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
fun real_head_of (@{term Trueprop} $ t) = real_head_of t
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
  | real_head_of (Const ("==>", _) $ _ $ t) = real_head_of t
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
  | real_head_of (Const (@{const_name all}, _) $ Abs (_, _, t)) = real_head_of t
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
  | real_head_of (Const (@{const_name All}, _) $ Abs (_, _, t)) = real_head_of t
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
  | real_head_of (Const ("HOL.induct_forall", _) $ Abs (_, _, t)) = real_head_of t
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
  | real_head_of t = head_of t  
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
ML {* 
2637
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
    43
fun mk_vc_compat (avoid, avoid_trm) prems concl_args params = 
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
  let
2637
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
    45
    val vc_goal = concl_args
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
    46
      |> HOLogic.mk_tuple
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
      |> mk_fresh_star avoid_trm 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
      |> HOLogic.mk_Trueprop
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
      |> (curry Logic.list_implies) prems
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
      |> (curry list_all_free) params
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
  in 
2637
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
    52
    if null avoid then [] else [vc_goal]
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
  end
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
fun map_term prop f trm =
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
  if prop trm 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
  then f trm
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
  else case trm of
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
    (t1 $ t2) => map_term prop f t1 $ map_term prop f t2
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
  | Abs (x, T, t) => Abs (x, T, map_term prop f t)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
  | _ => trm
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
fun add_p_c p (c, c_ty) trm =
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
  let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
    val (P, args) = strip_comb trm
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
    val (P_name, P_ty) = dest_Free P
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
    val (ty_args, bool) = strip_type P_ty
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
    val args' = map (mk_perm p) args
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
  in
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
    list_comb (Free (P_name, (c_ty :: ty_args) ---> bool),  c :: args')
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
    |> (fn t => HOLogic.all_const c_ty $ lambda c t )
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
    |> (fn t => HOLogic.all_const @{typ perm} $  lambda p t)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
  end
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
fun induct_forall_const T = Const ("HOL.induct_forall", (T --> @{typ bool}) --> @{typ bool})
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
fun mk_induct_forall (a, T) t =  induct_forall_const T $ Abs (a, T, t)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
fun add_c_prop qnt Ps (c, c_name, c_ty) trm =
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
  let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
    fun add t = 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
      let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
        val (P, args) = strip_comb t
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
        val (P_name, P_ty) = dest_Free P
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
        val (ty_args, bool) = strip_type P_ty
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
        val args' = args
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
          |> qnt ? map (incr_boundvars 1)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
      in
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
        list_comb (Free (P_name, (c_ty :: ty_args) ---> bool), c :: args')
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
        |> qnt ? mk_induct_forall (c_name, c_ty)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
      end
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
  in
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
    map_term (member (op =) Ps o head_of) add trm
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
  end
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
fun prep_prem Ps c_name c_ty (avoid, avoid_trm) (params, prems, concl) =
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
  let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
    val prems' = prems
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
      |> map (incr_boundvars 1) 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
      |> map (add_c_prop true Ps (Bound 0, c_name, c_ty))
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
    val avoid_trm' = avoid_trm
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
      |> (curry list_abs_free) (params @ [(c_name, c_ty)])
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
      |> strip_abs_body
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
      |> (fn t => mk_fresh_star_ty c_ty t (Bound 0))
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
      |> HOLogic.mk_Trueprop
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
    val prems'' = 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
      if null avoid 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
      then prems' 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
      else avoid_trm' :: prems'
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
    val concl' = concl
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
      |> incr_boundvars 1 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
      |> add_c_prop false Ps (Bound 0, c_name, c_ty)  
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
  in
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
    mk_full_horn (params @ [(c_name, c_ty)]) prems'' concl'
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
  end
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
fun same_name (Free (a1, _), Free (a2, _)) = (a1 = a2)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
  | same_name (Var (a1, _), Var (a2, _)) = (a1 = a2)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
  | same_name (Const (a1, _), Const (a2, _)) = (a1 = a2)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
  | same_name _ = false
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
(* local abbreviations *)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
fun eqvt_stac ctxt = Nominal_Permeq.eqvt_strict_tac ctxt @{thms permute_minus_cancel} []  
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
fun eqvt_srule ctxt = Nominal_Permeq.eqvt_strict_rule ctxt @{thms permute_minus_cancel} []  
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
val all_elims = 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
  let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
     fun spec' ct = Drule.instantiate' [SOME (ctyp_of_term ct)] [NONE, SOME ct] @{thm spec}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
  in
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
    fold (fn ct => fn th => th RS spec' ct)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
  end
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
fun helper_tac flag prm p ctxt =
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
  Subgoal.SUBPROOF (fn {context, prems, ...} =>
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
    let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
      val prems' = prems
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
        |> map (minus_permute_elim p)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
        |> map (eqvt_srule context)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
     
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
      val prm' = (prems' MRS prm)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
        |> flag ? (all_elims [p])
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
        |> flag ? (simplify (HOL_basic_ss addsimps @{thms permute_minus_cancel}))
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
    in
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
      simp_tac (HOL_ss addsimps (prm' :: @{thms induct_forall_def })) 1
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
    end) ctxt
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
fun non_binder_tac prem intr_cvars Ps ctxt = 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
  Subgoal.SUBPROOF (fn {context, params, prems, ...} =>
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
    let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
      val thy = ProofContext.theory_of context
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
      val (prms, p, _) = split_last2 (map snd params)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
      val prm_tys = map (fastype_of o term_of) prms
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
      val cperms = map (cterm_of thy o perm_const) prm_tys
0865caafbfe6 file with most of the strong rule induction development
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 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
      val prem' = cterm_instantiate (intr_cvars ~~ p_prms) prem
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
      (* for inductive-premises*)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
      fun tac1 prm = helper_tac true prm p context
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
      (* for non-inductive premises *)   
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
      fun tac2 prm =  
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
        EVERY' [ minus_permute_intro_tac p, 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
                 eqvt_stac context, 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
                 helper_tac false prm p context ]
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
      fun select prm (t, i) =
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
        (if member same_name Ps (real_head_of t) then tac1 prm else tac2 prm) i
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
    in
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
      EVERY1 [rtac prem', RANGE (map (SUBGOAL o select) prems) ]
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
    end) ctxt
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
2637
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   196
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
ML {*
2637
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   198
fun fresh_thm ctxt fresh_thms p c prms avoid_trm =
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   199
  let
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   200
    val conj1 = 
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   201
      mk_fresh_star (mk_perm (Bound 0) (mk_perm p avoid_trm)) c
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   202
    val conj2 =
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   203
      mk_fresh_star_ty @{typ perm} (mk_supp (HOLogic.mk_tuple (map (mk_perm p) prms))) (Bound 0)
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   204
    val fresh_goal = mk_exists ("q", @{typ perm}) (HOLogic.mk_conj (conj1, conj2))
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   205
      |> HOLogic.mk_Trueprop
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   206
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   207
    val _ = tracing ("fresh goal: " ^ Syntax.string_of_term ctxt fresh_goal)
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   208
  in 
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   209
    Goal.prove ctxt [] [] fresh_goal
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   210
      (K (HEADGOAL (rtac @{thm at_set_avoiding2})))
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   211
  end
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   212
*}
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   213
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   214
ML {*
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   215
fun binder_tac prem intr_cvars param_trms Ps fresh_thms avoid avoid_trm ctxt = 
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
  Subgoal.FOCUS (fn {context, params, ...} =>
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
    let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
      val thy = ProofContext.theory_of context
2637
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   219
      val (prms, p, c) = split_last2 (map snd params)
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   220
      val prm_trms = map term_of prms
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   221
      val prm_tys = map fastype_of prm_trms
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
      val cperms = map (cterm_of thy o perm_const) prm_tys
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
      val p_prms = map2 (fn ct1 => fn ct2 => Thm.mk_binop ct1 p ct2) cperms prms 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
      val prem' = cterm_instantiate (intr_cvars ~~ p_prms) prem
2637
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   225
      val avoid_trm' = subst_free (param_trms ~~ prm_trms) avoid_trm      
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   226
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   227
      val fthm = fresh_thm context fresh_thms (term_of p) (term_of c) (map term_of prms) avoid_trm'
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
    in
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
      Skip_Proof.cheat_tac thy
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
      (* EVERY1 [rtac prem'] *)  
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
    end) ctxt
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
ML {*
2637
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   235
fun case_tac ctxt fresh_thms Ps (avoid, avoid_trm) intr_cvars param_trms prem =
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
  let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
    val tac1 = non_binder_tac prem intr_cvars Ps ctxt
2637
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   238
    val tac2 = binder_tac prem intr_cvars param_trms Ps fresh_thms avoid avoid_trm ctxt
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
  in 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
    EVERY' [ rtac @{thm allI}, rtac @{thm allI}, eqvt_stac ctxt,
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
             if null avoid then tac1 else tac2 ]
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
  end
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
ML {*
2637
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   246
fun prove_sinduct_tac raw_induct fresh_thms Ps avoids avoid_trms intr_cvars param_trms {prems, context} =
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
  let
2637
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   248
    val cases_tac = map4 (case_tac context fresh_thms Ps) (avoids ~~avoid_trms) intr_cvars param_trms prems
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
  in 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
    EVERY1 [ DETERM o rtac raw_induct, RANGE cases_tac ]
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
  end
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
ML {*
2637
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   255
val normalise = @{lemma "(Q --> (!p c. P p c)) ==> (!!c. Q ==> P (0::perm) c)" by simp}
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
fun prove_strong_inductive rule_names avoids raw_induct intrs ctxt =
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   260
  let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   261
    val thy = ProofContext.theory_of ctxt
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   262
    val ((_, [raw_induct']), ctxt') = Variable.import true [raw_induct] ctxt
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
    val (ind_prems, ind_concl) = raw_induct'
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
      |> prop_of
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
      |> Logic.strip_horn
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
      |>> map strip_full_horn
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
    val params = map (fn (x, _, _) => x) ind_prems
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   269
    val param_trms = (map o map) Free params  
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   270
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   271
    val intr_vars_tys = map (fn t => rev (Term.add_vars (prop_of t) [])) intrs
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
    val intr_vars = (map o map) fst intr_vars_tys
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
    val intr_vars_substs = map2 (curry (op ~~)) intr_vars param_trms
2637
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   274
    val intr_cvars = (map o map) (cterm_of thy o Var) intr_vars_tys      
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   276
    val (intr_prems, intr_concls) = intrs
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   277
      |> map prop_of
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   278
      |> map2 subst_Vars intr_vars_substs
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
      |> map Logic.strip_horn
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   280
      |> split_list
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   281
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
    val intr_concls_args = map (snd o strip_comb o HOLogic.dest_Trueprop) intr_concls 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   283
      
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   284
    val avoid_trms = avoids
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   285
      |> (map o map) (setify ctxt') 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   286
      |> map fold_union
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   287
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   288
    val vc_compat_goals = 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   289
      map4 mk_vc_compat (avoids ~~ avoid_trms) intr_prems intr_concls_args params
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   290
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   291
    val ([c_name, a, p], ctxt'') = Variable.variant_fixes ["c", "'a", "p"] ctxt'
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   292
    val c_ty = TFree (a, @{sort fs})
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   293
    val c = Free (c_name, c_ty)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   294
    val p = Free (p, @{typ perm})
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   295
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   296
    val (preconds, ind_concls) = ind_concl
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   297
      |> HOLogic.dest_Trueprop
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   298
      |> HOLogic.dest_conj 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   299
      |> map HOLogic.dest_imp
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   300
      |> split_list
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   301
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   302
    val Ps = map (fst o strip_comb) ind_concls
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   303
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   304
    val ind_concl' = ind_concls
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   305
      |> map (add_p_c p (c, c_ty))
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   306
      |> (curry (op ~~)) preconds  
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   307
      |> map HOLogic.mk_imp
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   308
      |> fold_conj
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   309
      |> HOLogic.mk_Trueprop
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   310
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   311
    val ind_prems' = ind_prems
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   312
      |> map2 (prep_prem Ps c_name c_ty) (avoids ~~ avoid_trms)   
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   313
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   314
    fun after_qed ctxt_outside fresh_thms ctxt = 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   315
      let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   316
        val thms = Goal.prove ctxt [] ind_prems' ind_concl' 
2637
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   317
          (prove_sinduct_tac raw_induct fresh_thms Ps avoids avoid_trms intr_cvars param_trms) 
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   318
          |> singleton (ProofContext.export ctxt ctxt_outside)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   319
          |> Datatype_Aux.split_conj_thm
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   320
          |> map (fn thm => thm RS normalise)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   321
          |> map (asm_full_simplify (HOL_basic_ss addsimps @{thms permute_zero induct_rulify})) 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   322
          |> map (Drule.rotate_prems (length ind_prems'))
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   323
          |> map zero_var_indexes
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   324
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   325
        val _ = tracing ("RESULTS\n" ^ cat_lines (map (Syntax.string_of_term ctxt o prop_of) thms))
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   326
      in
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
        ctxt
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
      end 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   329
  in
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   330
    Proof.theorem NONE (after_qed ctxt) ((map o map) (rpair []) vc_compat_goals) ctxt''
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   331
  end
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   332
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   333
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   334
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   335
fun prove_strong_inductive_cmd (pred_name, avoids) ctxt =
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   336
  let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   337
    val thy = ProofContext.theory_of ctxt;
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   338
    val ({names, ...}, {raw_induct, intrs, ...}) =
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   339
      Inductive.the_inductive ctxt (Sign.intern_const thy pred_name);
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   340
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   341
    val rule_names = 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   342
      hd names
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   343
      |> the o Induct.lookup_inductP ctxt
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   344
      |> fst o Rule_Cases.get
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   345
      |> map fst
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   346
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   347
    val _ = (case duplicates (op = o pairself fst) avoids of
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   348
        [] => ()
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   349
      | xs => error ("Duplicate case names: " ^ commas_quote (map fst xs)))
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   350
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   351
    val _ = (case subtract (op =) rule_names (map fst avoids) of
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   352
        [] => ()
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   353
      | xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs))
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   354
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   355
    val avoids_ordered = order_default (op =) [] rule_names avoids
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   356
      
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   357
    fun read_avoids avoid_trms intr =
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   358
      let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   359
        (* fixme hack *)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   360
        val (((_, ctrms), _), ctxt') = Variable.import true [intr] ctxt
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   361
        val trms = map (term_of o snd) ctrms
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   362
        val ctxt'' = fold Variable.declare_term trms ctxt' 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   363
      in
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   364
        map (Syntax.read_term ctxt'') avoid_trms 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   365
      end 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   366
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   367
    val avoid_trms = map2 read_avoids avoids_ordered intrs
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   368
  in
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   369
    prove_strong_inductive rule_names avoid_trms raw_induct intrs ctxt
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   370
  end
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   371
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   372
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   373
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   374
(* outer syntax *)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   375
local
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   376
  structure P = Parse;
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   377
  structure S = Scan
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   378
  
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   379
  val _ = Keyword.keyword "avoids"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   380
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   381
  val single_avoid_parser = 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   382
    P.name -- (P.$$$ ":" |-- P.and_list1 P.term)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   383
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   384
  val avoids_parser = 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   385
    S.optional (P.$$$ "avoids" |-- P.enum1 "|" single_avoid_parser) []
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   386
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   387
  val main_parser = P.xname -- avoids_parser
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   388
in
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   389
  val _ =
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   390
  Outer_Syntax.local_theory_to_proof "nominal_inductive"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   391
    "prove strong induction theorem for inductive predicate involving nominal datatypes"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   392
      Keyword.thy_goal (main_parser >> prove_strong_inductive_cmd)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   393
end
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   394
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   395
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   396
inductive
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   397
  Acc :: "('a::pt \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   398
where
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   399
  AccI: "(\<And>y. R y x \<Longrightarrow> Acc R y) \<Longrightarrow> Acc R x"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   400
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   401
(*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   402
equivariance Acc
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   403
*)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   404
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   405
lemma Acc_eqvt [eqvt]:
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   406
  fixes p::"perm"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   407
  assumes a: "Acc R x"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   408
  shows "Acc (p \<bullet> R) (p \<bullet> x)"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   409
using a
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   410
apply(induct)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   411
apply(rule AccI)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   412
apply(rotate_tac 1)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   413
apply(drule_tac x="-p \<bullet> y" in meta_spec)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   414
apply(simp)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   415
apply(drule meta_mp)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   416
apply(rule_tac p="p" in permute_boolE)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   417
apply(perm_simp add: permute_minus_cancel)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   418
apply(assumption)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   419
apply(assumption)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   420
done
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   421
 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   422
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   423
nominal_inductive Acc .
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   424
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   425
section {* Typing *}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   426
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   427
nominal_datatype ty =
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   428
  TVar string
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   429
| TFun ty ty ("_ \<rightarrow> _") 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   430
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   431
lemma ty_fresh:
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   432
  fixes x::"name"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   433
  and   T::"ty"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   434
  shows "atom x \<sharp> T"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   435
apply (nominal_induct T rule: ty.strong_induct)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   436
apply (simp_all add: ty.fresh pure_fresh)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   437
done
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   438
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   439
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   440
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   441
inductive
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   442
  valid :: "(name \<times> ty) list \<Rightarrow> bool"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   443
where
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   444
  "valid []"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   445
| "\<lbrakk>atom x \<sharp> Gamma; valid Gamma\<rbrakk> \<Longrightarrow> valid ((x, T)#Gamma)"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   446
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   447
inductive
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   448
  typing :: "(name\<times>ty) list \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _" [60,60,60] 60) 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   449
where
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   450
    t_Var[intro]: "\<lbrakk>valid \<Gamma>; (x, T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   451
  | t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : T1 \<rightarrow> T2; \<Gamma> \<turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : T2"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   452
  | t_Lam[intro]: "\<lbrakk>atom x \<sharp> \<Gamma>; (x, T1) # \<Gamma> \<turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam x t : T1 \<rightarrow> T2"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   453
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   454
thm typing.intros
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   455
thm typing.induct
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   456
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   457
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   458
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   459
equivariance valid
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   460
equivariance typing
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   461
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   462
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   463
nominal_inductive typing
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   464
  avoids t_Lam: "x"
2637
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   465
      (* | t_Var: "x" *)
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   466
  apply -
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   467
  apply(simp_all add: fresh_star_def ty_fresh lam.fresh)?
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   468
  done
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   469
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   470
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   471
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   472
lemma
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   473
  fixes c::"'a::fs"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   474
  assumes a: "\<Gamma> \<turnstile> t : T" 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   475
  and a1: "\<And>\<Gamma> x T c. \<lbrakk>valid \<Gamma>; (x, T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> P c \<Gamma> (Var x) T"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   476
  and a2: "\<And>\<Gamma> t1 T1 T2 t2 c. \<lbrakk>\<Gamma> \<turnstile> t1 : T1 \<rightarrow> T2; \<And>d. P d \<Gamma> t1 T1 \<rightarrow> T2; \<Gamma> \<turnstile> t2 : T1; \<And>d. P d \<Gamma> t2 T1\<rbrakk> 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   477
           \<Longrightarrow> P c \<Gamma> (App t1 t2) T2"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   478
  and a3: "\<And>x \<Gamma> T1 t T2 c. \<lbrakk>{atom x} \<sharp>* c; atom x \<sharp> \<Gamma>; (x, T1) # \<Gamma> \<turnstile> t : T2; \<And>d. P d ((x, T1) # \<Gamma>) t T2\<rbrakk> 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   479
           \<Longrightarrow> P c \<Gamma> (Lam x t) T1 \<rightarrow> T2"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   480
  shows "P c \<Gamma> t T"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   481
proof -
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   482
  from a have "\<And>p c. P c (p \<bullet> \<Gamma>) (p \<bullet> t) (p \<bullet> T)"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   483
  proof (induct)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   484
    case (t_Var \<Gamma> x T p c)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   485
    then show ?case
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   486
      apply -
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   487
      apply(perm_strict_simp)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   488
      thm a1
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   489
      apply(rule a1)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   490
      apply(drule_tac p="p" in permute_boolI)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   491
      apply(perm_strict_simp add: permute_minus_cancel)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   492
      apply(assumption)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   493
      apply(rotate_tac 1)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   494
      apply(drule_tac p="p" in permute_boolI)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   495
      apply(perm_strict_simp add: permute_minus_cancel)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   496
      apply(assumption)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   497
      done
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   498
  next
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   499
    case (t_App \<Gamma> t1 T1 T2 t2 p c)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   500
    then show ?case
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   501
      apply -
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   502
      apply(perm_strict_simp)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   503
      apply(rule a2)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   504
      apply(drule_tac p="p" in permute_boolI)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   505
      apply(perm_strict_simp add: permute_minus_cancel)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   506
      apply(assumption)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   507
      apply(assumption)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   508
      apply(rotate_tac 2)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   509
      apply(drule_tac p="p" in permute_boolI)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   510
      apply(perm_strict_simp add: permute_minus_cancel)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   511
      apply(assumption)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   512
      apply(assumption)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   513
      done
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   514
  next
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   515
    case (t_Lam x \<Gamma> T1 t T2 p c)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   516
    then show ?case
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   517
      apply -
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   518
      apply(subgoal_tac "\<exists>q. (q \<bullet> {p \<bullet> atom x}) \<sharp>* c \<and> 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   519
        supp (p \<bullet> \<Gamma>, p \<bullet> Lam x t, p \<bullet> (T1 \<rightarrow> T2)) \<sharp>* q")
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   520
      apply(erule exE)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   521
      apply(rule_tac t="p \<bullet> \<Gamma>" and  s="(q + p) \<bullet> \<Gamma>" in subst)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   522
      apply(simp only: permute_plus)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   523
      apply(rule supp_perm_eq)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   524
      apply(simp add: supp_Pair fresh_star_Un)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   525
      apply(rule_tac t="p \<bullet> Lam x t" and  s="(q + p) \<bullet> Lam x t" in subst)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   526
      apply(simp only: permute_plus)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   527
      apply(rule supp_perm_eq)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   528
      apply(simp add: supp_Pair fresh_star_Un)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   529
      apply(rule_tac t="p \<bullet> (T1 \<rightarrow> T2)" and  s="(q + p) \<bullet> (T1 \<rightarrow> T2)" in subst)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   530
      apply(simp only: permute_plus)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   531
      apply(rule supp_perm_eq)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   532
      apply(simp add: supp_Pair fresh_star_Un)
2637
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   533
      (* apply(perm_simp) *) 
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   534
      apply(simp (no_asm) only: eqvts)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   535
      apply(rule a3)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   536
      apply(simp only: eqvts permute_plus)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   537
      apply(rule_tac p="- (q + p)" in permute_boolE)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   538
      apply(perm_strict_simp add: permute_minus_cancel)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   539
      apply(assumption)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   540
       apply(rule_tac p="- (q + p)" in permute_boolE)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   541
      apply(perm_strict_simp add: permute_minus_cancel)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   542
      apply(assumption)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   543
      apply(perm_strict_simp)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   544
      apply(simp only:)
2637
3890483c674f final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
parents: 2636
diff changeset
   545
      thm at_set_avoiding2
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   546
      apply(rule at_set_avoiding2)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   547
      apply(simp add: finite_supp)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   548
      apply(simp add: finite_supp)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   549
      apply(simp add: finite_supp)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   550
      apply(rule_tac p="-p" in permute_boolE)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   551
      apply(perm_strict_simp add: permute_minus_cancel)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   552
	--"supplied by the user"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   553
      apply(simp add: fresh_star_Pair)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   554
      sorry
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   555
  qed
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   556
  then have "P c (0 \<bullet> \<Gamma>) (0 \<bullet> t) (0 \<bullet> T)" .
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   557
  then show "P c \<Gamma> t T" by simp
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   558
qed
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   559
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   560
*)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   561
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   562
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   563
end
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   564
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   565
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   566