src/HOL/Nominal/nominal_permeq.ML
author wenzelm
Mon, 22 May 2017 00:23:25 +0200
changeset 65897 94b0da1b242e
parent 62913 13252110a6fe
child 67710 cc2db3239932
permissions -rw-r--r--
back to scala-2.11.8 due to apparent non-termination of HOL-Codegenerator_Test;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19477
diff changeset
     1
(*  Title:      HOL/Nominal/nominal_permeq.ML
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32733
diff changeset
     2
    Author:     Christian Urban, TU Muenchen
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32733
diff changeset
     3
    Author:     Julien Narboux, TU Muenchen
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     4
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32733
diff changeset
     5
Methods for simplifying permutations and for analysing equations
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
     6
involving permutations.
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19477
diff changeset
     7
*)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     8
20431
eef4e9081bea added a FIXME-comment
urbanc
parents: 20289
diff changeset
     9
(*
eef4e9081bea added a FIXME-comment
urbanc
parents: 20289
diff changeset
    10
FIXMES:
eef4e9081bea added a FIXME-comment
urbanc
parents: 20289
diff changeset
    11
eef4e9081bea added a FIXME-comment
urbanc
parents: 20289
diff changeset
    12
 - allow the user to give an explicit set S in the
eef4e9081bea added a FIXME-comment
urbanc
parents: 20289
diff changeset
    13
   fresh_guess tactic which is then verified
eef4e9081bea added a FIXME-comment
urbanc
parents: 20289
diff changeset
    14
eef4e9081bea added a FIXME-comment
urbanc
parents: 20289
diff changeset
    15
 - the perm_compose tactic does not do an "outermost
eef4e9081bea added a FIXME-comment
urbanc
parents: 20289
diff changeset
    16
   rewriting" and can therefore not deal with goals
eef4e9081bea added a FIXME-comment
urbanc
parents: 20289
diff changeset
    17
   like
eef4e9081bea added a FIXME-comment
urbanc
parents: 20289
diff changeset
    18
eef4e9081bea added a FIXME-comment
urbanc
parents: 20289
diff changeset
    19
      [(a,b)] o pi1 o pi2 = ....
eef4e9081bea added a FIXME-comment
urbanc
parents: 20289
diff changeset
    20
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    21
   rather it tries to permute pi1 over pi2, which
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    22
   results in a failure when used with the
20431
eef4e9081bea added a FIXME-comment
urbanc
parents: 20289
diff changeset
    23
   perm_(full)_simp tactics
eef4e9081bea added a FIXME-comment
urbanc
parents: 20289
diff changeset
    24
eef4e9081bea added a FIXME-comment
urbanc
parents: 20289
diff changeset
    25
*)
eef4e9081bea added a FIXME-comment
urbanc
parents: 20289
diff changeset
    26
eef4e9081bea added a FIXME-comment
urbanc
parents: 20289
diff changeset
    27
19987
b97607d4d9a5 - put declarations inside a structure (NominalPermeq)
berghofe
parents: 19858
diff changeset
    28
signature NOMINAL_PERMEQ =
b97607d4d9a5 - put declarations inside a structure (NominalPermeq)
berghofe
parents: 19858
diff changeset
    29
sig
25997
0c0f5d990d7d Cleaned up simproc code.
berghofe
parents: 24571
diff changeset
    30
  val perm_simproc_fun : simproc
0c0f5d990d7d Cleaned up simproc code.
berghofe
parents: 24571
diff changeset
    31
  val perm_simproc_app : simproc
0c0f5d990d7d Cleaned up simproc code.
berghofe
parents: 24571
diff changeset
    32
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
    33
  val perm_simp_tac : Proof.context -> int -> tactic
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
    34
  val perm_extend_simp_tac : Proof.context -> int -> tactic
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
    35
  val supports_tac : Proof.context -> int -> tactic
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
    36
  val finite_guess_tac : Proof.context -> int -> tactic
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
    37
  val fresh_guess_tac : Proof.context -> int -> tactic
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    38
30549
d2d7874648bd simplified method setup;
wenzelm
parents: 30510
diff changeset
    39
  val perm_simp_meth : (Proof.context -> Proof.method) context_parser
d2d7874648bd simplified method setup;
wenzelm
parents: 30510
diff changeset
    40
  val perm_simp_meth_debug : (Proof.context -> Proof.method) context_parser
d2d7874648bd simplified method setup;
wenzelm
parents: 30510
diff changeset
    41
  val perm_extend_simp_meth : (Proof.context -> Proof.method) context_parser
d2d7874648bd simplified method setup;
wenzelm
parents: 30510
diff changeset
    42
  val perm_extend_simp_meth_debug : (Proof.context -> Proof.method) context_parser
d2d7874648bd simplified method setup;
wenzelm
parents: 30510
diff changeset
    43
  val supports_meth : (Proof.context -> Proof.method) context_parser
d2d7874648bd simplified method setup;
wenzelm
parents: 30510
diff changeset
    44
  val supports_meth_debug : (Proof.context -> Proof.method) context_parser
d2d7874648bd simplified method setup;
wenzelm
parents: 30510
diff changeset
    45
  val finite_guess_meth : (Proof.context -> Proof.method) context_parser
d2d7874648bd simplified method setup;
wenzelm
parents: 30510
diff changeset
    46
  val finite_guess_meth_debug : (Proof.context -> Proof.method) context_parser
d2d7874648bd simplified method setup;
wenzelm
parents: 30510
diff changeset
    47
  val fresh_guess_meth : (Proof.context -> Proof.method) context_parser
d2d7874648bd simplified method setup;
wenzelm
parents: 30510
diff changeset
    48
  val fresh_guess_meth_debug : (Proof.context -> Proof.method) context_parser
