src/HOL/Nominal/nominal_atoms.ML
author paulson <lp15@cam.ac.uk>
Thu, 18 Apr 2024 13:07:34 +0100
changeset 80133 e414bcc5a39e
parent 79336 032a31db4c6f
permissions -rw-r--r--
Acknowledgement of Ata Keskin for his Martingales material
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32952
diff changeset
     1
(*  Title:      HOL/Nominal/nominal_atoms.ML
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
     2
    Author:     Christian Urban and Stefan Berghofer, TU Muenchen
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
     3
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
     4
Declaration of atom types to be used in nominal datatypes.
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
     5
*)
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
     6
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
     7
signature NOMINAL_ATOMS =
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
     8
sig
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
     9
  val create_nom_typedecls : string list -> theory -> theory
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    10
  type atom_info
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    11
  val get_atom_infos : theory -> atom_info Symtab.table
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    12
  val get_atom_info : theory -> string -> atom_info option
28372
291e7a158e95 Added some more theorems to NominalData.
berghofe
parents: 28083
diff changeset
    13
  val the_atom_info : theory -> string -> atom_info
28729
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    14
  val fs_class_of : theory -> string -> string
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    15
  val pt_class_of : theory -> string -> string
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    16
  val cp_class_of : theory -> string -> string -> string
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    17
  val at_inst_of : theory -> string -> thm
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    18
  val pt_inst_of : theory -> string -> thm
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    19
  val cp_inst_of : theory -> string -> string -> thm
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    20
  val dj_thm_of : theory -> string -> string -> thm
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    21
  val atoms_of : theory -> string list
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    22
  val mk_permT : typ -> typ
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    23
end
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    24
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    25
structure NominalAtoms : NOMINAL_ATOMS =
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    26
struct
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    27
23894
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
    28
