Nominal/Ex/Typing.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 05 Jan 2011 16:51:27 +0000
changeset 2638 e1e2ca92760b
parent 2637 3890483c674f
child 2639 a8fc346deda3
permissions -rw-r--r--
strong rule inductions; as an example the weakening lemma works
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 {*
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    24
fun mk_cplus p q = Thm.capply (Thm.capply @{cterm "plus::perm \<Rightarrow> perm \<Rightarrow> perm"} p) q 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    25
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
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
    27
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    28
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
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
    30
  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
    31
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
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
    33
  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
    34
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
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
    38
  | 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
    39
  | 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
    40
  | 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
    41
  | 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
    42
  | 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
    43
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
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
    46
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
    47
  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
    48
    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
    49
      |> HOLogic.mk_tuple
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
      |> mk_fresh_star avoid_trm 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
      |> HOLogic.mk_Trueprop
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
      |> (curry Logic.list_implies) prems
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
      |> (curry list_all_free) params
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    54
    val finite_goal = avoid_trm
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    55
      |> mk_finite
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    56
      |> HOLogic.mk_Trueprop
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    57
      |> (curry Logic.list_implies) prems
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    58
      |> (curry list_all_free) params
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
  in 
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    60
    if null avoid then [] else [vc_goal, finite_goal]
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
  end
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
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
    66
  if prop trm 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
  then f trm
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
  else case trm of
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
    (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
    70
  | 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
    71
  | _ => trm
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
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
    76
  let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
    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
    78
    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
    79
    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
    80
    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
    81
  in
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
    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
    83
    |> (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
    84
    |> (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
    85
  end
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
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
    90
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
    91
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
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
    95
  let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
    fun add t = 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
      let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
        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
    99
        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
   100
        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
   101
        val args' = args
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
          |> qnt ? map (incr_boundvars 1)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
      in
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
        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
   105
        |> 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
   106
      end
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
  in
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
    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
   109
  end
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
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
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
   114
  let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
    val prems' = prems
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
      |> map (incr_boundvars 1) 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
      |> 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
   118
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
    val avoid_trm' = avoid_trm
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
      |> (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
   121
      |> strip_abs_body
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
      |> (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
   123
      |> HOLogic.mk_Trueprop
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
    val prems'' = 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
      if null avoid 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
      then prems' 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
      else avoid_trm' :: prems'
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
    val concl' = concl
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
      |> incr_boundvars 1 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
      |> 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
   133
  in
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
    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
   135
  end
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
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
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
   141
  | 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
   142
  | 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
   143
  | same_name _ = false
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
ML {*
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   147
fun map7 _ [] [] [] [] [] [] [] = []
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   148
  | map7 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (r :: rs) (s :: ss) = 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   149
      f x y z u v r s :: map7 f xs ys zs us vs rs ss
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   150
*}
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   151
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   152
ML {*
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
(* local abbreviations *)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
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
   155
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
   156
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
val all_elims = 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
  let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
     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
   162
  in
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
    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
   164
  end
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
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
   169
  Subgoal.SUBPROOF (fn {context, prems, ...} =>
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
    let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
      val prems' = prems
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
        |> map (minus_permute_elim p)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
        |> map (eqvt_srule context)
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   174
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
      val prm' = (prems' MRS prm)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
        |> flag ? (all_elims [p])
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   177
        |> flag ? (eqvt_srule context)
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   178
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   179
      val _ = tracing ("prm':" ^ @{make_string} prm')
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
    in
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   181
      print_tac "start helper"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   182
      THEN asm_full_simp_tac (HOL_ss addsimps (prm' :: @{thms induct_forall_def})) 1
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   183
      THEN print_tac "final helper"
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
    end) ctxt
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
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
   189
  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
   190
    let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
      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
   192
      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
   193
      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
   194
      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
   195
      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
   196
      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
   197
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
      (* for inductive-premises*)
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   199
      fun tac1 prm = helper_tac true prm p context 
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
      (* for non-inductive premises *)   
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
      fun tac2 prm =  
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
        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
   204
                 eqvt_stac context, 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
                 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
   206
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
      fun select prm (t, i) =
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
        (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
   209
    in
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   210
      EVERY1 [eqvt_stac ctxt, rtac prem', RANGE (map (SUBGOAL o select) prems) ]
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
    end) ctxt
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
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
   214
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
ML {*
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   216
fun fresh_thm ctxt user_thm p c concl_args avoid_trm =
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
   217
  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
   218
    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
   219
      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
   220
    val conj2 =
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   221
      mk_fresh_star_ty @{typ perm} (mk_supp (HOLogic.mk_tuple (map (mk_perm p) concl_args))) (Bound 0)
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
   222
    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
   223
      |> 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
   224
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   225
    val ss = @{thms finite_supp supp_Pair finite_Un permute_finite} @ 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   226
             @{thms fresh_star_Pair fresh_star_permute_iff}
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   227
    val simp = asm_full_simp_tac (HOL_ss addsimps ss)
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
   228
  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
   229
    Goal.prove ctxt [] [] fresh_goal
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   230
      (K (HEADGOAL (rtac @{thm at_set_avoiding2} 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   231
          THEN_ALL_NEW EVERY' [cut_facts_tac user_thm, REPEAT o etac @{thm conjE}, simp])))
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   232
  end
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   233
*}
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   234
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   235
ML {* 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   236
val supp_perm_eq' = 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   237
  @{lemma "supp (p \<bullet> x) \<sharp>* q ==> p \<bullet> x == (q + p) \<bullet> x" by (simp add: supp_perm_eq)}
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   238
val fresh_star_plus =
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   239
  @{lemma "(q \<bullet> (p \<bullet> x)) \<sharp>* c ==> ((q + p) \<bullet> x) \<sharp>* c" by (simp add: permute_plus)}
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   240
*}
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   241
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   242
ML {*
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   243
fun binder_tac prem intr_cvars param_trms Ps user_thm avoid avoid_trm concl_args ctxt = 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   244
  Subgoal.FOCUS (fn {context = ctxt, params, prems, concl, ...} =>
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   245
    let
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   246
      val thy = ProofContext.theory_of ctxt
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   247
      val (prms, p, c) = split_last2 (map snd params)
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   248
      val prm_trms = map term_of prms
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   249
      val prm_tys = map fastype_of prm_trms
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   250
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   251
      val avoid_trm' = subst_free (param_trms ~~ prm_trms) avoid_trm 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   252
      val concl_args' = map (subst_free (param_trms ~~ prm_trms)) concl_args 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   253
      
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   254
      val user_thm' = map (cterm_instantiate (intr_cvars ~~ prms)) user_thm
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   255
        |> map (full_simplify (HOL_ss addsimps (@{thm fresh_star_Pair}::prems)))
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   256
      
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   257
      val fthm = fresh_thm ctxt user_thm' (term_of p) (term_of c) concl_args' avoid_trm'
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   258
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   259
      val (([(_, q)], fprop :: fresh_eqs), ctxt') = Obtain.result
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   260
              (K (EVERY1 [etac @{thm exE}, 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   261
                          full_simp_tac (HOL_basic_ss addsimps @{thms supp_Pair fresh_star_Un}), 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   262
                          REPEAT o etac @{thm conjE},
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   263
                          dtac fresh_star_plus,
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   264
                          REPEAT o dtac supp_perm_eq'])) [fthm] ctxt 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   265
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   266
      val expand_conv = Conv.try_conv (Conv.rewrs_conv fresh_eqs)
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   267
      fun expand_conv_bot ctxt = Conv.bottom_conv (K expand_conv) ctxt
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   268
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   269
      val cperms = map (cterm_of thy o perm_const) prm_tys
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   270
      val qp_prms = map2 (fn ct1 => fn ct2 => Thm.mk_binop ct1 (mk_cplus q p) ct2) cperms prms 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   271
      val prem' = cterm_instantiate (intr_cvars ~~ qp_prms) prem
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   272
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   273
      val fprop' = eqvt_srule ctxt' fprop 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   274
      val tac_fresh = simp_tac (HOL_basic_ss addsimps [fprop'])
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   275
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   276
      (* for inductive-premises*)
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   277
      fun tac1 prm = helper_tac true prm (mk_cplus q p) ctxt' 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   278
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   279
      (* for non-inductive premises *)   
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   280
      fun tac2 prm =  
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   281
        EVERY' [ minus_permute_intro_tac (mk_cplus q p), 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   282
                 eqvt_stac ctxt, 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   283
                 helper_tac false prm (mk_cplus q p) ctxt' ]
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   284
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   285
      fun select prm (t, i) =
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   286
        (if member same_name Ps (real_head_of t) then tac1 prm else tac2 prm) i
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   287
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   288
      val _ = tracing ("fthm:\n" ^ @{make_string} fthm)
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   289
      val _ = tracing ("fr_eqs:\n" ^ cat_lines (map @{make_string} fresh_eqs))
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   290
      val _ = tracing ("fprop:\n" ^ @{make_string} fprop)
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   291
      val _ = tracing ("fprop':\n" ^ @{make_string} fprop')
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   292
      val _ = tracing ("fperm:\n" ^ @{make_string} q)
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   293
      val _ = tracing ("prem':\n" ^ @{make_string} prem')
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   294
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   295
      val side_thm = Goal.prove ctxt' [] [] (term_of concl)
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   296
        (fn {context, ...} => 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   297
           EVERY1 [ CONVERSION (expand_conv_bot context),
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   298
                    eqvt_stac context,
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   299
                    rtac prem',
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   300
                    RANGE (tac_fresh :: map (SUBGOAL o select) prems),
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   301
                    K (print_tac "GOAL") ])
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   302
        |> singleton (ProofContext.export ctxt' ctxt)        
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   303
    in
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   304
      rtac side_thm 1
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   305
    end) ctxt
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   306
*}
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   307
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   308
ML {*
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   309
fun case_tac ctxt Ps avoid avoid_trm intr_cvars param_trms prem user_thm concl_args =
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   310
  let
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   311
    val tac1 = non_binder_tac prem intr_cvars Ps ctxt
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   312
    val tac2 = binder_tac prem intr_cvars param_trms Ps user_thm avoid avoid_trm concl_args ctxt
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   313
  in 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   314
    EVERY' [ rtac @{thm allI}, rtac @{thm allI}, if null avoid then tac1 else tac2 ]
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
   315
  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
   316
*}
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
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
   318
ML {*
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   319
fun prove_sinduct_tac raw_induct user_thms Ps avoids avoid_trms intr_cvars param_trms concl_args 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   320
  {prems, context} =
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   321
  let
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   322
    val cases_tac = 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   323
      map7 (case_tac context Ps) avoids avoid_trms intr_cvars param_trms prems user_thms concl_args
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   324
  in 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   325
    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
   326
  end
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   329
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
   330
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
   331
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   332
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   333
ML {* Local_Theory.note *}
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   334
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   335
ML {*
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   336
fun prove_strong_inductive pred_names rule_names avoids raw_induct intrs ctxt =
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   337
  let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   338
    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
   339
    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
   340
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   341
    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
   342
      |> prop_of
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   343
      |> Logic.strip_horn
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   344
      |>> map strip_full_horn
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   345
    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
   346
    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
   347
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   348
    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
   349
    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
   350
    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
   351
    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
   352
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   353
    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
   354
      |> map prop_of
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   355
      |> 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
   356
      |> map Logic.strip_horn
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   357
      |> split_list
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   358
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   359
    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
   360
      
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   361
    val avoid_trms = avoids
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   362
      |> (map o map) (setify ctxt') 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   363
      |> map fold_union
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   364
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   365
    val vc_compat_goals = 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   366
      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
   367
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   368
    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
   369
    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
   370
    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
   371
    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
   372
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   373
    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
   374
      |> HOLogic.dest_Trueprop
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   375
      |> HOLogic.dest_conj 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   376
      |> map HOLogic.dest_imp
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   377
      |> split_list
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 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
   380
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   381
    val ind_concl' = ind_concls
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   382
      |> 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
   383
      |> (curry (op ~~)) preconds  
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   384
      |> map HOLogic.mk_imp
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   385
      |> fold_conj
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   386
      |> HOLogic.mk_Trueprop
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   387
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   388
    val ind_prems' = ind_prems
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   389
      |> 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
   390
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   391
    fun after_qed ctxt_outside user_thms ctxt = 
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   392
      let
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   393
        val strong_ind_thms = Goal.prove ctxt [] ind_prems' ind_concl' 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   394
        (prove_sinduct_tac raw_induct user_thms Ps avoids avoid_trms intr_cvars param_trms intr_concls_args) 
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   395
          |> 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
   396
          |> Datatype_Aux.split_conj_thm
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   397
          |> 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
   398
          |> 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
   399
          |> 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
   400
          |> map zero_var_indexes
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   401
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   402
        val qualified_thm_name = pred_names
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   403
          |> map Long_Name.base_name
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   404
          |> space_implode "_"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   405
          |> (fn s => Binding.qualify false s (Binding.name "strong_induct"))
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   406
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   407
        val attrs = 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   408
          [ Attrib.internal (K (Rule_Cases.consumes 1)),
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   409
            Attrib.internal (K (Rule_Cases.case_names rule_names)) ]
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   410
        val _ = tracing ("RESULTS\n" ^ cat_lines (map (Syntax.string_of_term ctxt o prop_of) strong_ind_thms))
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   411
        val _ = tracing ("rule_names: " ^ commas rule_names)
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   412
        val _ = tracing ("pred_names: " ^ commas pred_names)
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   413
      in
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   414
        ctxt
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   415
        |> Local_Theory.note ((qualified_thm_name, attrs), strong_ind_thms)    
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   416
        |> snd   
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   417
      end
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   418
  in
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   419
    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
   420
  end
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
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   424
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
   425
  let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   426
    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
   427
    val ({names, ...}, {raw_induct, intrs, ...}) =
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   428
      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
   429
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   430
    val rule_names = 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   431
      hd names
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   432
      |> 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
   433
      |> fst o Rule_Cases.get
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   434
      |> map fst
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   435
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   436
    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
   437
        [] => ()
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   438
      | 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
   439
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   440
    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
   441
        [] => ()
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   442
      | 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
   443
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   444
    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
   445
      
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   446
    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
   447
      let
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   448
        (* fixme hack *)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   449
        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
   450
        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
   451
        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
   452
      in
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   453
        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
   454
      end 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   455
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   456
    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
   457
  in
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   458
    prove_strong_inductive names rule_names avoid_trms raw_induct intrs ctxt
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   459
  end
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   460
*}
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
ML {*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   463
(* outer syntax *)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   464
local
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   465
  structure P = Parse;
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   466
  structure S = Scan
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   467
  
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   468
  val _ = Keyword.keyword "avoids"
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
  val single_avoid_parser = 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   471
    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
   472
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   473
  val avoids_parser = 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   474
    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
   475
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   476
  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
   477
in
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   478
  val _ =
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   479
  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
   480
    "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
   481
      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
   482
end
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   483
*}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   484
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   485
inductive
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   486
  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
   487
where
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   488
  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
   489
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   490
(*
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   491
equivariance Acc
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   492
*)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   493
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   494
lemma Acc_eqvt [eqvt]:
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   495
  fixes p::"perm"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   496
  assumes a: "Acc R x"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   497
  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
   498
using a
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   499
apply(induct)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   500
apply(rule AccI)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   501
apply(rotate_tac 1)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   502
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
   503
apply(simp)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   504
apply(drule meta_mp)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   505
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
   506
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
   507
apply(assumption)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   508
apply(assumption)
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   509
done
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   510
 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   511
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   512
nominal_inductive Acc .
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   513
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   514
thm Acc.strong_induct
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   515
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   516
section {* Typing *}
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   517
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   518
nominal_datatype ty =
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   519
  TVar string
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   520
| TFun ty ty ("_ \<rightarrow> _") 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   521
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   522
lemma ty_fresh:
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   523
  fixes x::"name"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   524
  and   T::"ty"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   525
  shows "atom x \<sharp> T"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   526
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
   527
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
   528
done
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   529
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   530
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   531
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   532
inductive
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   533
  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
   534
where
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   535
  v_Nil[intro]: "valid []"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   536
| v_Cons[intro]: "\<lbrakk>atom x \<sharp> Gamma; valid Gamma\<rbrakk> \<Longrightarrow> valid ((x, T)#Gamma)"
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   537
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   538
inductive
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   539
  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
   540
where
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   541
    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
   542
  | 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
   543
  | 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
   544
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   545
thm typing.intros
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   546
thm typing.induct
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   547
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   548
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   549
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   550
equivariance valid
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   551
equivariance typing
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   552
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   553
nominal_inductive typing
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   554
  avoids t_Lam: "x"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   555
  apply -
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   556
  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
   557
  done
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   558
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   559
thm typing.strong_induct
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   560
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   561
abbreviation
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   562
  "sub_context" :: "(name \<times> ty) list \<Rightarrow> (name \<times> ty) list \<Rightarrow> bool" ("_ \<subseteq> _" [60,60] 60) 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   563
where
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   564
  "\<Gamma>1 \<subseteq> \<Gamma>2 \<equiv> \<forall>x T. (x, T) \<in> set \<Gamma>1 \<longrightarrow> (x, T) \<in> set \<Gamma>2"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   565
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   566
text {* Now it comes: The Weakening Lemma *}
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   567
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   568
text {*
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   569
  The first version is, after setting up the induction, 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   570
  completely automatic except for use of atomize. *}
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   571
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   572
lemma weakening_version2: 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   573
  fixes \<Gamma>1 \<Gamma>2::"(name \<times> ty) list"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   574
  and   t ::"lam"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   575
  and   \<tau> ::"ty"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   576
  assumes a: "\<Gamma>1 \<turnstile> t : T"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   577
  and     b: "valid \<Gamma>2" 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   578
  and     c: "\<Gamma>1 \<subseteq> \<Gamma>2"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   579
  shows "\<Gamma>2 \<turnstile> t : T"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   580
using a b c
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   581
proof (nominal_induct \<Gamma>1 t T avoiding: \<Gamma>2 rule: typing.strong_induct)
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   582
  case (t_Var \<Gamma>1 x T)  (* variable case *)
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   583
  have "\<Gamma>1 \<subseteq> \<Gamma>2" by fact 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   584
  moreover  
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   585
  have "valid \<Gamma>2" by fact 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   586
  moreover 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   587
  have "(x,T)\<in> set \<Gamma>1" by fact
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   588
  ultimately show "\<Gamma>2 \<turnstile> Var x : T" by auto
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   589
next
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   590
  case (t_Lam x \<Gamma>1 T1 t T2) (* lambda case *)
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   591
  have vc: "atom x \<sharp> \<Gamma>2" by fact   (* variable convention *)
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   592
  have ih: "\<lbrakk>valid ((x, T1) # \<Gamma>2); (x, T1) # \<Gamma>1 \<subseteq> (x, T1) # \<Gamma>2\<rbrakk> \<Longrightarrow> (x, T1) # \<Gamma>2 \<turnstile> t : T2" by fact
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   593
  have "\<Gamma>1 \<subseteq> \<Gamma>2" by fact
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   594
  then have "(x, T1) # \<Gamma>1 \<subseteq> (x, T1) # \<Gamma>2" by simp
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   595
  moreover
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   596
  have "valid \<Gamma>2" by fact
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   597
  then have "valid ((x, T1) # \<Gamma>2)" using vc by (simp add: v_Cons)
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   598
  ultimately have "(x, T1) # \<Gamma>2 \<turnstile> t : T2" using ih by simp
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   599
  with vc show "\<Gamma>2 \<turnstile> Lam x t : T1 \<rightarrow> T2" by auto
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   600
qed (auto) (* app case *)
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   601
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   602
lemma weakening_version1: 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   603
  fixes \<Gamma>1 \<Gamma>2::"(name \<times> ty) list"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   604
  assumes a: "\<Gamma>1 \<turnstile> t : T" 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   605
  and     b: "valid \<Gamma>2" 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   606
  and     c: "\<Gamma>1 \<subseteq> \<Gamma>2"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   607
  shows "\<Gamma>2 \<turnstile> t : T"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   608
using a b c
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   609
apply (nominal_induct \<Gamma>1 t T avoiding: \<Gamma>2 rule: typing.strong_induct)
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   610
apply (auto | atomize)+
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   611
done
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
   612
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   613
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   614
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   615
end
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   616
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   617
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   618