19987
b97607d4d9a5 - put declarations inside a structure (NominalPermeq)
berghofe
parents: 19858
diff changeset
    49
end
b97607d4d9a5 - put declarations inside a structure (NominalPermeq)
berghofe
parents: 19858
diff changeset
    50
b97607d4d9a5 - put declarations inside a structure (NominalPermeq)
berghofe
parents: 19858
diff changeset
    51
structure NominalPermeq : NOMINAL_PERMEQ =
b97607d4d9a5 - put declarations inside a structure (NominalPermeq)
berghofe
parents: 19858
diff changeset
    52
struct
b97607d4d9a5 - put declarations inside a structure (NominalPermeq)
berghofe
parents: 19858
diff changeset
    53
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    54
(* some lemmas needed below *)
24519
5c435b2ea137 made theorem-references safe
urbanc
parents: 22808
diff changeset
    55
val finite_emptyI = @{thm "finite.emptyI"};
5c435b2ea137 made theorem-references safe
urbanc
parents: 22808
diff changeset
    56
val finite_Un     = @{thm "finite_Un"};
5c435b2ea137 made theorem-references safe
urbanc
parents: 22808
diff changeset
    57
val conj_absorb   = @{thm "conj_absorb"};
5c435b2ea137 made theorem-references safe
urbanc
parents: 22808
diff changeset
    58
val not_false     = @{thm "not_False_eq_True"}
44684
8dde3352d5c4 assert Pure equations for theorem references; tuned
haftmann
parents: 43278
diff changeset
    59
val perm_fun_def  = Simpdata.mk_eq @{thm "Nominal.perm_fun_def"};
24519
5c435b2ea137 made theorem-references safe
urbanc
parents: 22808
diff changeset
    60
val perm_eq_app   = @{thm "Nominal.pt_fun_app_eq"};
44684
8dde3352d5c4 assert Pure equations for theorem references; tuned
haftmann
parents: 43278
diff changeset
    61
val supports_def  = Simpdata.mk_eq @{thm "Nominal.supports_def"};
8dde3352d5c4 assert Pure equations for theorem references; tuned
haftmann
parents: 43278
diff changeset
    62
val fresh_def     = Simpdata.mk_eq @{thm "Nominal.fresh_def"};
24519
5c435b2ea137 made theorem-references safe
urbanc
parents: 22808
diff changeset
    63
val fresh_prod    = @{thm "Nominal.fresh_prod"};
5c435b2ea137 made theorem-references safe
urbanc
parents: 22808
diff changeset
    64
val fresh_unit    = @{thm "Nominal.fresh_unit"};
5c435b2ea137 made theorem-references safe
urbanc
parents: 22808
diff changeset
    65
val supports_rule = @{thm "supports_finite"};
5c435b2ea137 made theorem-references safe
urbanc
parents: 22808
diff changeset
    66
val supp_prod     = @{thm "supp_prod"};
5c435b2ea137 made theorem-references safe
urbanc
parents: 22808
diff changeset
    67
val supp_unit     = @{thm "supp_unit"};
5c435b2ea137 made theorem-references safe
urbanc
parents: 22808
diff changeset
    68
val pt_perm_compose_aux = @{thm "pt_perm_compose_aux"};
5c435b2ea137 made theorem-references safe
urbanc
parents: 22808
diff changeset
    69
val cp1_aux             = @{thm "cp1_aux"};
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    70
val perm_aux_fold       = @{thm "perm_aux_fold"};
24519
5c435b2ea137 made theorem-references safe
urbanc
parents: 22808
diff changeset
    71
val supports_fresh_rule = @{thm "supports_fresh"};
21669
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21588
diff changeset
    72
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    73
(* needed in the process of fully simplifying permutations *)
24519
5c435b2ea137 made theorem-references safe
urbanc
parents: 22808
diff changeset
    74
val strong_congs = [@{thm "if_cong"}]
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    75
(* needed to avoid warnings about overwritten congs *)
24519
5c435b2ea137 made theorem-references safe
urbanc
parents: 22808
diff changeset
    76
val weak_congs   = [@{thm "if_weak_cong"}]
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    77
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    78
(* debugging *)
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    79
fun DEBUG ctxt (msg,tac) =
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    80
    CHANGED (EVERY [print_tac ctxt ("before "^msg), tac, print_tac ctxt ("after "^msg)]);
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    81
fun NO_DEBUG _ (_,tac) = CHANGED tac;
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    82
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19350
diff changeset
    83
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    84
(* simproc that deals with instances of permutations in front *)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    85
(* of applications; just adding this rule to the simplifier   *)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    86
(* would loop; it also needs careful tuning with the simproc  *)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    87
(* for functions to avoid further possibilities for looping   *)
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    88
fun perm_simproc_app' ctxt ct =
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    89
  let
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    90
    val thy = Proof_Context.theory_of ctxt
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    91
    val redex = Thm.term_of ct
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    92
    (* the "application" case is only applicable when the head of f is not a *)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    93
    (* constant or when (f x) is a permuation with two or more arguments     *)
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    94
    fun applicable_app t =
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    95
          (case (strip_comb t) of
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56230
diff changeset
    96
              (Const (@{const_name Nominal.perm},_),ts) => (length ts) >= 2
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    97
            | (Const _,_) => false
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    98
            | _ => true)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    99
  in
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   100
    case redex of