val finite_emptyI = @{thm "finite.emptyI"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
    29
val Collect_const = @{thm "Collect_const"};
21669
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21377
diff changeset
    30
59940
087d81f5213e local setup of induction tools, with restricted access to auxiliary consts;
wenzelm
parents: 59936
diff changeset
    31
val inductive_forall_def = @{thm HOL.induct_forall_def};
24569
c80e1871098b Added equivariance lemmas for induct_forall.
berghofe
parents: 24527
diff changeset
    32
21669
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21377
diff changeset
    33
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22794
diff changeset
    34
(* theory data *)
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    35
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    36
type atom_info =
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    37
  {pt_class : string,
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    38
   fs_class : string,
28729
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    39
   cp_classes : string Symtab.table,
28372
291e7a158e95 Added some more theorems to NominalData.
berghofe
parents: 28083
diff changeset
    40
   at_inst : thm,
291e7a158e95 Added some more theorems to NominalData.
berghofe
parents: 28083
diff changeset
    41
   pt_inst : thm,
28729
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    42
   cp_inst : thm Symtab.table,
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    43
   dj_thms : thm Symtab.table};
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    44
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 32960
diff changeset
    45
structure NominalData = Theory_Data
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22794
diff changeset
    46
(
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    47
  type T = atom_info Symtab.table;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    48
  val empty = Symtab.empty;
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 32960
diff changeset
    49
  fun merge data = Symtab.merge (K true) data;
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22794
diff changeset
    50
);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    51
28729
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    52
fun make_atom_info ((((((pt_class, fs_class), cp_classes), at_inst), pt_inst), cp_inst), dj_thms) =
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    53
  {pt_class = pt_class,
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    54
   fs_class = fs_class,
28372
291e7a158e95 Added some more theorems to NominalData.
berghofe
parents: 28083
diff changeset
    55
   cp_classes = cp_classes,
291e7a158e95 Added some more theorems to NominalData.
berghofe
parents: 28083
diff changeset
    56
   at_inst = at_inst,
291e7a158e95 Added some more theorems to NominalData.
berghofe
parents: 28083
diff changeset
    57
   pt_inst = pt_inst,
28729
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    58
   cp_inst = cp_inst,
28372
291e7a158e95 Added some more theorems to NominalData.
berghofe
parents: 28083
diff changeset
    59
   dj_thms = dj_thms};
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    60
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    61
val get_atom_infos = NominalData.get;
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    62
val get_atom_info = Symtab.lookup o NominalData.get;
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
    63
28729
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    64
fun gen_lookup lookup name = case lookup name of
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    65
    SOME info => info
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    66
  | NONE => error ("Unknown atom type " ^ quote name);
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    67
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    68
fun the_atom_info thy = gen_lookup (get_atom_info thy);
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    69
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    70
fun gen_lookup' f thy = the_atom_info thy #> f;
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    71
fun gen_lookup'' f thy =
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    72
  gen_lookup' (f #> Symtab.lookup #> gen_lookup) thy;
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    73
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    74
val fs_class_of = gen_lookup' #fs_class;
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    75
val pt_class_of = gen_lookup' #pt_class;
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    76
val at_inst_of = gen_lookup' #at_inst;
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    77
val pt_inst_of = gen_lookup' #pt_inst;
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    78
val cp_class_of = gen_lookup'' #cp_classes;
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    79
val cp_inst_of = gen_lookup'' #cp_inst;
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
    80
val dj_thm_of = gen_lookup'' #dj_thms;
28372
291e7a158e95 Added some more theorems to NominalData.
berghofe
parents: 28083
diff changeset
    81
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    82
fun atoms_of thy = map fst (Symtab.dest (NominalData.get thy));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    83
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
    84
fun mk_permT T = HOLogic.listT (HOLogic.mk_prodT (T, T));
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    85
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    86
fun mk_Cons x xs =
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    87
  let val T = fastype_of x
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
    88
  in Const (\<^const_name>\<open>Cons\<close>, T --> HOLogic.listT T --> HOLogic.listT T) $ x $ xs end;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    89
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
    90
fun add_thms_string args = Global_Theory.add_thms ((map o apfst o apfst) Binding.name args);
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
    91
fun add_thmss_string args = Global_Theory.add_thmss ((map o apfst o apfst) Binding.name args);
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
    92
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    93
(* this function sets up all matters related to atom-  *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    94
(* kinds; the user specifies a list of atom-kind names *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    95
(* atom_decl <ak1> ... <akn>                           *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    96
fun create_nom_typedecls ak_names thy =
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    97
  let
24527
888d56a8d9d3 modified proofs so that they are not using claset()
urbanc
parents: 24218
diff changeset
    98
    
24677
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
    99
    val (_,thy1) = 
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   100
    fold_map (fn ak => fn thy => 
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   101
          let val dt = ((Binding.name ak, [], NoSyn), [(Binding.name ak, [\<^typ>\<open>nat\<close>], NoSyn)])
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58300
diff changeset
   102
              val (dt_names, thy1) = BNF_LFP_Compat.add_datatype [BNF_LFP_Compat.Kill_Type_Args] [dt] thy;
31783
cfbe9609ceb1 add_datatypes does not yield particular rules any longer
haftmann
parents: 31781
diff changeset
   103
            
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58300
diff changeset
   104
              val injects = maps (#inject o BNF_LFP_Compat.the_info thy1 []) dt_names;
24677
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   105
              val ak_type = Type (Sign.intern_type thy1 ak,[])
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   106
              val ak_sign = Sign.intern_const thy1 ak 
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   107
              
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   108
              val inj_type = \<^typ>\<open>nat\<close> --> ak_type
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   109
              val inj_on_type = inj_type --> \<^typ>\<open>nat set\<close> --> \<^typ>\<open>bool\<close>
24677
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   110
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   111
              (* first statement *)
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   112
              val stmnt1 = HOLogic.mk_Trueprop 
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   113
                  (Const (\<^const_name>\<open>inj_on\<close>,inj_on_type) $ 
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   114
                         Const (ak_sign,inj_type) $ HOLogic.mk_UNIV \<^typ>\<open>nat\<close>)
24677
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   115
31783
cfbe9609ceb1 add_datatypes does not yield particular rules any longer
haftmann
parents: 31781
diff changeset
   116
              val simp1 = @{thm inj_on_def} :: injects;
24677
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   117
              
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   118
              fun proof1 ctxt = EVERY [simp_tac (put_simpset HOL_basic_ss ctxt addsimps simp1) 1,
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   119
                resolve_tac ctxt @{thms ballI} 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   120
                resolve_tac ctxt @{thms ballI} 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   121
                resolve_tac ctxt @{thms impI} 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   122
                assume_tac ctxt 1]
24677
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   123
             
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   124
              val (inj_thm,thy2) = 
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   125
                   add_thms_string [((ak^"_inj",Goal.prove_global thy1 [] [] stmnt1 (proof1 o #context)), [])] thy1
24677
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   126
              
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   127
              (* second statement *)
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   128
              val y = Free ("y",ak_type)  
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   129
              val stmnt2 = HOLogic.mk_Trueprop
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   130
                  (HOLogic.mk_exists ("x",\<^typ>\<open>nat\<close>,HOLogic.mk_eq (y,Const (ak_sign,inj_type) $ Bound 0)))
24677
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   131
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   132
              val proof2 = fn {prems, context = ctxt} =>
59826
442b09c0f898 tuned signature;
wenzelm
parents: 59802
diff changeset
   133
                Induct_Tacs.case_tac ctxt "y" [] NONE 1 THEN
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   134
                asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps simp1) 1 THEN
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   135
                resolve_tac ctxt @{thms exI} 1 THEN
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   136
                resolve_tac ctxt @{thms refl} 1
24677
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   137
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   138
              (* third statement *)
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   139
              val (inject_thm,thy3) =
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   140
                  add_thms_string [((ak^"_injection",Goal.prove_global thy2 [] [] stmnt2 proof2), [])] thy2
24677
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   141
  
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   142
              val stmnt3 = HOLogic.mk_Trueprop
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   143
                           (HOLogic.mk_not
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   144
                              (Const (\<^const_name>\<open>finite\<close>, HOLogic.mk_setT ak_type --> HOLogic.boolT) $
24677
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   145
                                  HOLogic.mk_UNIV ak_type))
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   146
             
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   147
              val simp2 = [@{thm image_def},@{thm bex_UNIV}]@inject_thm
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   148
              val simp3 = [@{thm UNIV_def}]
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   149
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   150
              fun proof3 ctxt = EVERY [cut_facts_tac inj_thm 1,
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   151
                dresolve_tac ctxt @{thms range_inj_infinite} 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   152
                asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps simp2) 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   153
                simp_tac (put_simpset HOL_basic_ss ctxt addsimps simp3) 1]
24677
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   154
           
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   155
              val (inf_thm,thy4) =  
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   156
                    add_thms_string [((ak^"_infinite",Goal.prove_global thy3 [] [] stmnt3 (proof3 o #context)), [])] thy3
24677
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   157
          in 
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   158
            ((inj_thm,inject_thm,inf_thm),thy4)
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   159
          end) ak_names thy
c6295d2dce48 changed the representation of atoms to datatypes over nats
urbanc
parents: 24569
diff changeset
   160
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   161
    (* produces a list consisting of pairs:         *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   162
    (*  fst component is the atom-kind name         *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   163
    (*  snd component is its type                   *)
21289
920b7b893d9c deleted all uses of sign_of as it's now the identity function
urbanc
parents: 20377
diff changeset
   164
    val full_ak_names = map (Sign.intern_type thy1) ak_names;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   165
    val ak_names_types = ak_names ~~ map (Type o rpair []) full_ak_names;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   166
     
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   167
    (* declares a swapping function for every atom-kind, it is         *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   168
    (* const swap_<ak> :: <akT> * <akT> => <akT> => <akT>              *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   169
    (* swap_<ak> (a,b) c = (if a=c then b (else if b=c then a else c)) *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   170
    (* overloades then the general swap-function                       *) 
20179
a2f3f523c84b adaption to argument change in primrec_package
haftmann
parents: 20097
diff changeset
   171
    val (swap_eqs, thy3) = fold_map (fn (ak_name, T) => fn thy =>
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   172
      let
41562
90fb3d7474df Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents: 41163
diff changeset
   173
        val thy' = Sign.add_path "rec" thy;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   174
        val swapT = HOLogic.mk_prodT (T, T) --> T --> T;
41562
90fb3d7474df Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents: 41163
diff changeset
   175
        val swap_name = "swap_" ^ ak_name;
90fb3d7474df Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents: 41163
diff changeset
   176
        val full_swap_name = Sign.full_bname thy' swap_name;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   177
        val a = Free ("a", T);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   178
        val b = Free ("b", T);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   179
        val c = Free ("c", T);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   180
        val ab = Free ("ab", HOLogic.mk_prodT (T, T))
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   181
        val cif = Const (\<^const_name>\<open>If\<close>, HOLogic.boolT --> T --> T --> T);
41562
90fb3d7474df Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents: 41163
diff changeset
   182
        val cswap_akname = Const (full_swap_name, swapT);
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   183
        val cswap = Const (\<^const_name>\<open>Nominal.swap\<close>, swapT)
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   184
41562
90fb3d7474df Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents: 41163
diff changeset
   185
        val name = swap_name ^ "_def";
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   186
        val def1 = HOLogic.mk_Trueprop (HOLogic.mk_eq
41562
90fb3d7474df Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents: 41163
diff changeset
   187
                (Free (swap_name, swapT) $ HOLogic.mk_prod (a,b) $ c,
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   188
                    cif $ HOLogic.mk_eq (a,c) $ b $ (cif $ HOLogic.mk_eq (b,c) $ a $ c)))
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   189
        val def2 = Logic.mk_equals (cswap $ ab $ c, cswap_akname $ ab $ c)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   190
      in
41562
90fb3d7474df Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents: 41163
diff changeset
   191
        thy' |>
60003
ba8fa0c38d66 renamed ML funs
blanchet
parents: 59940
diff changeset
   192
        BNF_LFP_Compat.primrec_global
41562
90fb3d7474df Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents: 41163
diff changeset
   193
          [(Binding.name swap_name, SOME swapT, NoSyn)]
63352
4eaf35781b23 tuned signature;
wenzelm
parents: 63182
diff changeset
   194
          [((Binding.empty_atts, def1), [], [])] ||>
41562
90fb3d7474df Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents: 41163
diff changeset
   195
        Sign.parent_path ||>>
79336
032a31db4c6f clarified signature: downgrade old-style Global_Theory.add_defs to Global_Theory.add_def without attributes;
wenzelm
parents: 74561
diff changeset
   196
        fold_map Global_Theory.add_def_unchecked_overloaded [(Binding.name name, def2)] |>> (snd o fst)
26398
fccb74555d9e removed redundant axiomatizations of XXX_infinite (fact already proven);
wenzelm
parents: 26343
diff changeset
   197
      end) ak_names_types thy1;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   198
    
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   199
    (* declares a permutation function for every atom-kind acting  *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   200
    (* on such atoms                                               *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   201
    (* const <ak>_prm_<ak> :: (<akT> * <akT>)list => akT => akT    *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   202
    (* <ak>_prm_<ak> []     a = a                                  *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   203
    (* <ak>_prm_<ak> (x#xs) a = swap_<ak> x (perm xs a)            *)
20179
a2f3f523c84b adaption to argument change in primrec_package
haftmann
parents: 20097
diff changeset
   204
    val (prm_eqs, thy4) = fold_map (fn (ak_name, T) => fn thy =>
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   205
      let
41562
90fb3d7474df Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents: 41163
diff changeset
   206
        val thy' = Sign.add_path "rec" thy;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   207
        val swapT = HOLogic.mk_prodT (T, T) --> T --> T;
41562
90fb3d7474df Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents: 41163
diff changeset
   208
        val swap_name = Sign.full_bname thy' ("swap_" ^ ak_name)
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   209
        val prmT = mk_permT T --> T --> T;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   210
        val prm_name = ak_name ^ "_prm_" ^ ak_name;
41562
90fb3d7474df Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents: 41163
diff changeset
   211
        val prm = Free (prm_name, prmT);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   212
        val x  = Free ("x", HOLogic.mk_prodT (T, T));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   213
        val xs = Free ("xs", mk_permT T);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   214
        val a  = Free ("a", T) ;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   215
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   216
        val cnil  = Const (\<^const_name>\<open>Nil\<close>, mk_permT T);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   217
        
41562
90fb3d7474df Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents: 41163
diff changeset
   218
        val def1 = HOLogic.mk_Trueprop (HOLogic.mk_eq (prm $ cnil $ a, a));
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   219
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   220
        val def2 = HOLogic.mk_Trueprop (HOLogic.mk_eq
41562
90fb3d7474df Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents: 41163
diff changeset
   221
                   (prm $ mk_Cons x xs $ a,
90fb3d7474df Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents: 41163
diff changeset
   222
                    Const (swap_name, swapT) $ x $ (prm $ xs $ a)));
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   223
      in
41562
90fb3d7474df Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents: 41163
diff changeset
   224
        thy' |>
60003
ba8fa0c38d66 renamed ML funs
blanchet
parents: 59940
diff changeset
   225
        BNF_LFP_Compat.primrec_global
41562
90fb3d7474df Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents: 41163
diff changeset
   226
          [(Binding.name prm_name, SOME prmT, NoSyn)]
63352
4eaf35781b23 tuned signature;
wenzelm
parents: 63182
diff changeset
   227
          (map (fn def => ((Binding.empty_atts, def), [], [])) [def1, def2]) ||>
41562
90fb3d7474df Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents: 41163
diff changeset
   228
        Sign.parent_path
20179
a2f3f523c84b adaption to argument change in primrec_package
haftmann
parents: 20097
diff changeset
   229
      end) ak_names_types thy3;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   230
    
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   231
    (* defines permutation functions for all combinations of atom-kinds; *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   232
    (* there are a trivial cases and non-trivial cases                   *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   233
    (* non-trivial case:                                                 *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   234
    (* <ak>_prm_<ak>_def:  perm pi a == <ak>_prm_<ak> pi a               *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   235
    (* trivial case with <ak> != <ak'>                                   *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   236
    (* <ak>_prm<ak'>_def[simp]:  perm pi a == a                          *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   237
    (*                                                                   *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   238
    (* the trivial cases are added to the simplifier, while the non-     *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   239
    (* have their own rules proved below                                 *)  
18366
78b4f225b640 Adapted to new type of PureThy.add_defs_i.
berghofe
parents: 18355
diff changeset
   240
    val (perm_defs, thy5) = fold_map (fn (ak_name, T) => fn thy =>
78b4f225b640 Adapted to new type of PureThy.add_defs_i.
berghofe
parents: 18355
diff changeset
   241
      fold_map (fn (ak_name', T') => fn thy' =>
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   242
        let
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   243
          val perm_def_name = ak_name ^ "_prm_" ^ ak_name';
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   244
          val pi = Free ("pi", mk_permT T);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   245
          val a  = Free ("a", T');
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   246
          val cperm = Const (\<^const_name>\<open>Nominal.perm\<close>, mk_permT T --> T' --> T');
41562
90fb3d7474df Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents: 41163
diff changeset
   247
          val thy'' = Sign.add_path "rec" thy'
90fb3d7474df Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents: 41163
diff changeset
   248
          val cperm_def = Const (Sign.full_bname thy'' perm_def_name, mk_permT T --> T' --> T');
90fb3d7474df Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents: 41163
diff changeset
   249
          val thy''' = Sign.parent_path thy'';
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   250
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   251
          val name = ak_name ^ "_prm_" ^ ak_name' ^ "_def";
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   252
          val def = Logic.mk_equals
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   253
                    (cperm $ pi $ a, if ak_name = ak_name' then cperm_def $ pi $ a else a)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   254
        in
79336
032a31db4c6f clarified signature: downgrade old-style Global_Theory.add_defs to Global_Theory.add_def without attributes;
wenzelm
parents: 74561
diff changeset
   255
          Global_Theory.add_def_unchecked_overloaded (Binding.name name, def) thy'''
18366
78b4f225b640 Adapted to new type of PureThy.add_defs_i.
berghofe
parents: 18355
diff changeset
   256
        end) ak_names_types thy) ak_names_types thy4;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   257
    
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   258
    (* proves that every atom-kind is an instance of at *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   259
    (* lemma at_<ak>_inst:                              *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   260
    (* at TYPE(<ak>)                                    *)
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   261
    val (prm_cons_thms,thy6) = 
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   262
      thy5 |> add_thms_string (map (fn (ak_name, T) =>
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   263
      let
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   264
        val ak_name_qu = Sign.full_bname thy5 (ak_name);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   265
        val i_type = Type(ak_name_qu,[]);
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   266
        val cat = Const (\<^const_name>\<open>Nominal.at\<close>, Term.itselfT i_type --> HOLogic.boolT);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   267
        val at_type = Logic.mk_type i_type;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   268
        fun proof ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   269
          simp_tac (put_simpset HOL_ss ctxt
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   270
            addsimps maps (Global_Theory.get_thms thy5)
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   271
                                  ["at_def",
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   272
                                   ak_name ^ "_prm_" ^ ak_name ^ "_def",
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   273
                                   ak_name ^ "_prm_" ^ ak_name ^ ".simps",
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   274
                                   "swap_" ^ ak_name ^ "_def",
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   275
                                   "swap_" ^ ak_name ^ ".simps",
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   276
                                   ak_name ^ "_infinite"]) 1;            
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   277
        val name = "at_"^ak_name^ "_inst";
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   278
        val statement = HOLogic.mk_Trueprop (cat $ at_type);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   279
      in 
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   280
        ((name, Goal.prove_global thy5 [] [] statement (proof o #context)), [])
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   281
      end) ak_names_types);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   282
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   283
    (* declares a perm-axclass for every atom-kind               *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   284
    (* axclass pt_<ak>                                           *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   285
    (* pt_<ak>1[simp]: perm [] x = x                             *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   286
    (* pt_<ak>2:       perm (pi1@pi2) x = perm pi1 (perm pi2 x)  *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   287
    (* pt_<ak>3:       pi1 ~ pi2 ==> perm pi1 x = perm pi2 x     *)
18438
b8d867177916 made the changes according to Florian's re-arranging of
urbanc
parents: 18437
diff changeset
   288
     val (pt_ax_classes,thy7) =  fold_map (fn (ak_name, T) => fn thy =>
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   289
      let 
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   290
          val cl_name = "pt_"^ak_name;
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   291
          val ty = TFree("'a", \<^sort>\<open>type\<close>);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   292
          val x   = Free ("x", ty);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   293
          val pi1 = Free ("pi1", mk_permT T);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   294
          val pi2 = Free ("pi2", mk_permT T);
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   295
          val cperm = Const (\<^const_name>\<open>Nominal.perm\<close>, mk_permT T --> ty --> ty);
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   296
          val cnil  = Const (\<^const_name>\<open>Nil\<close>, mk_permT T);
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   297
          val cappend = Const (\<^const_name>\<open>append\<close>, mk_permT T --> mk_permT T --> mk_permT T);
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   298
          val cprm_eq = Const (\<^const_name>\<open>Nominal.prm_eq\<close>, mk_permT T --> mk_permT T --> HOLogic.boolT);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   299
          (* nil axiom *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   300
          val axiom1 = HOLogic.mk_Trueprop (HOLogic.mk_eq 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   301
                       (cperm $ cnil $ x, x));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   302
          (* append axiom *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   303
          val axiom2 = HOLogic.mk_Trueprop (HOLogic.mk_eq
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   304
                       (cperm $ (cappend $ pi1 $ pi2) $ x, cperm $ pi1 $ (cperm $ pi2 $ x)));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   305
          (* perm-eq axiom *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   306
          val axiom3 = Logic.mk_implies
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   307
                       (HOLogic.mk_Trueprop (cprm_eq $ pi1 $ pi2),
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   308
                        HOLogic.mk_Trueprop (HOLogic.mk_eq (cperm $ pi1 $ x, cperm $ pi2 $ x)));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   309
      in
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   310
          Axclass.define_class (Binding.name cl_name, \<^sort>\<open>type\<close>) []
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   311
                [((Binding.name (cl_name ^ "1"), [Simplifier.simp_add]), [axiom1]),
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   312
                 ((Binding.name (cl_name ^ "2"), []), [axiom2]),                           
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   313
                 ((Binding.name (cl_name ^ "3"), []), [axiom3])] thy
18438
b8d867177916 made the changes according to Florian's re-arranging of
urbanc
parents: 18437
diff changeset
   314
      end) ak_names_types thy6;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   315
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   316
    (* proves that every pt_<ak>-type together with <ak>-type *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   317
    (* instance of pt                                         *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   318
    (* lemma pt_<ak>_inst:                                    *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   319
    (* pt TYPE('x::pt_<ak>) TYPE(<ak>)                        *)
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   320
    val (prm_inst_thms,thy8) = 
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   321
      thy7 |> add_thms_string (map (fn (ak_name, T) =>
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   322
      let
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   323
        val ak_name_qu = Sign.full_bname thy7 ak_name;
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   324
        val pt_name_qu = Sign.full_bname thy7 ("pt_"^ak_name);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   325
        val i_type1 = TFree("'x",[pt_name_qu]);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   326
        val i_type2 = Type(ak_name_qu,[]);
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 54742
diff changeset
   327
        val cpt =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   328
          Const (\<^const_name>\<open>Nominal.pt\<close>, (Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   329
        val pt_type = Logic.mk_type i_type1;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   330
        val at_type = Logic.mk_type i_type2;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   331
        fun proof ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   332
          simp_tac (put_simpset HOL_ss ctxt addsimps maps (Global_Theory.get_thms thy7)
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   333
                                  ["pt_def",
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   334
                                   "pt_" ^ ak_name ^ "1",
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   335
                                   "pt_" ^ ak_name ^ "2",
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   336
                                   "pt_" ^ ak_name ^ "3"]) 1;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   337
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   338
        val name = "pt_"^ak_name^ "_inst";
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   339
        val statement = HOLogic.mk_Trueprop (cpt $ pt_type $ at_type);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   340
      in 
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   341
        ((name, Goal.prove_global thy7 [] [] statement (proof o #context)), []) 
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   342
      end) ak_names_types);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   343
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   344
     (* declares an fs-axclass for every atom-kind       *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   345
     (* axclass fs_<ak>                                  *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   346
     (* fs_<ak>1: finite ((supp x)::<ak> set)            *)
18438
b8d867177916 made the changes according to Florian's re-arranging of
urbanc
parents: 18437
diff changeset
   347
     val (fs_ax_classes,thy11) =  fold_map (fn (ak_name, T) => fn thy =>
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   348
       let 
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   349
          val cl_name = "fs_"^ak_name;
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   350
          val pt_name = Sign.full_bname thy ("pt_"^ak_name);
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   351
          val ty = TFree("'a",\<^sort>\<open>type\<close>);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   352
          val x   = Free ("x", ty);
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   353
          val csupp    = Const (\<^const_name>\<open>Nominal.supp\<close>, ty --> HOLogic.mk_setT T);
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   354
          val cfinite  = Const (\<^const_name>\<open>finite\<close>, HOLogic.mk_setT T --> HOLogic.boolT)
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   355
          
22274
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 21669
diff changeset
   356
          val axiom1   = HOLogic.mk_Trueprop (cfinite $ (csupp $ x));
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   357
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   358
       in  
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 51671
diff changeset
   359
        Axclass.define_class (Binding.name cl_name, [pt_name]) []
30345
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30235
diff changeset
   360
          [((Binding.name (cl_name ^ "1"), []), [axiom1])] thy
18438
b8d867177916 made the changes according to Florian's re-arranging of
urbanc
parents: 18437
diff changeset
   361
       end) ak_names_types thy8; 
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   362
         
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   363
     (* proves that every fs_<ak>-type together with <ak>-type   *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   364
     (* instance of fs-type                                      *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   365
     (* lemma abst_<ak>_inst:                                    *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   366
     (* fs TYPE('x::pt_<ak>) TYPE (<ak>)                         *)
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   367
     val (fs_inst_thms,thy12) = 
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   368
       thy11 |> add_thms_string (map (fn (ak_name, T) =>
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   369
       let
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   370
         val ak_name_qu = Sign.full_bname thy11 ak_name;
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   371
         val fs_name_qu = Sign.full_bname thy11 ("fs_"^ak_name);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   372
         val i_type1 = TFree("'x",[fs_name_qu]);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   373
         val i_type2 = Type(ak_name_qu,[]);
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   374
         val cfs = Const (\<^const_name>\<open>Nominal.fs\<close>,
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   375
                                 (Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   376
         val fs_type = Logic.mk_type i_type1;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   377
         val at_type = Logic.mk_type i_type2;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   378
         fun proof ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   379
          simp_tac (put_simpset HOL_ss ctxt addsimps maps (Global_Theory.get_thms thy11)
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   380
                                   ["fs_def",
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   381
                                    "fs_" ^ ak_name ^ "1"]) 1;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   382
    
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   383
         val name = "fs_"^ak_name^ "_inst";
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   384
         val statement = HOLogic.mk_Trueprop (cfs $ fs_type $ at_type);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   385
       in 
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   386
         ((name, Goal.prove_global thy11 [] [] statement (proof o #context)), []) 
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   387
       end) ak_names_types);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   388
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   389
       (* declares for every atom-kind combination an axclass            *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   390
       (* cp_<ak1>_<ak2> giving a composition property                   *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   391
       (* cp_<ak1>_<ak2>1: pi1 o pi2 o x = (pi1 o pi2) o (pi1 o x)       *)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   392
        val (cp_ax_classes,thy12b) = fold_map (fn (ak_name, T) => fn thy =>
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   393
         fold_map (fn (ak_name', T') => fn thy' =>
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   394
             let
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   395
               val cl_name = "cp_"^ak_name^"_"^ak_name';
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   396
               val ty = TFree("'a",\<^sort>\<open>type\<close>);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   397
               val x   = Free ("x", ty);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   398
               val pi1 = Free ("pi1", mk_permT T);
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   399
               val pi2 = Free ("pi2", mk_permT T');                  
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   400
               val cperm1 = Const (\<^const_name>\<open>Nominal.perm\<close>, mk_permT T  --> ty --> ty);
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   401
               val cperm2 = Const (\<^const_name>\<open>Nominal.perm\<close>, mk_permT T' --> ty --> ty);
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   402
               val cperm3 = Const (\<^const_name>\<open>Nominal.perm\<close>, mk_permT T  --> mk_permT T' --> mk_permT T');
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   403
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   404
               val ax1   = HOLogic.mk_Trueprop 
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   405
                           (HOLogic.mk_eq (cperm1 $ pi1 $ (cperm2 $ pi2 $ x), 
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   406
                                           cperm2 $ (cperm3 $ pi1 $ pi2) $ (cperm1 $ pi1 $ x)));
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   407
               in  
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   408
                 Axclass.define_class (Binding.name cl_name, \<^sort>\<open>type\<close>) []
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   409
                   [((Binding.name (cl_name ^ "1"), []), [ax1])] thy'  
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   410
               end) ak_names_types thy) ak_names_types thy12;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   411
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   412
        (* proves for every <ak>-combination a cp_<ak1>_<ak2>_inst theorem;     *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   413
        (* lemma cp_<ak1>_<ak2>_inst:                                           *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   414
        (* cp TYPE('a::cp_<ak1>_<ak2>) TYPE(<ak1>) TYPE(<ak2>)                  *)
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   415
        val (cp_thms,thy12c) = fold_map (fn (ak_name, T) => fn thy =>
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   416
         fold_map (fn (ak_name', T') => fn thy' =>
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   417
           let
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   418
             val ak_name_qu  = Sign.full_bname thy' (ak_name);
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   419
             val ak_name_qu' = Sign.full_bname thy' (ak_name');
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   420
             val cp_name_qu  = Sign.full_bname thy' ("cp_"^ak_name^"_"^ak_name');
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   421
             val i_type0 = TFree("'a",[cp_name_qu]);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   422
             val i_type1 = Type(ak_name_qu,[]);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   423
             val i_type2 = Type(ak_name_qu',[]);
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   424
             val ccp = Const (\<^const_name>\<open>Nominal.cp\<close>,
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   425
                             (Term.itselfT i_type0)-->(Term.itselfT i_type1)-->
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   426
                                                      (Term.itselfT i_type2)-->HOLogic.boolT);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   427
             val at_type  = Logic.mk_type i_type1;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   428
             val at_type' = Logic.mk_type i_type2;
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   429
             val cp_type  = Logic.mk_type i_type0;
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   430
             val cp1      = Global_Theory.get_thm thy' ("cp_" ^ ak_name ^ "_" ^ ak_name' ^ "1");
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   431
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   432
             val name = "cp_"^ak_name^ "_"^ak_name'^"_inst";
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   433
             val statement = HOLogic.mk_Trueprop (ccp $ cp_type $ at_type $ at_type');
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   434
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   435
             fun proof ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   436
              simp_tac (put_simpset HOL_basic_ss ctxt
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   437
                  addsimps maps (Global_Theory.get_thms thy') ["cp_def"]) 1
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   438
                THEN EVERY [resolve_tac ctxt [allI] 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   439
                  resolve_tac ctxt [allI] 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   440
                  resolve_tac ctxt [allI] 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   441
                  resolve_tac ctxt [cp1] 1];
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   442
           in
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   443
             yield_singleton add_thms_string ((name,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   444
               Goal.prove_global thy' [] [] statement (proof o #context)), []) thy'
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   445
           end) 
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   446
           ak_names_types thy) ak_names_types thy12b;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   447
       
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   448
        (* proves for every non-trivial <ak>-combination a disjointness   *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   449
        (* theorem; i.e. <ak1> != <ak2>                                   *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   450
        (* lemma ds_<ak1>_<ak2>:                                          *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   451
        (* dj TYPE(<ak1>) TYPE(<ak2>)                                     *)
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   452
        val (dj_thms, thy12d) = fold_map (fn (ak_name,T) => fn thy =>
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   453
          fold_map (fn (ak_name',T') => fn thy' =>
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   454
          (if not (ak_name = ak_name') 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   455
           then 
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   456
               let
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   457
                 val ak_name_qu  = Sign.full_bname thy' ak_name;
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   458
                 val ak_name_qu' = Sign.full_bname thy' ak_name';
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   459
                 val i_type1 = Type(ak_name_qu,[]);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   460
                 val i_type2 = Type(ak_name_qu',[]);
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   461
                 val cdj = Const (\<^const_name>\<open>Nominal.disjoint\<close>,
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   462
                           (Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   463
                 val at_type  = Logic.mk_type i_type1;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   464
                 val at_type' = Logic.mk_type i_type2;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   465
                 fun proof ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   466
                  simp_tac (put_simpset HOL_ss ctxt
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   467
                    addsimps maps (Global_Theory.get_thms thy')
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   468
                                           ["disjoint_def",
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   469
                                            ak_name ^ "_prm_" ^ ak_name' ^ "_def",
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   470
                                            ak_name' ^ "_prm_" ^ ak_name ^ "_def"]) 1;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   471
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   472
                 val name = "dj_"^ak_name^"_"^ak_name';
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   473
                 val statement = HOLogic.mk_Trueprop (cdj $ at_type $ at_type');
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   474
               in
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   475
                add_thms_string [((name, Goal.prove_global thy' [] [] statement (proof o #context)), [])] thy'
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   476
               end
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   477
           else 
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   478
            ([],thy')))  (* do nothing branch, if ak_name = ak_name' *) 
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   479
            ak_names_types thy) ak_names_types thy12c;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   480
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   481
     (********  pt_<ak> class instances  ********)
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   482
     (*=========================================*)
18279
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   483
     (* some abbreviations for theorems *)
23894
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   484
      val pt1           = @{thm "pt1"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   485
      val pt2           = @{thm "pt2"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   486
      val pt3           = @{thm "pt3"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   487
      val at_pt_inst    = @{thm "at_pt_inst"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   488
      val pt_unit_inst  = @{thm "pt_unit_inst"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   489
      val pt_prod_inst  = @{thm "pt_prod_inst"}; 
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   490
      val pt_nprod_inst = @{thm "pt_nprod_inst"}; 
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   491
      val pt_list_inst  = @{thm "pt_list_inst"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   492
      val pt_optn_inst  = @{thm "pt_option_inst"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   493
      val pt_noptn_inst = @{thm "pt_noption_inst"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   494
      val pt_fun_inst   = @{thm "pt_fun_inst"};
45979
296d9a9c8d24 treatment of type constructor `set`
haftmann
parents: 45839
diff changeset
   495
      val pt_set_inst   = @{thm "pt_set_inst"};
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   496
18435
318d2c271040 tuned one comment
urbanc
parents: 18432
diff changeset
   497
     (* for all atom-kind combinations <ak>/<ak'> show that        *)
318d2c271040 tuned one comment
urbanc
parents: 18432
diff changeset
   498
     (* every <ak> is an instance of pt_<ak'>; the proof for       *)
318d2c271040 tuned one comment
urbanc
parents: 18432
diff changeset
   499
     (* ak!=ak' is by definition; the case ak=ak' uses at_pt_inst. *)
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   500
     val thy13 = fold (fn ak_name => fn thy =>
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   501
        fold (fn ak_name' => fn thy' =>
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   502
         let
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   503
           val qu_name =  Sign.full_bname thy' ak_name';
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   504
           val cls_name = Sign.full_bname thy' ("pt_"^ak_name);
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   505
           val at_inst  = Global_Theory.get_thm thy' ("at_" ^ ak_name' ^ "_inst");
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   506
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58354
diff changeset
   507
           fun proof1 ctxt = EVERY [Class.intro_classes_tac ctxt [],
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   508
             resolve_tac ctxt [(at_inst RS at_pt_inst) RS pt1] 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   509
             resolve_tac ctxt [(at_inst RS at_pt_inst) RS pt2] 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   510
             resolve_tac ctxt [(at_inst RS at_pt_inst) RS pt3] 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   511
             assume_tac ctxt 1];
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   512
           fun proof2 ctxt =
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58354
diff changeset
   513
             Class.intro_classes_tac ctxt [] THEN
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   514
             REPEAT (asm_simp_tac
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   515
              (put_simpset HOL_basic_ss ctxt addsimps
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   516
                maps (Global_Theory.get_thms thy') [ak_name ^ "_prm_" ^ ak_name' ^ "_def"]) 1);
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   517
         in
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   518
           thy'
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 51671
diff changeset
   519
           |> Axclass.prove_arity (qu_name,[],[cls_name])
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58354
diff changeset
   520
              (fn ctxt => if ak_name = ak_name' then proof1 ctxt else proof2 ctxt)
26484
28dd7c7a9a2e eliminated non-linear access to thy1 and thy12c;
wenzelm
parents: 26398
diff changeset
   521
         end) ak_names thy) ak_names thy12d;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   522
18430
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   523
     (* show that                       *)
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   524
     (*      fun(pt_<ak>,pt_<ak>)       *)
18579
002d371401f5 changed the name of the type "nOption" to "noption".
urbanc
parents: 18438
diff changeset
   525
     (*      noption(pt_<ak>)           *)
18430
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   526
     (*      option(pt_<ak>)            *)
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   527
     (*      list(pt_<ak>)              *)
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   528
     (*      *(pt_<ak>,pt_<ak>)         *)
18600
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   529
     (*      nprod(pt_<ak>,pt_<ak>)     *)
18430
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   530
     (*      unit                       *)
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   531
     (*      set(pt_<ak>)               *)
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   532
     (* are instances of pt_<ak>        *)
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   533
     val thy18 = fold (fn ak_name => fn thy =>
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   534
       let
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   535
          val cls_name = Sign.full_bname thy ("pt_"^ak_name);
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   536
          val at_thm   = Global_Theory.get_thm thy ("at_"^ak_name^"_inst");
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   537
          val pt_inst  = Global_Theory.get_thm thy ("pt_"^ak_name^"_inst");
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   538
51671
0d142a78fb7c formal proof context for axclass proofs;
wenzelm
parents: 46961
diff changeset
   539
          fun pt_proof thm ctxt =
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58354
diff changeset
   540
              EVERY [Class.intro_classes_tac ctxt [],
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   541
                resolve_tac ctxt [thm RS pt1] 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   542
                resolve_tac ctxt [thm RS pt2] 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   543
                resolve_tac ctxt [thm RS pt3] 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   544
                assume_tac ctxt 1];
18430
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   545
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   546
          val pt_thm_fun   = at_thm RS (pt_inst RS (pt_inst RS pt_fun_inst));
46180
72ee700e1d8f Corrected pt_set_inst, added missing cp_set_inst, deleted obsolete
berghofe
parents: 45979
diff changeset
   547
          val pt_thm_set   = pt_inst RS pt_set_inst;
18430
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   548
          val pt_thm_noptn = pt_inst RS pt_noptn_inst; 
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   549
          val pt_thm_optn  = pt_inst RS pt_optn_inst; 
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   550
          val pt_thm_list  = pt_inst RS pt_list_inst;
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   551
          val pt_thm_prod  = pt_inst RS (pt_inst RS pt_prod_inst);
18600
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   552
          val pt_thm_nprod = pt_inst RS (pt_inst RS pt_nprod_inst);
18430
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   553
          val pt_thm_unit  = pt_unit_inst;
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   554
       in
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   555
        thy
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   556
        |> Axclass.prove_arity (\<^type_name>\<open>fun\<close>,[[cls_name],[cls_name]],[cls_name]) (pt_proof pt_thm_fun)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   557
        |> Axclass.prove_arity (\<^type_name>\<open>set\<close>,[[cls_name]],[cls_name]) (pt_proof pt_thm_set)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   558
        |> Axclass.prove_arity (\<^type_name>\<open>noption\<close>,[[cls_name]],[cls_name]) (pt_proof pt_thm_noptn) 
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   559
        |> Axclass.prove_arity (\<^type_name>\<open>option\<close>,[[cls_name]],[cls_name]) (pt_proof pt_thm_optn)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   560
        |> Axclass.prove_arity (\<^type_name>\<open>list\<close>,[[cls_name]],[cls_name]) (pt_proof pt_thm_list)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   561
        |> Axclass.prove_arity (\<^type_name>\<open>prod\<close>,[[cls_name],[cls_name]],[cls_name]) (pt_proof pt_thm_prod)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   562
        |> Axclass.prove_arity (\<^type_name>\<open>nprod\<close>,[[cls_name],[cls_name]],[cls_name]) (pt_proof pt_thm_nprod)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   563
        |> Axclass.prove_arity (\<^type_name>\<open>unit\<close>,[],[cls_name]) (pt_proof pt_thm_unit)
18430
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   564
     end) ak_names thy13; 
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   565
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   566
       (********  fs_<ak> class instances  ********)
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   567
       (*=========================================*)
18279
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   568
       (* abbreviations for some lemmas *)
23894
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   569
       val fs1            = @{thm "fs1"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   570
       val fs_at_inst     = @{thm "fs_at_inst"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   571
       val fs_unit_inst   = @{thm "fs_unit_inst"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   572
       val fs_prod_inst   = @{thm "fs_prod_inst"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   573
       val fs_nprod_inst  = @{thm "fs_nprod_inst"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   574
       val fs_list_inst   = @{thm "fs_list_inst"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   575
       val fs_option_inst = @{thm "fs_option_inst"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   576
       val dj_supp        = @{thm "dj_supp"};
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   577
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   578
       (* shows that <ak> is an instance of fs_<ak>     *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   579
       (* uses the theorem at_<ak>_inst                 *)
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   580
       val thy20 = fold (fn ak_name => fn thy =>
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   581
        fold (fn ak_name' => fn thy' =>
18437
ee0283e5dfe3 added proofs to show that every atom-kind combination
urbanc
parents: 18436
diff changeset
   582
        let
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   583
           val qu_name =  Sign.full_bname thy' ak_name';
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   584
           val qu_class = Sign.full_bname thy' ("fs_"^ak_name);
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   585
           fun proof ctxt =
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   586
               (if ak_name = ak_name'
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   587
                then
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   588
                  let val at_thm = Global_Theory.get_thm thy' ("at_"^ak_name^"_inst") in
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   589
                    EVERY [Class.intro_classes_tac ctxt [],
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   590
                      resolve_tac ctxt [(at_thm RS fs_at_inst) RS fs1] 1]
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   591
                  end
18437
ee0283e5dfe3 added proofs to show that every atom-kind combination
urbanc
parents: 18436
diff changeset
   592
                else
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   593
                  let val dj_inst = Global_Theory.get_thm thy' ("dj_"^ak_name'^"_"^ak_name);
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   594
                      val simp_s =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   595
                        put_simpset HOL_basic_ss ctxt addsimps [dj_inst RS dj_supp, finite_emptyI];
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58354
diff changeset
   596
                  in EVERY [Class.intro_classes_tac ctxt [], asm_simp_tac simp_s 1] end)
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   597
        in
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   598
         Axclass.prove_arity (qu_name,[],[qu_class]) proof thy'
18437
ee0283e5dfe3 added proofs to show that every atom-kind combination
urbanc
parents: 18436
diff changeset
   599
        end) ak_names thy) ak_names thy18;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   600
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   601
       (* shows that                  *)
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   602
       (*    unit                     *)
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   603
       (*    *(fs_<ak>,fs_<ak>)       *)
18600
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   604
       (*    nprod(fs_<ak>,fs_<ak>)   *)
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   605
       (*    list(fs_<ak>)            *)
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   606
       (*    option(fs_<ak>)          *) 
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   607
       (* are instances of fs_<ak>    *)
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   608
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   609
       val thy24 = fold (fn ak_name => fn thy => 
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   610
        let
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   611
          val cls_name = Sign.full_bname thy ("fs_"^ak_name);
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   612
          val fs_inst  = Global_Theory.get_thm thy ("fs_"^ak_name^"_inst");
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   613
          fun fs_proof thm ctxt =
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   614
            EVERY [Class.intro_classes_tac ctxt [], resolve_tac ctxt [thm RS fs1] 1];
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   615
18600
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   616
          val fs_thm_unit  = fs_unit_inst;
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   617
          val fs_thm_prod  = fs_inst RS (fs_inst RS fs_prod_inst);
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   618
          val fs_thm_nprod = fs_inst RS (fs_inst RS fs_nprod_inst);
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   619
          val fs_thm_list  = fs_inst RS fs_list_inst;
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   620
          val fs_thm_optn  = fs_inst RS fs_option_inst;
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   621
        in 
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   622
         thy
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   623
         |> Axclass.prove_arity (\<^type_name>\<open>unit\<close>,[],[cls_name]) (fs_proof fs_thm_unit) 
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   624
         |> Axclass.prove_arity (\<^type_name>\<open>prod\<close>,[[cls_name],[cls_name]],[cls_name]) (fs_proof fs_thm_prod) 
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   625
         |> Axclass.prove_arity (\<^type_name>\<open>nprod\<close>,[[cls_name],[cls_name]],[cls_name]) (fs_proof fs_thm_nprod)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   626
         |> Axclass.prove_arity (\<^type_name>\<open>list\<close>,[[cls_name]],[cls_name]) (fs_proof fs_thm_list)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   627
         |> Axclass.prove_arity (\<^type_name>\<open>option\<close>,[[cls_name]],[cls_name]) (fs_proof fs_thm_optn)
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   628
        end) ak_names thy20;
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   629
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   630
       (********  cp_<ak>_<ai> class instances  ********)
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   631
       (*==============================================*)
18279
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   632
       (* abbreviations for some lemmas *)
23894
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   633
       val cp1             = @{thm "cp1"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   634
       val cp_unit_inst    = @{thm "cp_unit_inst"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   635
       val cp_bool_inst    = @{thm "cp_bool_inst"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   636
       val cp_prod_inst    = @{thm "cp_prod_inst"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   637
       val cp_list_inst    = @{thm "cp_list_inst"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   638
       val cp_fun_inst     = @{thm "cp_fun_inst"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   639
       val cp_option_inst  = @{thm "cp_option_inst"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   640
       val cp_noption_inst = @{thm "cp_noption_inst"};
46180
72ee700e1d8f Corrected pt_set_inst, added missing cp_set_inst, deleted obsolete
berghofe
parents: 45979
diff changeset
   641
       val cp_set_inst     = @{thm "cp_set_inst"};
23894
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   642
       val pt_perm_compose = @{thm "pt_perm_compose"};
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   643
23894
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   644
       val dj_pp_forget    = @{thm "dj_perm_perm_forget"};
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   645
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   646
       (* shows that <aj> is an instance of cp_<ak>_<ai>  *)
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   647
       (* for every  <ak>/<ai>-combination                *)
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   648
       val thy25 = fold (fn ak_name => fn thy =>
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   649
         fold (fn ak_name' => fn thy' =>
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   650
          fold (fn ak_name'' => fn thy'' =>
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   651
            let
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   652
              val name =  Sign.full_bname thy'' ak_name;
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   653
              val cls_name = Sign.full_bname thy'' ("cp_"^ak_name'^"_"^ak_name'');
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   654
              fun proof ctxt =
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   655
                (if (ak_name'=ak_name'') then 
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   656
                  (let
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   657
                    val pt_inst  = Global_Theory.get_thm thy'' ("pt_"^ak_name''^"_inst");
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   658
                    val at_inst  = Global_Theory.get_thm thy'' ("at_"^ak_name''^"_inst");
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   659
                  in
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58354
diff changeset
   660
                   EVERY [Class.intro_classes_tac ctxt [],
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   661
                     resolve_tac ctxt [at_inst RS (pt_inst RS pt_perm_compose)] 1]
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   662
                  end)
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   663
                else
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   664
                  (let
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   665
                     val dj_inst  = Global_Theory.get_thm thy'' ("dj_"^ak_name''^"_"^ak_name');
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   666
                     val simp_s = put_simpset HOL_basic_ss ctxt addsimps
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   667
                                        ((dj_inst RS dj_pp_forget)::
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   668
                                         (maps (Global_Theory.get_thms thy'')
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   669
                                           [ak_name' ^"_prm_"^ak_name^"_def",
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   670
                                            ak_name''^"_prm_"^ak_name^"_def"]));
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   671
                  in
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58354
diff changeset
   672
                    EVERY [Class.intro_classes_tac ctxt [], simp_tac simp_s 1]
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   673
                  end))
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   674
              in
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   675
                Axclass.prove_arity (name,[],[cls_name]) proof thy''
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   676
              end) ak_names thy') ak_names thy) ak_names thy24;
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   677
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   678
       (* shows that                                                    *) 
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   679
       (*      units                                                    *) 
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   680
       (*      products                                                 *)
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   681
       (*      lists                                                    *)
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   682
       (*      functions                                                *)
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   683
       (*      options                                                  *)
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   684
       (*      noptions                                                 *)
22536
35debf264656 made the type sets instance of the "cp" type-class
urbanc
parents: 22535
diff changeset
   685
       (*      sets                                                     *)
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   686
       (* are instances of cp_<ak>_<ai> for every <ak>/<ai>-combination *)
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   687
       val thy26 = fold (fn ak_name => fn thy =>
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   688
        fold (fn ak_name' => fn thy' =>
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   689
        let
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   690
            val cls_name = Sign.full_bname thy' ("cp_"^ak_name^"_"^ak_name');
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   691
            val cp_inst  = Global_Theory.get_thm thy' ("cp_"^ak_name^"_"^ak_name'^"_inst");
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   692
            val pt_inst  = Global_Theory.get_thm thy' ("pt_"^ak_name^"_inst");
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   693
            val at_inst  = Global_Theory.get_thm thy' ("at_"^ak_name^"_inst");
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   694
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   695
            fun cp_proof thm ctxt =
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60003
diff changeset
   696
              EVERY [Class.intro_classes_tac ctxt [], resolve_tac ctxt [thm RS cp1] 1];
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   697
          
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   698
            val cp_thm_unit = cp_unit_inst;
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   699
            val cp_thm_prod = cp_inst RS (cp_inst RS cp_prod_inst);
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   700
            val cp_thm_list = cp_inst RS cp_list_inst;
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   701
            val cp_thm_fun  = at_inst RS (pt_inst RS (cp_inst RS (cp_inst RS cp_fun_inst)));
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   702
            val cp_thm_optn = cp_inst RS cp_option_inst;
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   703
            val cp_thm_noptn = cp_inst RS cp_noption_inst;
46180
72ee700e1d8f Corrected pt_set_inst, added missing cp_set_inst, deleted obsolete
berghofe
parents: 45979
diff changeset
   704
            val cp_thm_set = cp_inst RS cp_set_inst;
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   705
        in
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   706
         thy'
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   707
         |> Axclass.prove_arity (\<^type_name>\<open>unit\<close>,[],[cls_name]) (cp_proof cp_thm_unit)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   708
         |> Axclass.prove_arity (\<^type_name>\<open>Product_Type.prod\<close>, [[cls_name],[cls_name]],[cls_name]) (cp_proof cp_thm_prod)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   709
         |> Axclass.prove_arity (\<^type_name>\<open>list\<close>,[[cls_name]],[cls_name]) (cp_proof cp_thm_list)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   710
         |> Axclass.prove_arity (\<^type_name>\<open>fun\<close>,[[cls_name],[cls_name]],[cls_name]) (cp_proof cp_thm_fun)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   711
         |> Axclass.prove_arity (\<^type_name>\<open>option\<close>,[[cls_name]],[cls_name]) (cp_proof cp_thm_optn)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   712
         |> Axclass.prove_arity (\<^type_name>\<open>noption\<close>,[[cls_name]],[cls_name]) (cp_proof cp_thm_noptn)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   713
         |> Axclass.prove_arity (\<^type_name>\<open>set\<close>,[[cls_name]],[cls_name]) (cp_proof cp_thm_set)
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   714
        end) ak_names thy) ak_names thy25;
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   715
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   716
     (* show that discrete nominal types are permutation types, finitely     *)
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   717
     (* supported and have the commutation property                          *)
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   718
     (* discrete types have a permutation operation defined as pi o x = x;   *)
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   719
     (* which renders the proofs to be simple "simp_all"-proofs.             *)
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   720
     val thy32 =
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   721
        let
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   722
          fun discrete_pt_inst discrete_ty defn =
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   723
             fold (fn ak_name => fn thy =>
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   724
             let
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   725
               val qu_class = Sign.full_bname thy ("pt_"^ak_name);
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   726
               fun proof ctxt =
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58354
diff changeset
   727
                Class.intro_classes_tac ctxt [] THEN
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   728
                REPEAT (asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [Simpdata.mk_eq defn]) 1);
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   729
             in 
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   730
               Axclass.prove_arity (discrete_ty, [], [qu_class]) proof thy
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   731
             end) ak_names;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   732
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   733
          fun discrete_fs_inst discrete_ty defn = 
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   734
             fold (fn ak_name => fn thy =>
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   735
             let
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   736
               val qu_class = Sign.full_bname thy ("fs_"^ak_name);
44684
8dde3352d5c4 assert Pure equations for theorem references; tuned
haftmann
parents: 41562
diff changeset
   737
               val supp_def = Simpdata.mk_eq @{thm "Nominal.supp_def"};
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   738
               fun proof ctxt =
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58354
diff changeset
   739
                Class.intro_classes_tac ctxt [] THEN
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   740
                asm_simp_tac (put_simpset HOL_ss ctxt
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   741
                  addsimps [supp_def, Collect_const, finite_emptyI, Simpdata.mk_eq defn]) 1;
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   742
             in 
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   743
               Axclass.prove_arity (discrete_ty, [], [qu_class]) proof thy
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   744
             end) ak_names;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   745
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   746
          fun discrete_cp_inst discrete_ty defn = 
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   747
             fold (fn ak_name' => (fold (fn ak_name => fn thy =>
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   748
             let
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28729
diff changeset
   749
               val qu_class = Sign.full_bname thy ("cp_"^ak_name^"_"^ak_name');
44684
8dde3352d5c4 assert Pure equations for theorem references; tuned
haftmann
parents: 41562
diff changeset
   750
               val supp_def = Simpdata.mk_eq @{thm "Nominal.supp_def"};
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   751
               fun proof ctxt =
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58354
diff changeset
   752
                Class.intro_classes_tac ctxt [] THEN
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   753
                asm_simp_tac (put_simpset HOL_ss ctxt addsimps [Simpdata.mk_eq defn]) 1;
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   754
             in
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   755
               Axclass.prove_arity (discrete_ty, [], [qu_class]) proof thy
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   756
             end) ak_names)) ak_names;
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   757
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   758
        in
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   759
         thy26
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   760
         |> discrete_pt_inst \<^type_name>\<open>nat\<close> @{thm perm_nat_def}
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   761
         |> discrete_fs_inst \<^type_name>\<open>nat\<close> @{thm perm_nat_def}
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   762
         |> discrete_cp_inst \<^type_name>\<open>nat\<close> @{thm perm_nat_def}
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   763
         |> discrete_pt_inst \<^type_name>\<open>bool\<close> @{thm perm_bool_def}
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   764
         |> discrete_fs_inst \<^type_name>\<open>bool\<close> @{thm perm_bool_def}
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   765
         |> discrete_cp_inst \<^type_name>\<open>bool\<close> @{thm perm_bool_def}
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   766
         |> discrete_pt_inst \<^type_name>\<open>int\<close> @{thm perm_int_def}
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   767
         |> discrete_fs_inst \<^type_name>\<open>int\<close> @{thm perm_int_def}
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   768
         |> discrete_cp_inst \<^type_name>\<open>int\<close> @{thm perm_int_def}
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   769
         |> discrete_pt_inst \<^type_name>\<open>char\<close> @{thm perm_char_def}
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   770
         |> discrete_fs_inst \<^type_name>\<open>char\<close> @{thm perm_char_def}
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   771
         |> discrete_cp_inst \<^type_name>\<open>char\<close> @{thm perm_char_def}
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   772
        end;
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   773
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   774
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   775
       (* abbreviations for some lemmas *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   776
       (*===============================*)
23894
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   777
       val abs_fun_pi          = @{thm "Nominal.abs_fun_pi"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   778
       val abs_fun_pi_ineq     = @{thm "Nominal.abs_fun_pi_ineq"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   779
       val abs_fun_eq          = @{thm "Nominal.abs_fun_eq"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   780
       val abs_fun_eq'         = @{thm "Nominal.abs_fun_eq'"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   781
       val abs_fun_fresh       = @{thm "Nominal.abs_fun_fresh"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   782
       val abs_fun_fresh'      = @{thm "Nominal.abs_fun_fresh'"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   783
       val dj_perm_forget      = @{thm "Nominal.dj_perm_forget"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   784
       val dj_pp_forget        = @{thm "Nominal.dj_perm_perm_forget"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   785
       val fresh_iff           = @{thm "Nominal.fresh_abs_fun_iff"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   786
       val fresh_iff_ineq      = @{thm "Nominal.fresh_abs_fun_iff_ineq"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   787
       val abs_fun_supp        = @{thm "Nominal.abs_fun_supp"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   788
       val abs_fun_supp_ineq   = @{thm "Nominal.abs_fun_supp_ineq"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   789
       val pt_swap_bij         = @{thm "Nominal.pt_swap_bij"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   790
       val pt_swap_bij'        = @{thm "Nominal.pt_swap_bij'"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   791
       val pt_fresh_fresh      = @{thm "Nominal.pt_fresh_fresh"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   792
       val pt_bij              = @{thm "Nominal.pt_bij"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   793
       val pt_perm_compose     = @{thm "Nominal.pt_perm_compose"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   794
       val pt_perm_compose'    = @{thm "Nominal.pt_perm_compose'"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   795
       val perm_app            = @{thm "Nominal.pt_fun_app_eq"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   796
       val at_fresh            = @{thm "Nominal.at_fresh"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   797
       val at_fresh_ineq       = @{thm "Nominal.at_fresh_ineq"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   798
       val at_calc             = @{thms "Nominal.at_calc"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   799
       val at_swap_simps       = @{thms "Nominal.at_swap_simps"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   800
       val at_supp             = @{thm "Nominal.at_supp"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   801
       val dj_supp             = @{thm "Nominal.dj_supp"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   802
       val fresh_left_ineq     = @{thm "Nominal.pt_fresh_left_ineq"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   803
       val fresh_left          = @{thm "Nominal.pt_fresh_left"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   804
       val fresh_right_ineq    = @{thm "Nominal.pt_fresh_right_ineq"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   805
       val fresh_right         = @{thm "Nominal.pt_fresh_right"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   806
       val fresh_bij_ineq      = @{thm "Nominal.pt_fresh_bij_ineq"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   807
       val fresh_bij           = @{thm "Nominal.pt_fresh_bij"};
26773
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26484
diff changeset
   808
       val fresh_star_bij_ineq = @{thms "Nominal.pt_fresh_star_bij_ineq"};
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26484
diff changeset
   809
       val fresh_star_bij      = @{thms "Nominal.pt_fresh_star_bij"};
23894
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   810
       val fresh_eqvt          = @{thm "Nominal.pt_fresh_eqvt"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   811
       val fresh_eqvt_ineq     = @{thm "Nominal.pt_fresh_eqvt_ineq"};
30086
2023fb9fbf31 Added equivariance lemmas for fresh_star.
berghofe
parents: 29585
diff changeset
   812
       val fresh_star_eqvt     = @{thms "Nominal.pt_fresh_star_eqvt"};
2023fb9fbf31 Added equivariance lemmas for fresh_star.
berghofe
parents: 29585
diff changeset
   813
       val fresh_star_eqvt_ineq= @{thms "Nominal.pt_fresh_star_eqvt_ineq"};
23894
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   814
       val set_diff_eqvt       = @{thm "Nominal.pt_set_diff_eqvt"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   815
       val in_eqvt             = @{thm "Nominal.pt_in_eqvt"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   816
       val eq_eqvt             = @{thm "Nominal.pt_eq_eqvt"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   817
       val all_eqvt            = @{thm "Nominal.pt_all_eqvt"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   818
       val ex_eqvt             = @{thm "Nominal.pt_ex_eqvt"};
28011
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27691
diff changeset
   819
       val ex1_eqvt            = @{thm "Nominal.pt_ex1_eqvt"};
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27691
diff changeset
   820
       val the_eqvt            = @{thm "Nominal.pt_the_eqvt"};
23894
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   821
       val pt_pi_rev           = @{thm "Nominal.pt_pi_rev"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   822
       val pt_rev_pi           = @{thm "Nominal.pt_rev_pi"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   823
       val at_exists_fresh     = @{thm "Nominal.at_exists_fresh"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   824
       val at_exists_fresh'    = @{thm "Nominal.at_exists_fresh'"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   825
       val fresh_perm_app_ineq = @{thm "Nominal.pt_fresh_perm_app_ineq"};
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   826
       val fresh_perm_app      = @{thm "Nominal.pt_fresh_perm_app"};    
23894
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   827
       val fresh_aux           = @{thm "Nominal.pt_fresh_aux"};  
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   828
       val pt_perm_supp_ineq   = @{thm "Nominal.pt_perm_supp_ineq"};
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23158
diff changeset
   829
       val pt_perm_supp        = @{thm "Nominal.pt_perm_supp"};
26090
ec111fa4f8c5 added eqvt-flag to subseteq-lemma
urbanc
parents: 25919
diff changeset
   830
       val subseteq_eqvt       = @{thm "Nominal.pt_subseteq_eqvt"};
22786
d8d7a53ffb63 fixes last commit
narboux
parents: 22785
diff changeset
   831
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   832
       (* Now we collect and instantiate some lemmas w.r.t. all atom      *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   833
       (* types; this allows for example to use abs_perm (which is a      *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   834
       (* collection of theorems) instead of thm abs_fun_pi with explicit *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   835
       (* instantiations.                                                 *)
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   836
       val (_, thy33) =
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   837
         let
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51717
diff changeset
   838
             val ctxt32 = Proof_Context.init_global thy32;
18651
0cb5a8f501aa added the thms-collection "pt_id" (collection of all pt_<ak>1 lemmas)
urbanc
parents: 18626
diff changeset
   839
18279
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   840
             (* takes a theorem thm and a list of theorems [t1,..,tn]            *)
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   841
             (* produces a list of theorems of the form [t1 RS thm,..,tn RS thm] *) 
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   842
             fun instR thm thms = map (fn ti => ti RS thm) thms;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   843
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32952
diff changeset
   844
             (* takes a theorem thm and a list of theorems [(t1,m1),..,(tn,mn)]  *)
26773
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26484
diff changeset
   845
             (* produces a list of theorems of the form [[t1,m1] MRS thm,..,[tn,mn] MRS thm] *) 
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26484
diff changeset
   846
             fun instRR thm thms = map (fn (ti,mi) => [ti,mi] MRS thm) thms;
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26484
diff changeset
   847
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   848
             (* takes two theorem lists (hopefully of the same length ;o)                *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   849
             (* produces a list of theorems of the form                                  *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   850
             (* [t1 RS m1,..,tn RS mn] where [t1,..,tn] is thms1 and [m1,..,mn] is thms2 *) 
18279
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   851
             fun inst_zip thms1 thms2 = map (fn (t1,t2) => t1 RS t2) (thms1 ~~ thms2);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   852
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   853
             (* takes a theorem list of the form [l1,...,ln]              *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   854
             (* and a list of theorem lists of the form                   *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   855
             (* [[h11,...,h1m],....,[hk1,....,hkm]                        *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   856
             (* produces the list of theorem lists                        *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   857
             (* [[l1 RS h11,...,l1 RS h1m],...,[ln RS hk1,...,ln RS hkm]] *)
18279
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   858
             fun inst_mult thms thmss = map (fn (t,ts) => instR t ts) (thms ~~ thmss);
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   859
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   860
             (* FIXME: these lists do not need to be created dynamically again *)
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   861
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   862
             
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   863
             (* list of all at_inst-theorems *)
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   864
             val ats = map (fn ak => Global_Theory.get_thm thy32 ("at_"^ak^"_inst")) ak_names
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   865
             (* list of all pt_inst-theorems *)
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   866
             val pts = map (fn ak => Global_Theory.get_thm thy32 ("pt_"^ak^"_inst")) ak_names
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   867
             (* list of all cp_inst-theorems as a collection of lists*)
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   868
             val cps = 
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   869
                 let fun cps_fun ak1 ak2 =  Global_Theory.get_thm thy32 ("cp_"^ak1^"_"^ak2^"_inst")
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   870
                 in map (fn aki => (map (cps_fun aki) ak_names)) ak_names end; 
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   871
             (* list of all cp_inst-theorems that have different atom types *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   872
             val cps' = 
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   873
                let fun cps'_fun ak1 ak2 = 
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   874
                if ak1=ak2 then NONE else SOME (Global_Theory.get_thm thy32 ("cp_"^ak1^"_"^ak2^"_inst"))
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 31784
diff changeset
   875
                in map (fn aki => (map_filter I (map (cps'_fun aki) ak_names))) ak_names end;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   876
             (* list of all dj_inst-theorems *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   877
             val djs = 
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   878
               let fun djs_fun ak1 ak2 = 
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   879
                     if ak1=ak2 then NONE else SOME(Global_Theory.get_thm thy32 ("dj_"^ak2^"_"^ak1))
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   880
               in map_filter I (map_product djs_fun ak_names ak_names) end;
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   881
             (* list of all fs_inst-theorems *)
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   882
             val fss = map (fn ak => Global_Theory.get_thm thy32 ("fs_"^ak^"_inst")) ak_names
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
   883
             (* list of all at_inst-theorems *)
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   884
             val fs_axs = map (fn ak => Global_Theory.get_thm thy32 ("fs_"^ak^"1")) ak_names
20097
be2d96bbf39c replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents: 20046
diff changeset
   885
25538
58e8ba3b792b map_product and fold_product
haftmann
parents: 24867
diff changeset
   886
             fun inst_pt thms = maps (fn ti => instR ti pts) thms;
58e8ba3b792b map_product and fold_product
haftmann
parents: 24867
diff changeset
   887
             fun inst_at thms = maps (fn ti => instR ti ats) thms;
58e8ba3b792b map_product and fold_product
haftmann
parents: 24867
diff changeset
   888
             fun inst_fs thms = maps (fn ti => instR ti fss) thms;
58e8ba3b792b map_product and fold_product
haftmann
parents: 24867
diff changeset
   889
             fun inst_cp thms cps = flat (inst_mult thms cps);
26773
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26484
diff changeset
   890
             fun inst_pt_at thms = maps (fn ti => instRR ti (pts ~~ ats)) thms;
25538
58e8ba3b792b map_product and fold_product
haftmann
parents: 24867
diff changeset
   891
             fun inst_dj thms = maps (fn ti => instR ti djs) thms;
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   892
             fun inst_pt_pt_at_cp thms = inst_cp (inst_zip ats (inst_zip pts (inst_pt thms))) cps;
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   893
             fun inst_pt_at_fs thms = inst_zip (inst_fs [fs1]) (inst_zip ats (inst_pt thms));
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   894
             fun inst_pt_pt_at_cp thms =
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   895
                 let val i_pt_pt_at = inst_zip ats (inst_zip pts (inst_pt thms));
18436
9649e24bc10e added thms to perm_compose (so far only composition
urbanc
parents: 18435
diff changeset
   896
                     val i_pt_pt_at_cp = inst_cp i_pt_pt_at cps';
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   897
                 in i_pt_pt_at_cp end;
18396
b3e7da94b51f added a fresh_left lemma that contains all instantiation
urbanc
parents: 18381
diff changeset
   898
             fun inst_pt_pt_at_cp_dj thms = inst_zip djs (inst_pt_pt_at_cp thms);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   899
           in
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   900
            thy32 
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   901
            |>   add_thmss_string [(("alpha", inst_pt_at [abs_fun_eq]),[])]
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   902
            ||>> add_thmss_string [(("alpha'", inst_pt_at [abs_fun_eq']),[])]
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   903
            ||>> add_thmss_string [(("alpha_fresh", inst_pt_at [abs_fun_fresh]),[])]
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   904
            ||>> add_thmss_string [(("alpha_fresh'", inst_pt_at [abs_fun_fresh']),[])]
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   905
            ||>> add_thmss_string [(("perm_swap", inst_pt_at [pt_swap_bij] @ inst_pt_at [pt_swap_bij']),[])]
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   906
            ||>> add_thmss_string 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32952
diff changeset
   907
              let val thms1 = inst_at at_swap_simps
27399
1fb3d1219c12 added facts to lemma swap_simps and tuned lemma calc_atms
urbanc
parents: 27216
diff changeset
   908
                  and thms2 = inst_dj [dj_perm_forget]
1fb3d1219c12 added facts to lemma swap_simps and tuned lemma calc_atms
urbanc
parents: 27216
diff changeset
   909
              in [(("swap_simps", thms1 @ thms2),[])] end 
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   910
            ||>> add_thmss_string 
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   911
              let val thms1 = inst_pt_at [pt_pi_rev];
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   912
                  val thms2 = inst_pt_at [pt_rev_pi];
19139
af69e41eab5d improved the decision-procedure for permutations;
urbanc
parents: 19134
diff changeset
   913
              in [(("perm_pi_simp",thms1 @ thms2),[])] end
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   914
            ||>> add_thmss_string [(("perm_fresh_fresh", inst_pt_at [pt_fresh_fresh]),[])]
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   915
            ||>> add_thmss_string [(("perm_bij", inst_pt_at [pt_bij]),[])]
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   916
            ||>> add_thmss_string 
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   917
              let val thms1 = inst_pt_at [pt_perm_compose];
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   918
                  val thms2 = instR cp1 (Library.flat cps');
18436
9649e24bc10e added thms to perm_compose (so far only composition
urbanc
parents: 18435
diff changeset
   919
              in [(("perm_compose",thms1 @ thms2),[])] end
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   920
            ||>> add_thmss_string [(("perm_compose'",inst_pt_at [pt_perm_compose']),[])] 
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   921
            ||>> add_thmss_string [(("perm_app", inst_pt_at [perm_app]),[])]
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   922
            ||>> add_thmss_string [(("supp_atm", (inst_at [at_supp]) @ (inst_dj [dj_supp])),[])]
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   923
            ||>> add_thmss_string [(("exists_fresh", inst_at [at_exists_fresh]),[])]
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   924
            ||>> add_thmss_string [(("exists_fresh'", inst_at [at_exists_fresh']),[])]
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   925
            ||>> add_thmss_string
24569
c80e1871098b Added equivariance lemmas for induct_forall.
berghofe
parents: 24527
diff changeset
   926
              let
c80e1871098b Added equivariance lemmas for induct_forall.
berghofe
parents: 24527
diff changeset
   927
                val thms1 = inst_pt_at [all_eqvt];
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51717
diff changeset
   928
                val thms2 = map (fold_rule ctxt32 [inductive_forall_def]) thms1
24569
c80e1871098b Added equivariance lemmas for induct_forall.
berghofe
parents: 24527
diff changeset
   929
              in
c80e1871098b Added equivariance lemmas for induct_forall.
berghofe
parents: 24527
diff changeset
   930
                [(("all_eqvt", thms1 @ thms2), [NominalThmDecls.eqvt_force_add])]
c80e1871098b Added equivariance lemmas for induct_forall.
berghofe
parents: 24527
diff changeset
   931
              end
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   932
            ||>> add_thmss_string [(("ex_eqvt", inst_pt_at [ex_eqvt]),[NominalThmDecls.eqvt_force_add])]
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   933
            ||>> add_thmss_string [(("ex1_eqvt", inst_pt_at [ex1_eqvt]),[NominalThmDecls.eqvt_force_add])]
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   934
            ||>> add_thmss_string [(("the_eqvt", inst_pt_at [the_eqvt]),[NominalThmDecls.eqvt_force_add])]
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   935
            ||>> add_thmss_string 
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   936
              let val thms1 = inst_at [at_fresh]
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   937
                  val thms2 = inst_dj [at_fresh_ineq]
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   938
              in [(("fresh_atm", thms1 @ thms2),[])] end
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   939
            ||>> add_thmss_string
27399
1fb3d1219c12 added facts to lemma swap_simps and tuned lemma calc_atms
urbanc
parents: 27216
diff changeset
   940
              let val thms1 = inst_at at_calc
1fb3d1219c12 added facts to lemma swap_simps and tuned lemma calc_atms
urbanc
parents: 27216
diff changeset
   941
                  and thms2 = inst_dj [dj_perm_forget]
1fb3d1219c12 added facts to lemma swap_simps and tuned lemma calc_atms
urbanc
parents: 27216
diff changeset
   942
              in [(("calc_atm", thms1 @ thms2),[])] end
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   943
            ||>> add_thmss_string
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   944
              let val thms1 = inst_pt_at [abs_fun_pi]
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   945
                  and thms2 = inst_pt_pt_at_cp [abs_fun_pi_ineq]
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   946
              in [(("abs_perm", thms1 @ thms2),[NominalThmDecls.eqvt_add])] end
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   947
            ||>> add_thmss_string
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   948
              let val thms1 = inst_dj [dj_perm_forget]
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   949
                  and thms2 = inst_dj [dj_pp_forget]
18279
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   950
              in [(("perm_dj", thms1 @ thms2),[])] end
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   951
            ||>> add_thmss_string
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   952
              let val thms1 = inst_pt_at_fs [fresh_iff]
18626
b6596f579e40 added some lemmas to the collection "abs_fresh"
urbanc
parents: 18600
diff changeset
   953
                  and thms2 = inst_pt_at [fresh_iff]
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   954
                  and thms3 = inst_pt_pt_at_cp_dj [fresh_iff_ineq]
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   955
            in [(("abs_fresh", thms1 @ thms2 @ thms3),[])] end
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   956
            ||>> add_thmss_string
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   957
              let val thms1 = inst_pt_at [abs_fun_supp]
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   958
                  and thms2 = inst_pt_at_fs [abs_fun_supp]
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   959
                  and thms3 = inst_pt_pt_at_cp_dj [abs_fun_supp_ineq]
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   960
              in [(("abs_supp", thms1 @ thms2 @ thms3),[])] end
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   961
            ||>> add_thmss_string
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   962
              let val thms1 = inst_pt_at [fresh_left]
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   963
                  and thms2 = inst_pt_pt_at_cp [fresh_left_ineq]
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   964
              in [(("fresh_left", thms1 @ thms2),[])] end
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   965
            ||>> add_thmss_string
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   966
              let val thms1 = inst_pt_at [fresh_right]
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   967
                  and thms2 = inst_pt_pt_at_cp [fresh_right_ineq]
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   968
              in [(("fresh_right", thms1 @ thms2),[])] end
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   969
            ||>> add_thmss_string
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   970
              let val thms1 = inst_pt_at [fresh_bij]
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   971
                  and thms2 = inst_pt_pt_at_cp [fresh_bij_ineq]
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   972
              in [(("fresh_bij", thms1 @ thms2),[])] end
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   973
            ||>> add_thmss_string
26773
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26484
diff changeset
   974
              let val thms1 = inst_pt_at fresh_star_bij
30086
2023fb9fbf31 Added equivariance lemmas for fresh_star.
berghofe
parents: 29585
diff changeset
   975
                  and thms2 = maps (fn ti => inst_pt_pt_at_cp [ti]) fresh_star_bij_ineq
26773
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26484
diff changeset
   976
              in [(("fresh_star_bij", thms1 @ thms2),[])] end
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   977
            ||>> add_thmss_string
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   978
              let val thms1 = inst_pt_at [fresh_eqvt]
22535
cbee450f88a6 added extended the lemma for equivariance of freshness
urbanc
parents: 22418
diff changeset
   979
                  and thms2 = inst_pt_pt_at_cp_dj [fresh_eqvt_ineq]
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   980
              in [(("fresh_eqvt", thms1 @ thms2),[NominalThmDecls.eqvt_add])] end
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   981
            ||>> add_thmss_string
30086
2023fb9fbf31 Added equivariance lemmas for fresh_star.
berghofe
parents: 29585
diff changeset
   982
              let val thms1 = inst_pt_at fresh_star_eqvt
2023fb9fbf31 Added equivariance lemmas for fresh_star.
berghofe
parents: 29585
diff changeset
   983
                  and thms2 = maps (fn ti => inst_pt_pt_at_cp_dj [ti]) fresh_star_eqvt_ineq
2023fb9fbf31 Added equivariance lemmas for fresh_star.
berghofe
parents: 29585
diff changeset
   984
              in [(("fresh_star_eqvt", thms1 @ thms2),[NominalThmDecls.eqvt_add])] end
2023fb9fbf31 Added equivariance lemmas for fresh_star.
berghofe
parents: 29585
diff changeset
   985
            ||>> add_thmss_string
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   986
              let val thms1 = inst_pt_at [in_eqvt]
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   987
              in [(("in_eqvt", thms1),[NominalThmDecls.eqvt_add])] end
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   988
            ||>> add_thmss_string
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   989
              let val thms1 = inst_pt_at [eq_eqvt]
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   990
              in [(("eq_eqvt", thms1),[NominalThmDecls.eqvt_add])] end
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   991
            ||>> add_thmss_string
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   992
              let val thms1 = inst_pt_at [set_diff_eqvt]
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   993
              in [(("set_diff_eqvt", thms1),[NominalThmDecls.eqvt_add])] end
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   994
            ||>> add_thmss_string
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   995
              let val thms1 = inst_pt_at [subseteq_eqvt]
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   996
              in [(("subseteq_eqvt", thms1),[NominalThmDecls.eqvt_add])] end
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
   997
            ||>> add_thmss_string
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   998
              let val thms1 = inst_pt_at [fresh_aux]
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
   999
                  and thms2 = inst_pt_pt_at_cp_dj [fresh_perm_app_ineq] 
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
  1000
              in  [(("fresh_aux", thms1 @ thms2),[])] end  
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
  1001
            ||>> add_thmss_string
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
  1002
              let val thms1 = inst_pt_at [fresh_perm_app]
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
  1003
                  and thms2 = inst_pt_pt_at_cp_dj [fresh_perm_app_ineq] 
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
  1004
              in  [(("fresh_perm_app", thms1 @ thms2),[])] end 
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
  1005
            ||>> add_thmss_string
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
  1006
              let val thms1 = inst_pt_at [pt_perm_supp]
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
  1007
                  and thms2 = inst_pt_pt_at_cp [pt_perm_supp_ineq] 
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
  1008
              in  [(("supp_eqvt", thms1 @ thms2),[NominalThmDecls.eqvt_add])] end  
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29128
diff changeset
  1009
            ||>> add_thmss_string [(("fin_supp",fs_axs),[])]
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 26090
diff changeset
  1010
           end;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
  1011
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
  1012
    in 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
  1013
      NominalData.map (fold Symtab.update (full_ak_names ~~ map make_atom_info
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
  1014
        (pt_ax_classes ~~
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22274
diff changeset
  1015
         fs_ax_classes ~~
28729
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
  1016
         map (fn cls => Symtab.make (full_ak_names ~~ cls)) cp_ax_classes ~~
28372
291e7a158e95 Added some more theorems to NominalData.
berghofe
parents: 28083
diff changeset
  1017
         prm_cons_thms ~~ prm_inst_thms ~~
28729
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
  1018
         map (fn ths => Symtab.make (full_ak_names ~~ ths)) cp_thms ~~
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
  1019
         map (fn thss => Symtab.make
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 31784
diff changeset
  1020
           (map_filter (fn (s, [th]) => SOME (s, th) | _ => NONE)
28729
cfd169f1dae2 Some more functions for accessing information about atoms.
berghofe
parents: 28372
diff changeset
  1021
              (full_ak_names ~~ thss))) dj_thms))) thy33
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
  1022
    end;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
  1023
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
  1024
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
  1025
(* syntax und parsing *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
  1026
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24712
diff changeset
  1027
val _ =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
  1028
  Outer_Syntax.command \<^command_keyword>\<open>atom_decl\<close> "declare new kinds of atoms"
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 33522
diff changeset
  1029
    (Scan.repeat1 Parse.name >> (Toplevel.theory o create_nom_typedecls));
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
  1030
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
  1031
end;