19169
20a73345dd6e fixed a problem where a permutation is not analysed
urbanc
parents: 19163
diff changeset
   101
        (* case pi o (f x) == (pi o f) (pi o x)          *)
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56230
diff changeset
   102
        (Const(@{const_name Nominal.perm},
83b3c110f22d more antiquotations;
wenzelm
parents: 56230
diff changeset
   103
          Type(@{type_name fun},
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   104
            [Type(@{type_name list}, [Type(@{type_name prod},[Type(n,_),_])]),_])) $ pi $ (f $ x)) =>
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   105
            (if (applicable_app f) then
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   106
              let
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30280
diff changeset
   107
                val name = Long_Name.base_name n
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   108
                val at_inst = Global_Theory.get_thm thy ("at_" ^ name ^ "_inst")
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   109
                val pt_inst = Global_Theory.get_thm thy ("pt_" ^ name ^ "_inst")
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   110
              in SOME ((at_inst RS (pt_inst RS perm_eq_app)) RS eq_reflection) end
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   111
            else NONE)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   112
      | _ => NONE
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   113
  end
19139
af69e41eab5d improved the decision-procedure for permutations;
urbanc
parents: 18434
diff changeset
   114
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   115
val perm_simproc_app =
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   116
  Simplifier.make_simproc @{context} "perm_simproc_app"
62913
13252110a6fe eliminated unused simproc identifier;
wenzelm
parents: 61144
diff changeset
   117
   {lhss = [@{term "Nominal.perm pi x"}], proc = K perm_simproc_app'}
25997
0c0f5d990d7d Cleaned up simproc code.
berghofe
parents: 24571
diff changeset
   118
24519
5c435b2ea137 made theorem-references safe
urbanc
parents: 22808
diff changeset
   119
(* a simproc that deals with permutation instances in front of functions  *)
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   120
fun perm_simproc_fun' ctxt ct =
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   121
   let
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   122
     val redex = Thm.term_of ct
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   123
     fun applicable_fun t =
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   124
       (case (strip_comb t) of
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   125
          (Abs _ ,[]) => true
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56230
diff changeset
   126
        | (Const (@{const_name Nominal.perm},_),_) => false
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   127
        | (Const _, _) => true
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32733
diff changeset
   128
        | _ => false)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   129
   in
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   130
     case redex of
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   131
       (* case pi o f == (%x. pi o (f ((rev pi)o x))) *)
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   132
       (Const(@{const_name Nominal.perm},_) $ pi $ f)  =>
44830
haftmann
parents: 44693
diff changeset
   133
          (if applicable_fun f then SOME perm_fun_def else NONE)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   134
      | _ => NONE
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   135
   end
19139
af69e41eab5d improved the decision-procedure for permutations;
urbanc
parents: 18434
diff changeset
   136
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   137
val perm_simproc_fun =
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   138
  Simplifier.make_simproc @{context} "perm_simproc_fun"
62913
13252110a6fe eliminated unused simproc identifier;
wenzelm
parents: 61144
diff changeset
   139
   {lhss = [@{term "Nominal.perm pi x"}], proc = K perm_simproc_fun'}
25997
0c0f5d990d7d Cleaned up simproc code.
berghofe
parents: 24571
diff changeset
   140
28322
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   141
(* function for simplyfying permutations          *)
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   142
(* stac contains the simplifiation tactic that is *)
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   143
(* applied (see (no_asm) options below            *)
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   144
fun perm_simp_gen stac dyn_thms eqvt_thms ctxt i =
56230
3e449273942a more standard method_setup;
wenzelm
parents: 55990
diff changeset
   145
    ("general simplification of permutations", fn st => SUBGOAL (fn _ =>
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   146
    let
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   147
       val ctxt' = ctxt
60359
cb8828b859a1 clarified context;
wenzelm
parents: 59621
diff changeset
   148
         addsimps (maps (Proof_Context.get_thms ctxt) dyn_thms @ eqvt_thms)
25997
0c0f5d990d7d Cleaned up simproc code.
berghofe
parents: 24571
diff changeset
   149
         addsimprocs [perm_simproc_fun, perm_simproc_app]
45620
f2a587696afb modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 44830
diff changeset
   150
         |> fold Simplifier.del_cong weak_congs
f2a587696afb modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 44830
diff changeset
   151
         |> fold Simplifier.add_cong strong_congs
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19350
diff changeset
   152
    in
56230
3e449273942a more standard method_setup;
wenzelm
parents: 55990
diff changeset
   153
      stac ctxt' i
3e449273942a more standard method_setup;
wenzelm
parents: 55990
diff changeset
   154
    end) i st);
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19350
diff changeset
   155
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   156
(* general simplification of permutations and permutation that arose from eqvt-problems *)
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   157
fun perm_simp stac ctxt =
22610
c8b5133045f3 tuned slightly the previous commit
urbanc
parents: 22609
diff changeset
   158
    let val simps = ["perm_swap","perm_fresh_fresh","perm_bij","perm_pi_simp","swap_simps"]
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   159
    in
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   160
        perm_simp_gen stac simps [] ctxt
22610
c8b5133045f3 tuned slightly the previous commit
urbanc
parents: 22609
diff changeset
   161
    end;
c8b5133045f3 tuned slightly the previous commit
urbanc
parents: 22609
diff changeset
   162
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   163
fun eqvt_simp stac ctxt =
22610
c8b5133045f3 tuned slightly the previous commit
urbanc
parents: 22609
diff changeset
   164
    let val simps = ["perm_swap","perm_fresh_fresh","perm_pi_simp"]
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   165
        val eqvts_thms = NominalThmDecls.get_eqvt_thms ctxt;
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   166
    in
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   167
        perm_simp_gen stac simps eqvts_thms ctxt
22610
c8b5133045f3 tuned slightly the previous commit
urbanc
parents: 22609
diff changeset
   168
    end;
22562
80b814fe284b rename bij and fresh into bijs and freshs and lookup for eqvts lemma using the ml value instead of the name
narboux
parents: 22541
diff changeset
   169
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   170
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   171
(* main simplification tactics for permutations *)
56491
a8ccf3d6a6e4 proper context for print_tac;
wenzelm
parents: 56253
diff changeset
   172
fun perm_simp_tac_gen_i stac tactical ctxt i = DETERM (tactical ctxt (perm_simp stac ctxt i));
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   173
fun eqvt_simp_tac_gen_i stac tactical ctxt i = DETERM (tactical ctxt (eqvt_simp stac ctxt i));
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   174
28322
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   175
val perm_simp_tac_i          = perm_simp_tac_gen_i simp_tac
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   176
val perm_asm_simp_tac_i      = perm_simp_tac_gen_i asm_simp_tac
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   177
val perm_full_simp_tac_i     = perm_simp_tac_gen_i full_simp_tac
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   178
val perm_asm_lr_simp_tac_i   = perm_simp_tac_gen_i asm_lr_simp_tac
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   179
val perm_asm_full_simp_tac_i = perm_simp_tac_gen_i asm_full_simp_tac
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   180
val eqvt_asm_full_simp_tac_i = eqvt_simp_tac_gen_i asm_full_simp_tac
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   181
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19350
diff changeset
   182
(* applies the perm_compose rule such that                             *)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19350
diff changeset
   183
(*   pi o (pi' o lhs) = rhs                                            *)
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   184
(* is transformed to                                                   *)
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19350
diff changeset
   185
(*  (pi o pi') o (pi' o lhs) = rhs                                     *)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19350
diff changeset
   186
(*                                                                     *)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19350
diff changeset
   187
(* this rule would loop in the simplifier, so some trick is used with  *)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19350
diff changeset
   188
(* generating perm_aux'es for the outermost permutation and then un-   *)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19350
diff changeset
   189
(* folding the definition                                              *)
25997
0c0f5d990d7d Cleaned up simproc code.
berghofe
parents: 24571
diff changeset
   190
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   191
fun perm_compose_simproc' ctxt ct =
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   192
  (case Thm.term_of ct of
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56230
diff changeset
   193
     (Const (@{const_name Nominal.perm}, Type (@{type_name fun}, [Type (@{type_name list},
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   194
       [Type (@{type_name Product_Type.prod}, [T as Type (tname,_),_])]),_])) $ pi1 $ (Const (@{const_name Nominal.perm},
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   195
         Type (@{type_name fun}, [Type (@{type_name list}, [Type (@{type_name Product_Type.prod}, [U as Type (uname,_),_])]),_])) $
25997
0c0f5d990d7d Cleaned up simproc code.
berghofe
parents: 24571
diff changeset
   196
          pi2 $ t)) =>
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19350
diff changeset
   197
    let
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   198
      val thy = Proof_Context.theory_of ctxt
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30280
diff changeset
   199
      val tname' = Long_Name.base_name tname
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30280
diff changeset
   200
      val uname' = Long_Name.base_name uname
25997
0c0f5d990d7d Cleaned up simproc code.
berghofe
parents: 24571
diff changeset
   201
    in
0c0f5d990d7d Cleaned up simproc code.
berghofe
parents: 24571
diff changeset
   202
      if pi1 <> pi2 then  (* only apply the composition rule in this case *)
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   203
        if T = U then
60801
7664e0916eec tuned signature;
wenzelm
parents: 60787
diff changeset
   204
          SOME (Thm.instantiate'
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   205
            [SOME (Thm.global_ctyp_of thy (fastype_of t))]
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   206
            [SOME (Thm.global_cterm_of thy pi1), SOME (Thm.global_cterm_of thy pi2), SOME (Thm.global_cterm_of thy t)]
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   207
            (mk_meta_eq ([Global_Theory.get_thm thy ("pt_"^tname'^"_inst"),
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   208
             Global_Theory.get_thm thy ("at_"^tname'^"_inst")] MRS pt_perm_compose_aux)))
25997
0c0f5d990d7d Cleaned up simproc code.
berghofe
parents: 24571
diff changeset
   209
        else
60801
7664e0916eec tuned signature;
wenzelm
parents: 60787
diff changeset
   210
          SOME (Thm.instantiate'
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   211
            [SOME (Thm.global_ctyp_of thy (fastype_of t))]
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   212
            [SOME (Thm.global_cterm_of thy pi1), SOME (Thm.global_cterm_of thy pi2), SOME (Thm.global_cterm_of thy t)]
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   213
            (mk_meta_eq (Global_Theory.get_thm thy ("cp_"^tname'^"_"^uname'^"_inst") RS
25997
0c0f5d990d7d Cleaned up simproc code.
berghofe
parents: 24571
diff changeset
   214
             cp1_aux)))
0c0f5d990d7d Cleaned up simproc code.
berghofe
parents: 24571
diff changeset
   215
      else NONE
0c0f5d990d7d Cleaned up simproc code.
berghofe
parents: 24571
diff changeset
   216
    end
0c0f5d990d7d Cleaned up simproc code.
berghofe
parents: 24571
diff changeset
   217
  | _ => NONE);
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19350
diff changeset
   218
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   219
val perm_compose_simproc =
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   220
  Simplifier.make_simproc @{context} "perm_compose"
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   221
   {lhss = [@{term "Nominal.perm pi1 (Nominal.perm pi2 t)"}],
62913
13252110a6fe eliminated unused simproc identifier;
wenzelm
parents: 61144
diff changeset
   222
    proc = K perm_compose_simproc'}
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19350
diff changeset
   223
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   224
fun perm_compose_tac ctxt i =
25997
0c0f5d990d7d Cleaned up simproc code.
berghofe
parents: 24571
diff changeset
   225
  ("analysing permutation compositions on the lhs",
0c0f5d990d7d Cleaned up simproc code.
berghofe
parents: 24571
diff changeset
   226
   fn st => EVERY
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60359
diff changeset
   227
     [resolve_tac ctxt [trans] i,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   228
      asm_full_simp_tac (empty_simpset ctxt addsimprocs [perm_compose_simproc]) i,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   229
      asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [perm_aux_fold]) i] st);
18012
23e6cfda8c4b Added (optional) arguments to the tactics
urbanc
parents: 17870
diff changeset
   230
58956
a816aa3ff391 proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
wenzelm
parents: 56492
diff changeset
   231
fun apply_cong_tac ctxt i = ("application of congruence", cong_tac ctxt i);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   232
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   233
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19350
diff changeset
   234
(* unfolds the definition of permutations     *)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19350
diff changeset
   235
(* applied to functions such that             *)
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   236
(*     pi o f = rhs                           *)
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19350
diff changeset
   237
(* is transformed to                          *)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   238
(*     %x. pi o (f ((rev pi) o x)) = rhs      *)
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60359
diff changeset
   239
fun unfold_perm_fun_def_tac ctxt i =
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   240
    ("unfolding of permutations on functions",
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60359
diff changeset
   241
      resolve_tac ctxt [perm_fun_def RS meta_eq_to_obj_eq RS trans] i)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   242
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19350
diff changeset
   243
(* applies the ext-rule such that      *)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19350
diff changeset
   244
(*                                     *)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   245
(*    f = g   goes to  /\x. f x = g x  *)
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60359
diff changeset
   246
fun ext_fun_tac ctxt i =
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60359
diff changeset
   247
  ("extensionality expansion of functions", resolve_tac ctxt @{thms ext} i);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   248
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   249
28322
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   250
(* perm_extend_simp_tac_i is perm_simp plus additional tactics        *)
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19350
diff changeset
   251
(* to decide equation that come from support problems             *)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19350
diff changeset
   252
(* since it contains looping rules the "recursion" - depth is set *)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19350
diff changeset
   253
(* to 10 - this seems to be sufficient in most cases              *)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   254
fun perm_extend_simp_tac_i tactical ctxt =
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   255
  let fun perm_extend_simp_tac_aux tactical ctxt n =
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32733
diff changeset
   256
          if n=0 then K all_tac
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   257
          else DETERM o
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60359
diff changeset
   258
               (FIRST'
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60359
diff changeset
   259
                 [fn i => tactical ctxt ("splitting conjunctions on the rhs", resolve_tac ctxt [conjI] i),
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60359
diff changeset
   260
                  fn i => tactical ctxt (perm_simp asm_full_simp_tac ctxt i),
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60359
diff changeset
   261
                  fn i => tactical ctxt (perm_compose_tac ctxt i),
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   262
                  fn i => tactical ctxt (apply_cong_tac ctxt i),
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60359
diff changeset
   263
                  fn i => tactical ctxt (unfold_perm_fun_def_tac ctxt i),
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60359
diff changeset
   264
                  fn i => tactical ctxt (ext_fun_tac ctxt i)]
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60359
diff changeset
   265
                THEN_ALL_NEW (TRY o (perm_extend_simp_tac_aux tactical ctxt (n-1))))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   266
  in perm_extend_simp_tac_aux tactical ctxt 10 end;
19151
66e841b1001e added a finite_guess tactic, which solves
urbanc
parents: 19144
diff changeset
   267
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   268
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   269
(* tactic that tries to solve "supports"-goals; first it *)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   270
(* unfolds the support definition and strips off the     *)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   271
(* intros, then applies eqvt_simp_tac                    *)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   272
fun supports_tac_i tactical ctxt i =
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   273
  let
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 35232
diff changeset
   274
     val simps        = [supports_def, Thm.symmetric fresh_def, fresh_prod]
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   275
  in
56491
a8ccf3d6a6e4 proper context for print_tac;
wenzelm
parents: 56253
diff changeset
   276
      EVERY [tactical ctxt ("unfolding of supports   ", simp_tac (put_simpset HOL_basic_ss ctxt addsimps simps) i),
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60359
diff changeset
   277
             tactical ctxt ("stripping of foralls    ", REPEAT_DETERM (resolve_tac ctxt [allI] i)),
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60359
diff changeset
   278
             tactical ctxt ("geting rid of the imps  ", resolve_tac ctxt [impI] i),
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60359
diff changeset
   279
             tactical ctxt ("eliminating conjuncts   ", REPEAT_DETERM (eresolve_tac ctxt [conjE] i)),
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60359
diff changeset
   280
             tactical ctxt ("applying eqvt_simp      ", eqvt_simp_tac_gen_i asm_full_simp_tac tactical ctxt i)]
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   281
  end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   282
19151
66e841b1001e added a finite_guess tactic, which solves
urbanc
parents: 19144
diff changeset
   283
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   284
(* tactic that guesses the finite-support of a goal        *)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   285
(* it first collects all free variables and tries to show  *)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   286
(* that the support of these free variables (op supports)  *)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   287
(* the goal                                                *)
20854
f9cf9e62d11c insert replacing ins ins_int ins_string
haftmann
parents: 20431
diff changeset
   288
fun collect_vars i (Bound j) vs = if j < i then vs else insert (op =) (Bound (j - i)) vs
f9cf9e62d11c insert replacing ins ins_int ins_string
haftmann
parents: 20431
diff changeset
   289
  | collect_vars i (v as Free _) vs = insert (op =) v vs
f9cf9e62d11c insert replacing ins ins_int ins_string
haftmann
parents: 20431
diff changeset
   290
  | collect_vars i (v as Var _) vs = insert (op =) v vs
19151
66e841b1001e added a finite_guess tactic, which solves
urbanc
parents: 19144
diff changeset
   291
  | collect_vars i (Const _) vs = vs
66e841b1001e added a finite_guess tactic, which solves
urbanc
parents: 19144
diff changeset
   292
  | collect_vars i (Abs (_, _, t)) vs = collect_vars (i+1) t vs
66e841b1001e added a finite_guess tactic, which solves
urbanc
parents: 19144
diff changeset
   293
  | collect_vars i (t $ u) vs = collect_vars i u (collect_vars i t vs);
66e841b1001e added a finite_guess tactic, which solves
urbanc
parents: 19144
diff changeset
   294
43278
1fbdcebb364b more robust exception pattern General.Subscript;
wenzelm
parents: 42364
diff changeset
   295
(* FIXME proper SUBGOAL/CSUBGOAL instead of cprems_of etc. *)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   296
fun finite_guess_tac_i tactical ctxt i st =
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 41489
diff changeset
   297
    let val goal = nth (cprems_of st) (i - 1)
19151
66e841b1001e added a finite_guess tactic, which solves
urbanc
parents: 19144
diff changeset
   298
    in
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58956
diff changeset
   299
      case Envir.eta_contract (Logic.strip_assums_concl (Thm.term_of goal)) of
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56230
diff changeset
   300
          _ $ (Const (@{const_name finite}, _) $ (Const (@{const_name Nominal.supp}, T) $ x)) =>
19151
66e841b1001e added a finite_guess tactic, which solves
urbanc
parents: 19144
diff changeset
   301
          let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58956
diff changeset
   302
            val ps = Logic.strip_params (Thm.term_of goal);
19151
66e841b1001e added a finite_guess tactic, which solves
urbanc
parents: 19144
diff changeset
   303
            val Ts = rev (map snd ps);
66e841b1001e added a finite_guess tactic, which solves
urbanc
parents: 19144
diff changeset
   304
            val vs = collect_vars 0 x [];
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 32960
diff changeset
   305
            val s = fold_rev (fn v => fn s =>
19151
66e841b1001e added a finite_guess tactic, which solves
urbanc
parents: 19144
diff changeset
   306
                HOLogic.pair_const (fastype_of1 (Ts, v)) (fastype_of1 (Ts, s)) $ v $ s)
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 32960
diff changeset
   307
              vs HOLogic.unit;
46219
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 45620
diff changeset
   308
            val s' = fold_rev Term.abs ps
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56230
diff changeset
   309
              (Const (@{const_name Nominal.supp}, fastype_of1 (Ts, s) -->
44692
ccfc7c193d2b modify nominal packages to better respect set/pred distinction
huffman
parents: 43278
diff changeset
   310
                Term.range_type T) $ s);
19151
66e841b1001e added a finite_guess tactic, which solves
urbanc
parents: 19144
diff changeset
   311
            val supports_rule' = Thm.lift_rule goal supports_rule;
66e841b1001e added a finite_guess tactic, which solves
urbanc
parents: 19144
diff changeset
   312
            val _ $ (_ $ S $ _) =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58956
diff changeset
   313
              Logic.strip_assums_concl (hd (Thm.prems_of supports_rule'));
60787
12cd198f07af updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
   314
            val supports_rule'' =
12cd198f07af updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
   315
              infer_instantiate ctxt
12cd198f07af updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
   316
                [(#1 (dest_Var (head_of S)), Thm.cterm_of ctxt s')] supports_rule';
60359
cb8828b859a1 clarified context;
wenzelm
parents: 59621
diff changeset
   317
            val fin_supp = Proof_Context.get_thms ctxt "fin_supp"
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   318
            val ctxt' = ctxt addsimps [supp_prod,supp_unit,finite_Un,finite_emptyI,conj_absorb]@fin_supp
19151
66e841b1001e added a finite_guess tactic, which solves
urbanc
parents: 19144
diff changeset
   319
          in
56491
a8ccf3d6a6e4 proper context for print_tac;
wenzelm
parents: 56253
diff changeset
   320
            (tactical ctxt ("guessing of the right supports-set",
58956
a816aa3ff391 proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
wenzelm
parents: 56492
diff changeset
   321
                      EVERY [compose_tac ctxt (false, supports_rule'', 2) i,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   322
                             asm_full_simp_tac ctxt' (i+1),
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   323
                             supports_tac_i tactical ctxt i])) st
19151
66e841b1001e added a finite_guess tactic, which solves
urbanc
parents: 19144
diff changeset
   324
          end
66e841b1001e added a finite_guess tactic, which solves
urbanc
parents: 19144
diff changeset
   325
        | _ => Seq.empty
66e841b1001e added a finite_guess tactic, which solves
urbanc
parents: 19144
diff changeset
   326
    end
43278
1fbdcebb364b more robust exception pattern General.Subscript;
wenzelm
parents: 42364
diff changeset
   327
    handle General.Subscript => Seq.empty
1fbdcebb364b more robust exception pattern General.Subscript;
wenzelm
parents: 42364
diff changeset
   328
(* FIXME proper SUBGOAL/CSUBGOAL instead of cprems_of etc. *)
19151
66e841b1001e added a finite_guess tactic, which solves
urbanc
parents: 19144
diff changeset
   329
22595
293934e41dfd make fresh_guess fail if it does not solve the subgoal
narboux
parents: 22578
diff changeset
   330
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   331
(* tactic that guesses whether an atom is fresh for an expression  *)
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   332
(* it first collects all free variables and tries to show that the *)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   333
(* support of these free variables (op supports) the goal          *)
43278
1fbdcebb364b more robust exception pattern General.Subscript;
wenzelm
parents: 42364
diff changeset
   334
(* FIXME proper SUBGOAL/CSUBGOAL instead of cprems_of etc. *)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   335
fun fresh_guess_tac_i tactical ctxt i st =
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   336
    let
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 41489
diff changeset
   337
        val goal = nth (cprems_of st) (i - 1)
60359
cb8828b859a1 clarified context;
wenzelm
parents: 59621
diff changeset
   338
        val fin_supp = Proof_Context.get_thms ctxt "fin_supp"
cb8828b859a1 clarified context;
wenzelm
parents: 59621
diff changeset
   339
        val fresh_atm = Proof_Context.get_thms ctxt "fresh_atm"
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   340
        val ctxt1 = ctxt addsimps [Thm.symmetric fresh_def,fresh_prod,fresh_unit,conj_absorb,not_false]@fresh_atm
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   341
        val ctxt2 = ctxt addsimps [supp_prod,supp_unit,finite_Un,finite_emptyI,conj_absorb]@fin_supp
19857
a0c36e0fc897 Added fresh_guess_tac.
berghofe
parents: 19494
diff changeset
   342
    in
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58956
diff changeset
   343
      case Logic.strip_assums_concl (Thm.term_of goal) of
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   344
          _ $ (Const (@{const_name Nominal.fresh}, Type ("fun", [T, _])) $ _ $ t) =>
19857
a0c36e0fc897 Added fresh_guess_tac.
berghofe
parents: 19494
diff changeset
   345
          let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58956
diff changeset
   346
            val ps = Logic.strip_params (Thm.term_of goal);
19857
a0c36e0fc897 Added fresh_guess_tac.
berghofe
parents: 19494
diff changeset
   347
            val Ts = rev (map snd ps);
a0c36e0fc897 Added fresh_guess_tac.
berghofe
parents: 19494
diff changeset
   348
            val vs = collect_vars 0 t [];
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 32960
diff changeset
   349
            val s = fold_rev (fn v => fn s =>
19857
a0c36e0fc897 Added fresh_guess_tac.
berghofe
parents: 19494
diff changeset
   350
                HOLogic.pair_const (fastype_of1 (Ts, v)) (fastype_of1 (Ts, s)) $ v $ s)
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 32960
diff changeset
   351
              vs HOLogic.unit;
46219
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 45620
diff changeset
   352
            val s' =
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 45620
diff changeset
   353
              fold_rev Term.abs ps
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56230
diff changeset
   354
                (Const (@{const_name Nominal.supp}, fastype_of1 (Ts, s) --> HOLogic.mk_setT T) $ s);
19857
a0c36e0fc897 Added fresh_guess_tac.
berghofe
parents: 19494
diff changeset
   355
            val supports_fresh_rule' = Thm.lift_rule goal supports_fresh_rule;
a0c36e0fc897 Added fresh_guess_tac.
berghofe
parents: 19494
diff changeset
   356
            val _ $ (_ $ S $ _) =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58956
diff changeset
   357
              Logic.strip_assums_concl (hd (Thm.prems_of supports_fresh_rule'));
60787
12cd198f07af updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
   358
            val supports_fresh_rule'' =
12cd198f07af updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
   359
              infer_instantiate ctxt
12cd198f07af updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
   360
                [(#1 (dest_Var (head_of S)), Thm.cterm_of ctxt s')] supports_fresh_rule';
19857
a0c36e0fc897 Added fresh_guess_tac.
berghofe
parents: 19494
diff changeset
   361
          in
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   362
            (tactical ctxt ("guessing of the right set that supports the goal",
58956
a816aa3ff391 proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
wenzelm
parents: 56492
diff changeset
   363
                      (EVERY [compose_tac ctxt (false, supports_fresh_rule'', 3) i,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   364
                             asm_full_simp_tac ctxt1 (i+2),
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   365
                             asm_full_simp_tac ctxt2 (i+1),
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   366
                             supports_tac_i tactical ctxt i]))) st
19857
a0c36e0fc897 Added fresh_guess_tac.
berghofe
parents: 19494
diff changeset
   367
          end
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   368
          (* when a term-constructor contains more than one binder, it is useful    *)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   369
          (* in nominal_primrecs to try whether the goal can be solved by an hammer *)
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   370
        | _ => (tactical ctxt ("if it is not of the form _\<sharp>_, then try the simplifier",
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   371
                          (asm_full_simp_tac (put_simpset HOL_ss ctxt addsimps [fresh_prod]@fresh_atm) i))) st
19857
a0c36e0fc897 Added fresh_guess_tac.
berghofe
parents: 19494
diff changeset
   372
    end
43278
1fbdcebb364b more robust exception pattern General.Subscript;
wenzelm
parents: 42364
diff changeset
   373
    handle General.Subscript => Seq.empty;
1fbdcebb364b more robust exception pattern General.Subscript;
wenzelm
parents: 42364
diff changeset
   374
(* FIXME proper SUBGOAL/CSUBGOAL instead of cprems_of etc. *)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   375
56492
29a1d14b944c more standard names;
wenzelm
parents: 56491
diff changeset
   376
val eqvt_simp_tac        = eqvt_asm_full_simp_tac_i NO_DEBUG;
28322
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   377
56492
29a1d14b944c more standard names;
wenzelm
parents: 56491
diff changeset
   378
val perm_simp_tac        = perm_asm_full_simp_tac_i NO_DEBUG;
29a1d14b944c more standard names;
wenzelm
parents: 56491
diff changeset
   379
val perm_extend_simp_tac = perm_extend_simp_tac_i NO_DEBUG;
29a1d14b944c more standard names;
wenzelm
parents: 56491
diff changeset
   380
val supports_tac         = supports_tac_i NO_DEBUG;
29a1d14b944c more standard names;
wenzelm
parents: 56491
diff changeset
   381
val finite_guess_tac     = finite_guess_tac_i NO_DEBUG;
29a1d14b944c more standard names;
wenzelm
parents: 56491
diff changeset
   382
val fresh_guess_tac      = fresh_guess_tac_i NO_DEBUG;
28322
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   383
56492
29a1d14b944c more standard names;
wenzelm
parents: 56491
diff changeset
   384
val dperm_simp_tac        = perm_asm_full_simp_tac_i DEBUG;
29a1d14b944c more standard names;
wenzelm
parents: 56491
diff changeset
   385
val dperm_extend_simp_tac = perm_extend_simp_tac_i DEBUG;
29a1d14b944c more standard names;
wenzelm
parents: 56491
diff changeset
   386
val dsupports_tac         = supports_tac_i DEBUG;
29a1d14b944c more standard names;
wenzelm
parents: 56491
diff changeset
   387
val dfinite_guess_tac     = finite_guess_tac_i DEBUG;
29a1d14b944c more standard names;
wenzelm
parents: 56491
diff changeset
   388
val dfresh_guess_tac      = fresh_guess_tac_i DEBUG;
28322
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   389
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   390
(* Code opied from the Simplifer for setting up the perm_simp method   *)
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   391
(* behaves nearly identical to the simp-method, for example can handle *)
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   392
(* options like (no_asm) etc.                                          *)
28322
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   393
val no_asmN = "no_asm";
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   394
val no_asm_useN = "no_asm_use";
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   395
val no_asm_simpN = "no_asm_simp";
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   396
val asm_lrN = "asm_lr";
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   397
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   398
val perm_simp_options =
56492
29a1d14b944c more standard names;
wenzelm
parents: 56491
diff changeset
   399
 (Args.parens (Args.$$$ no_asmN) >> K (perm_simp_tac_i NO_DEBUG) ||
29a1d14b944c more standard names;
wenzelm
parents: 56491
diff changeset
   400
  Args.parens (Args.$$$ no_asm_simpN) >> K (perm_asm_simp_tac_i NO_DEBUG) ||
29a1d14b944c more standard names;
wenzelm
parents: 56491
diff changeset
   401
  Args.parens (Args.$$$ no_asm_useN) >> K (perm_full_simp_tac_i NO_DEBUG) ||
29a1d14b944c more standard names;
wenzelm
parents: 56491
diff changeset
   402
  Args.parens (Args.$$$ asm_lrN) >> K (perm_asm_lr_simp_tac_i NO_DEBUG) ||
29a1d14b944c more standard names;
wenzelm
parents: 56491
diff changeset
   403
  Scan.succeed (perm_asm_full_simp_tac_i NO_DEBUG));
28322
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   404
30549
d2d7874648bd simplified method setup;
wenzelm
parents: 30510
diff changeset
   405
val perm_simp_meth =
33554
4601372337e4 eliminated some unused/obsolete Args.bang_facts;
wenzelm
parents: 33244
diff changeset
   406
  Scan.lift perm_simp_options --| Method.sections (Simplifier.simp_modifiers') >>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   407
  (fn tac => fn ctxt => SIMPLE_METHOD' (CHANGED_PROP o tac ctxt));
28322
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   408
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   409
(* setup so that the simpset is used which is active at the moment when the tactic is called *)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   410
fun local_simp_meth_setup tac =
30549
d2d7874648bd simplified method setup;
wenzelm
parents: 30510
diff changeset
   411
  Method.sections (Simplifier.simp_modifiers' @ Splitter.split_modifiers) >>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   412
  (K (SIMPLE_METHOD' o tac));
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   413
22595
293934e41dfd make fresh_guess fail if it does not solve the subgoal
narboux
parents: 22578
diff changeset
   414
(* uses HOL_basic_ss only and fails if the tactic does not solve the subgoal *)
293934e41dfd make fresh_guess fail if it does not solve the subgoal
narboux
parents: 22578
diff changeset
   415
22656
13302b2d0948 debug versions of finite_guess and fresh_guess do not fail if they can not solve the goal
narboux
parents: 22610
diff changeset
   416
fun basic_simp_meth_setup debug tac =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   417
  Scan.depend (fn context => Scan.succeed (Simplifier.map_ss (put_simpset HOL_basic_ss) context, ())) --
30549
d2d7874648bd simplified method setup;
wenzelm
parents: 30510
diff changeset
   418
  Method.sections (Simplifier.simp_modifiers' @ Splitter.split_modifiers) >>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46219
diff changeset
   419
  (K (SIMPLE_METHOD' o (if debug then tac else SOLVED' o tac)));
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   420
28322
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   421
val perm_simp_meth_debug        = local_simp_meth_setup dperm_simp_tac;
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   422
val perm_extend_simp_meth       = local_simp_meth_setup perm_extend_simp_tac;
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   423
val perm_extend_simp_meth_debug = local_simp_meth_setup dperm_extend_simp_tac;
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   424
val supports_meth               = local_simp_meth_setup supports_tac;
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   425
val supports_meth_debug         = local_simp_meth_setup dsupports_tac;
24571
a6d0428dea8e some cleaning up to do with contexts
urbanc
parents: 24519
diff changeset
   426
28322
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   427
val finite_guess_meth         = basic_simp_meth_setup false finite_guess_tac;
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   428
val finite_guess_meth_debug   = basic_simp_meth_setup true  dfinite_guess_tac;
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   429
val fresh_guess_meth          = basic_simp_meth_setup false fresh_guess_tac;
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28262
diff changeset
   430
val fresh_guess_meth_debug    = basic_simp_meth_setup true  dfresh_guess_tac;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   431
20289
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 19993
diff changeset
   432
end